From b7ed040f1efef5de31c216017cd824bb836b1589 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 12 Apr 2023 17:31:00 +0900 Subject: [PATCH 01/48] probabilistic language using mca's kernels Co-authored-by: Cyril Cohen Co-authored-by: @AyumuSaito --- _CoqProject | 2 + theories/Make | 2 + theories/prob_lang.v | 1003 ++++++++++++++++++++++++++++++++++++++ theories/prob_lang_wip.v | 148 ++++++ 4 files changed, 1155 insertions(+) create mode 100644 theories/prob_lang.v create mode 100644 theories/prob_lang_wip.v diff --git a/_CoqProject b/_CoqProject index 4a35dbab15..f318ea26ef 100644 --- a/_CoqProject +++ b/_CoqProject @@ -136,3 +136,5 @@ theories/showcase/pnt.v analysis_stdlib/Rstruct_topology.v analysis_stdlib/showcase/uniform_bigO.v +theories/prob_lang.v +theories/prob_lang_wip.v diff --git a/theories/Make b/theories/Make index 98c1b98895..c3b9137f8f 100644 --- a/theories/Make +++ b/theories/Make @@ -97,3 +97,5 @@ pi_irrational.v gauss_integral.v all_analysis.v showcase/summability.v +prob_lang.v +prob_lang_wip.v diff --git a/theories/prob_lang.v b/theories/prob_lang.v new file mode 100644 index 0000000000..bffc1621c9 --- /dev/null +++ b/theories/prob_lang.v @@ -0,0 +1,1003 @@ +(* mathcomp analysis (c) 2022 Inria and AIST. License: CeCILL-C. *) +From HB Require Import structures. +From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. +From mathcomp Require Import rat. +From mathcomp.classical Require Import mathcomp_extra boolp classical_sets. +From mathcomp.classical Require Import functions cardinality fsbigop. +Require Import reals ereal signed topology normedtype sequences esum measure. +Require Import lebesgue_measure numfun lebesgue_integral exp kernel. + +(******************************************************************************) +(* Semantics of a probabilistic programming language using s-finite kernels *) +(* *) +(* bernoulli r1 == Bernoulli probability with r1 a proof that *) +(* r : {nonneg R} is smaller than 1 *) +(* *) +(* sample P == sample according to the probability P *) +(* letin l k == execute l, augment the context, and execute k *) +(* ret mf == access the context with f and return the result *) +(* score mf == observe t from d, where f is the density of d and *) +(* t occurs in f *) +(* e.g., score (r e^(-r * t)) = observe t from exp(r) *) +(* pnormalize k P == normalize the kernel k into a probability kernel, *) +(* P is a default probability in case normalization is *) +(* not possible *) +(* ite mf k1 k2 == access the context with the boolean function f and *) +(* behaves as k1 or k2 according to the result *) +(* *) +(* poisson == Poisson distribution function *) +(* exp_density == density function for exponential distribution *) +(* *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. +Import Order.TTheory GRing.Theory Num.Def Num.ExtraDef Num.Theory. +Import numFieldTopology.Exports. + +Local Open Scope classical_set_scope. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. + +(* TODO: PR *) +Lemma onem1' (R : numDomainType) (p : R) : (p + `1- p = 1)%R. +Proof. by rewrite /onem addrCA subrr addr0. Qed. + +Lemma onem_nonneg_proof (R : numDomainType) (p : {nonneg R}) : + (p%:num <= 1 -> 0 <= `1-(p%:num))%R. +Proof. by rewrite /onem/= subr_ge0. Qed. + +Definition onem_nonneg (R : numDomainType) (p : {nonneg R}) + (p1 : (p%:num <= 1)%R) := + NngNum (onem_nonneg_proof p1). +(* /TODO: PR *) + +Section bernoulli. +Variables (R : realType) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). +Local Open Scope ring_scope. + +Definition bernoulli : set _ -> \bar R := + measure_add + [the measure _ _ of mscale p [the measure _ _ of dirac true]] + [the measure _ _ of mscale (onem_nonneg p1) [the measure _ _ of dirac false]]. + +HB.instance Definition _ := Measure.on bernoulli. + +Local Close Scope ring_scope. + +Let bernoulli_setT : bernoulli [set: _] = 1. +Proof. +rewrite /bernoulli/= /measure_add/= /msum 2!big_ord_recr/= big_ord0 add0e/=. +by rewrite /mscale/= !diracT !mule1 -EFinD onem1'. +Qed. + +HB.instance Definition _ := @Measure_isProbability.Build _ _ R bernoulli bernoulli_setT. + +End bernoulli. + +Section mscore. +Context d (T : measurableType d) (R : realType). +Variable f : T -> R. + +Definition mscore t : {measure set _ -> \bar R} := + let p := NngNum (normr_ge0 (f t)) in + [the measure _ _ of mscale p [the measure _ _ of dirac tt]]. + +Lemma mscoreE t U : mscore t U = if U == set0 then 0 else `| (f t)%:E |. +Proof. +rewrite /mscore/= /mscale/=; have [->|->] := set_unit U. + by rewrite eqxx dirac0 mule0. +by rewrite diracT mule1 (negbTE setT0). +Qed. + +Lemma measurable_fun_mscore U : measurable_fun setT f -> + measurable_fun setT (mscore ^~ U). +Proof. +move=> mr; under eq_fun do rewrite mscoreE/=. +have [U0|U0] := eqVneq U set0; first exact: measurable_cst. +by apply: measurableT_comp => //; exact: measurableT_comp. +Qed. + +End mscore. + +(* decomposition of score into finite kernels *) +Module SCORE. +Section score. +Context d (T : measurableType d) (R : realType). +Variable f : T -> R. + +Definition k (mf : measurable_fun setT f) i t U := + if i%:R%:E <= mscore f t U < i.+1%:R%:E then + mscore f t U + else + 0. + +Hypothesis mf : measurable_fun setT f. + +Lemma k0 i t : k mf i t (set0 : set unit) = 0 :> \bar R. +Proof. by rewrite /k measure0; case: ifP. Qed. + +Lemma k_ge0 i t B : 0 <= k mf i t B. +Proof. by rewrite /k; case: ifP. Qed. + +Lemma k_sigma_additive i t : semi_sigma_additive (k mf i t). +Proof. +move=> /= F mF tF mUF; rewrite /k /=. +have [F0|UF0] := eqVneq (\bigcup_n F n) set0. + rewrite F0 measure0 (_ : (fun _ => _) = cst 0). + by case: ifPn => _; exact: cvg_cst. + apply/funext => k; rewrite big1// => n _. + by move: F0 => /bigcup0P -> //; rewrite measure0; case: ifPn. +move: (UF0) => /eqP/bigcup0P/existsNP[m /not_implyP[_ /eqP Fm0]]. +rewrite [in X in _ --> X]mscoreE (negbTE UF0). +rewrite -(cvg_shiftn m.+1)/=. +case: ifPn => ir. + rewrite (_ : (fun _ => _) = cst `|(f t)%:E|); first exact: cvg_cst. + apply/funext => n. + rewrite big_mkord (bigD1 (widen_ord (leq_addl n _) (Ordinal (ltnSn m))))//=. + rewrite [in X in X + _]mscoreE (negbTE Fm0) ir big1 ?adde0// => /= j jk. + rewrite mscoreE; have /eqP -> : F j == set0. + have [/eqP//|Fjtt] := set_unit (F j). + move/trivIsetP : tF => /(_ j m Logic.I Logic.I jk). + by rewrite Fjtt setTI => /eqP; rewrite (negbTE Fm0). + by rewrite eqxx; case: ifP. +rewrite (_ : (fun _ => _) = cst 0); first exact: cvg_cst. +apply/funext => n. +rewrite big_mkord (bigD1 (widen_ord (leq_addl n _) (Ordinal (ltnSn m))))//=. +rewrite [in X in if X then _ else _]mscoreE (negbTE Fm0) (negbTE ir) add0e. +rewrite big1//= => j jm; rewrite mscoreE; have /eqP -> : F j == set0. + have [/eqP//|Fjtt] := set_unit (F j). + move/trivIsetP : tF => /(_ j m Logic.I Logic.I jm). + by rewrite Fjtt setTI => /eqP; rewrite (negbTE Fm0). +by rewrite eqxx; case: ifP. +Qed. + +HB.instance Definition _ i t := isMeasure.Build _ _ _ + (k mf i t) (k0 i t) (k_ge0 i t) (@k_sigma_additive i t). + +Lemma measurable_fun_k i U : measurable U -> measurable_fun setT (k mf i ^~ U). +Proof. +move=> /= mU; rewrite /k /= (_ : (fun x => _) = + (fun x => if i%:R%:E <= x < i.+1%:R%:E then x else 0) \o (mscore f ^~ U)) //. +apply: measurableT_comp => /=; last exact/measurable_fun_mscore. +rewrite (_ : (fun x => _) = (fun x => x * + (\1_(`[i%:R%:E, i.+1%:R%:E [%classic : set _) x)%:E)); last first. + apply/funext => x; case: ifPn => ix; first by rewrite indicE/= mem_set ?mule1. + by rewrite indicE/= memNset ?mule0// /= in_itv/=; exact/negP. +apply: emeasurable_funM => //=; apply/EFin_measurable_fun. +by rewrite (_ : \1__ = mindic R (emeasurable_itv `[(i%:R)%:E, (i.+1%:R)%:E[)). +Qed. + +Definition mk i t := [the measure _ _ of k mf i t]. + +HB.instance Definition _ i := + isKernel.Build _ _ _ _ _ (mk i) (measurable_fun_k i). + +Lemma mk_uub i : measure_fam_uub (mk i). +Proof. +exists i.+1%:R => /= t; rewrite /k mscoreE setT_unit. +by case: ifPn => //; case: ifPn => // _ /andP[]. +Qed. + +HB.instance Definition _ i := + Kernel_isFinite.Build _ _ _ _ _ (mk i) (mk_uub i). + +End score. +End SCORE. + +Section kscore. +Context d (T : measurableType d) (R : realType). +Variable f : T -> R. + +Definition kscore (mf : measurable_fun setT f) + : T -> {measure set _ -> \bar R} := + mscore f. + +Variable mf : measurable_fun setT f. + +Let measurable_fun_kscore U : measurable U -> + measurable_fun setT (kscore mf ^~ U). +Proof. by move=> /= _; exact: measurable_fun_mscore. Qed. + +HB.instance Definition _ := isKernel.Build _ _ T _ R + (kscore mf) measurable_fun_kscore. + +Import SCORE. + +Let sfinite_kscore : exists k : (R.-fker T ~> _)^nat, + forall x U, measurable U -> + kscore mf x U = mseries (k ^~ x) 0 U. +Proof. +rewrite /=; exists (fun i => [the R.-fker _ ~> _ of mk mf i]) => /= t U mU. +rewrite /mseries /kscore/= mscoreE; case: ifPn => [/eqP U0|U0]. + by apply/esym/eseries0 => i _; rewrite U0 measure0. +rewrite /mk /= /k /= mscoreE (negbTE U0). +apply/esym/cvg_lim => //. +rewrite -(cvg_shiftn `|floor (fine `|(f t)%:E|)|%N.+1)/=. +rewrite (_ : (fun _ => _) = cst `|(f t)%:E|); first exact: cvg_cst. +apply/funext => n. +pose floor_f := widen_ord (leq_addl n `|floor `|f t| |.+1) + (Ordinal (ltnSn `|floor `|f t| |)). +rewrite big_mkord (bigD1 floor_f)//= ifT; last first. + rewrite lee_fin lte_fin; apply/andP; split. + by rewrite natr_absz (@ger0_norm _ (floor `|f t|)) ?floor_ge0 ?floor_le. + rewrite -addn1 natrD natr_absz. + by rewrite (@ger0_norm _ (floor `|f t|)) ?floor_ge0 ?lt_succ_floor. +rewrite big1 ?adde0//= => j jk. +rewrite ifF// lte_fin lee_fin. +move: jk; rewrite neq_ltn/= => /orP[|] jr. +- suff : (j.+1%:R <= `|f t|)%R by rewrite leNgt => /negbTE ->; rewrite andbF. + rewrite (_ : j.+1%:R = j.+1%:~R)// floor_ge_int. + move: jr; rewrite -lez_nat => /le_trans; apply. + by rewrite -[leRHS](@ger0_norm _ (floor `|f t|)) ?floor_ge0. +- suff : (`|f t| < j%:R)%R by rewrite ltNge => /negbTE ->. + move: jr; rewrite -ltz_nat -(@ltr_int R) (@gez0_abs (floor `|f t|)) ?floor_ge0//. + by rewrite ltr_int -floor_lt_int. +Qed. + +HB.instance Definition _ := + @Kernel_isSFinite.Build _ _ _ _ _ (kscore mf) sfinite_kscore. + +End kscore. + +(* decomposition of ite into s-finite kernels *) +Module ITE. +Section ite. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). + +Section kiteT. +Variable k : R.-ker X ~> Y. + +Definition kiteT : X * bool -> {measure set Y -> \bar R} := + fun xb => if xb.2 then k xb.1 else [the measure _ _ of mzero]. + +Let measurable_fun_kiteT U : measurable U -> measurable_fun setT (kiteT ^~ U). +Proof. +move=> /= mcU; rewrite /kiteT. +rewrite (_ : (fun _ => _) = + (fun x => if x.2 then k x.1 U else mzero U)); last first. + by apply/funext => -[t b]/=; case: ifPn. +apply: (@measurable_fun_if_pair _ _ _ _ (k ^~ U) (fun=> mzero U)) => //. +exact/measurable_kernel. +Qed. + +#[export] +HB.instance Definition _ := isKernel.Build _ _ _ _ _ + kiteT measurable_fun_kiteT. +End kiteT. + +Section sfkiteT. +Variable k : R.-sfker X ~> Y. + +Let sfinite_kiteT : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> kiteT k x U = mseries (k_ ^~ x) 0 U. +Proof. +have [k_ hk /=] := sfinite_kernel k. +exists (fun n => [the _.-ker _ ~> _ of kiteT (k_ n)]) => /=. + move=> n; have /measure_fam_uubP[r k_r] := measure_uub (k_ n). + by exists r%:num => /= -[x []]; rewrite /kiteT//= /mzero//. +move=> [x b] U mU; rewrite /kiteT; case: ifPn => hb; first by rewrite hk. +by rewrite /mseries eseries0. +Qed. + +#[export] +HB.instance Definition _ t := @Kernel_isSFinite_subdef.Build _ _ _ _ _ + (kiteT k) sfinite_kiteT. +End sfkiteT. + +Section fkiteT. +Variable k : R.-fker X ~> Y. + +Let kiteT_uub : measure_fam_uub (kiteT k). +Proof. +have /measure_fam_uubP[M hM] := measure_uub k. +exists M%:num => /= -[]; rewrite /kiteT => t [|]/=; first exact: hM. +by rewrite /= /mzero. +Qed. + +#[export] +HB.instance Definition _ t := Kernel_isFinite.Build _ _ _ _ _ + (kiteT k) kiteT_uub. +End fkiteT. + +Section kiteF. +Variable k : R.-ker X ~> Y. + +Definition kiteF : X * bool -> {measure set Y -> \bar R} := + fun xb => if ~~ xb.2 then k xb.1 else [the measure _ _ of mzero]. + +Let measurable_fun_kiteF U : measurable U -> measurable_fun setT (kiteF ^~ U). +Proof. +move=> /= mcU; rewrite /kiteF. +rewrite (_ : (fun x => _) = + (fun x => if x.2 then mzero U else k x.1 U)); last first. + by apply/funext => -[t b]/=; rewrite if_neg//; case: ifPn. +apply: (@measurable_fun_if_pair _ _ _ _ (fun=> mzero U) (k ^~ U)) => //. +exact/measurable_kernel. +Qed. + +#[export] +HB.instance Definition _ := isKernel.Build _ _ _ _ _ + kiteF measurable_fun_kiteF. + +End kiteF. + +Section sfkiteF. +Variable k : R.-sfker X ~> Y. + +Let sfinite_kiteF : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> kiteF k x U = mseries (k_ ^~ x) 0 U. +Proof. +have [k_ hk /=] := sfinite_kernel k. +exists (fun n => [the _.-ker _ ~> _ of kiteF (k_ n)]) => /=. + move=> n; have /measure_fam_uubP[r k_r] := measure_uub (k_ n). + by exists r%:num => /= -[x []]; rewrite /kiteF//= /mzero//. +move=> [x b] U mU; rewrite /kiteF; case: ifPn => hb; first by rewrite hk. +by rewrite /mseries eseries0. +Qed. + +#[export] +HB.instance Definition _ := @Kernel_isSFinite_subdef.Build _ _ _ _ _ + (kiteF k) sfinite_kiteF. + +End sfkiteF. + +Section fkiteF. +Variable k : R.-fker X ~> Y. + +Let kiteF_uub : measure_fam_uub (kiteF k). +Proof. +have /measure_fam_uubP[M hM] := measure_uub k. +by exists M%:num => /= -[]; rewrite /kiteF/= => t; case => //=; rewrite /mzero. +Qed. + +#[export] +HB.instance Definition _ := Kernel_isFinite.Build _ _ _ _ _ + (kiteF k) kiteF_uub. + +End fkiteF. +End ite. +End ITE. + +Section ite. +Context d d' (T : measurableType d) (T' : measurableType d') (R : realType). +Variables (f : T -> bool) (u1 u2 : R.-sfker T ~> T'). + +Definition mite (mf : measurable_fun setT f) : T -> set T' -> \bar R := + fun t => if f t then u1 t else u2 t. + +Variables mf : measurable_fun setT f. + +Let mite0 t : mite mf t set0 = 0. +Proof. by rewrite /mite; case: ifPn. Qed. + +Let mite_ge0 t U : 0 <= mite mf t U. +Proof. by rewrite /mite; case: ifPn. Qed. + +Let mite_sigma_additive t : semi_sigma_additive (mite mf t). +Proof. +by rewrite /mite; case: ifPn => ft; exact: measure_semi_sigma_additive. +Qed. + +HB.instance Definition _ t := isMeasure.Build _ _ _ (mite mf t) + (mite0 t) (mite_ge0 t) (@mite_sigma_additive t). + +Import ITE. + +(* +Definition kite : R.-sfker T ~> T' := + kdirac mf \; kadd (kiteT u1) (kiteF u2). +*) +Definition kite := + [the R.-sfker _ ~> _ of kdirac mf] \; + [the R.-sfker _ ~> _ of kadd + [the R.-sfker _ ~> T' of kiteT u1] + [the R.-sfker _ ~> T' of kiteF u2] ]. + +End ite. + +Section insn2. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). + +Definition ret (f : X -> Y) (mf : measurable_fun setT f) + : R.-pker X ~> Y := [the R.-pker _ ~> _ of kdirac mf]. + +Definition sample (P : pprobability Y R) : R.-pker X ~> Y := + [the R.-pker _ ~> _ of kprobability (measurable_cst P)]. + +Definition normalize (k : R.-sfker X ~> Y) P : X -> probability Y R := + fun x => [the probability _ _ of mnormalize (k x) P]. + +Definition ite (f : X -> bool) (mf : measurable_fun setT f) + (k1 k2 : R.-sfker X ~> Y) : R.-sfker X ~> Y := + locked [the R.-sfker X ~> Y of kite k1 k2 mf]. + +End insn2. +Arguments ret {d d' X Y R f} mf. +Arguments sample {d d' X Y R}. + +Section insn2_lemmas. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). + +Lemma retE (f : X -> Y) (mf : measurable_fun setT f) x : + ret mf x = \d_(f x) :> (_ -> \bar R). +Proof. by []. Qed. + +Lemma sampleE (P : probability Y R) (x : X) : sample P x = P. +Proof. by []. Qed. + +Lemma normalizeE (f : R.-sfker X ~> Y) P x U : + normalize f P x U = + if (f x [set: Y] == 0) || (f x [set: Y] == +oo) then P U + else f x U * ((fine (f x [set: Y]))^-1)%:E. +Proof. by rewrite /normalize /= /mnormalize; case: ifPn. Qed. + +Lemma iteE (f : X -> bool) (mf : measurable_fun setT f) + (k1 k2 : R.-sfker X ~> Y) x : + ite mf k1 k2 x = if f x then k1 x else k2 x. +Proof. +apply/eq_measure/funext => U. +rewrite /ite; unlock => /=. +rewrite /kcomp/= integral_dirac//=. +rewrite indicT mul1e. +rewrite -/(measure_add (ITE.kiteT k1 (x, f x)) (ITE.kiteF k2 (x, f x))). +rewrite measure_addE. +rewrite /ITE.kiteT /ITE.kiteF/=. +by case: ifPn => fx /=; rewrite /mzero ?(adde0,add0e). +Qed. + +End insn2_lemmas. + +Section insn3. +Context d d' d3 (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). + +Definition letin (l : R.-sfker X ~> Y) (k : R.-sfker [the measurableType _ of (X * Y)%type] ~> Z) + : R.-sfker X ~> Z := + [the R.-sfker X ~> Z of l \; k]. + +End insn3. + +Section insn3_lemmas. +Context d d' d3 (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). + +Lemma letinE (l : R.-sfker X ~> Y) (k : R.-sfker [the measurableType _ of (X * Y)%type] ~> Z) x U : + letin l k x U = \int[l x]_y k (x, y) U. +Proof. by []. Qed. + +End insn3_lemmas. + +(* rewriting laws *) +Section letin_return. +Context d d' d3 (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). + +Lemma letin_kret (k : R.-sfker X ~> Y) + (f : X * Y -> Z) (mf : measurable_fun setT f) x U : + measurable U -> + letin k (ret mf) x U = k x (curry f x @^-1` U). +Proof. +move=> mU; rewrite letinE. +under eq_integral do rewrite retE. +rewrite integral_indic ?setIT// -[X in measurable X]setTI. +exact: (measurableT_comp mf). +Qed. + +Lemma letin_retk + (f : X -> Y) (mf : measurable_fun setT f) + (k : R.-sfker [the measurableType _ of (X * Y)%type] ~> Z) + x U : measurable U -> + letin (ret mf) k x U = k (x, f x) U. +Proof. +move=> mU; rewrite letinE retE integral_dirac ?indicT ?mul1e//. +exact: (measurableT_comp (measurable_kernel k _ mU)). +Qed. + +End letin_return. + +Section insn1. +Context d (X : measurableType d) (R : realType). + +Definition score (f : X -> R) (mf : measurable_fun setT f) + : R.-sfker X ~> _ := + [the R.-sfker X ~> _ of kscore mf]. + +End insn1. + +Section hard_constraint. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). + +Definition fail := + letin (score (@measurable_cst _ _ X _ setT (0%R : R))) + (ret (@measurable_cst _ _ _ Y setT point)). + +Lemma failE x U : fail x U = 0. +Proof. by rewrite /fail letinE ge0_integral_mscale//= normr0 mul0e. Qed. + +End hard_constraint. +Arguments fail {d d' X Y R}. + +Module Notations. + +Notation var1of2 := (@measurable_fst _ _ _ _). +Notation var2of2 := (@measurable_snd _ _ _ _). +Notation var1of3 := (measurableT_comp (@measurable_fst _ _ _ _) + (@measurable_fst _ _ _ _)). +Notation var2of3 := (measurableT_comp (@measurable_snd _ _ _ _) + (@measurable_fst _ _ _ _)). +Notation var3of3 := (@measurable_snd _ _ _ _). + +Notation mR := Real_sort__canonical__measure_Measurable. +Notation munit := Datatypes_unit__canonical__measure_Measurable. +Notation mbool := Datatypes_bool__canonical__measure_Measurable. + +End Notations. + +Section cst_fun. +Context d (T : measurableType d) (R : realType). + +Definition kr (r : R) := @measurable_cst _ _ T _ setT r. +Definition k3 : measurable_fun _ _ := kr 3%:R. +Definition k10 : measurable_fun _ _ := kr 10%:R. +Definition ktt := @measurable_cst _ _ T _ setT tt. + +End cst_fun. +Arguments kr {d T R}. +Arguments k3 {d T R}. +Arguments k10 {d T R}. +Arguments ktt {d T}. + +Section insn1_lemmas. +Import Notations. +Context d (T : measurableType d) (R : realType). + +Let kcomp_scoreE d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) + (g : R.-sfker [the measurableType _ of (T1 * unit)%type] ~> T2) + f (mf : measurable_fun setT f) r U : + (score mf \; g) r U = `|f r|%:E * g (r, tt) U. +Proof. +rewrite /= /kcomp /kscore /= ge0_integral_mscale//=. +by rewrite integral_dirac// indicT mul1e. +Qed. + +Lemma scoreE d' (T' : measurableType d') (x : T * T') (U : set T') (f : R -> R) + (r : R) (r0 : (0 <= r)%R) + (f0 : (forall r, 0 <= r -> 0 <= f r)%R) (mf : measurable_fun setT f) : + score (measurableT_comp mf var2of2) + (x, r) (curry (snd \o fst) x @^-1` U) = + (f r)%:E * \d_x.2 U. +Proof. by rewrite /score/= /mscale/= ger0_norm// f0. Qed. + +Lemma score_score (f : R -> R) (g : R * unit -> R) + (mf : measurable_fun setT f) + (mg : measurable_fun setT g) : + letin (score mf) (score mg) = + score (measurable_funM mf (measurableT_comp mg (measurable_pair2 tt))). +Proof. +apply/eq_sfkernel => x U. +rewrite {1}/letin; unlock. +by rewrite kcomp_scoreE/= /mscale/= diracE normrM muleA EFinM. +Qed. + +Import Notations. + +(* hard constraints to express score below 1 *) +Lemma score_fail (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : + score (kr r%:num) = + letin (sample [the probability _ _ of bernoulli r1] : R.-pker T ~> _) + (ite var2of2 (ret ktt) fail). +Proof. +apply/eq_sfkernel => x U. +rewrite letinE/= /sample; unlock. +rewrite integral_measure_add//= ge0_integral_mscale//= ge0_integral_mscale//=. +rewrite integral_dirac//= integral_dirac//= !indicT/= !mul1e. +by rewrite /mscale/= iteE//= iteE//= failE mule0 adde0 ger0_norm. +Qed. + +End insn1_lemmas. + +Section letin_ite. +Context d d2 d3 (T : measurableType d) (T2 : measurableType d2) + (Z : measurableType d3) (R : realType). +Variables (k1 k2 : R.-sfker T ~> Z) (u : R.-sfker [the measurableType _ of (T * Z)%type] ~> T2) + (f : T -> bool) (mf : measurable_fun setT f) + (t : T) (U : set T2). + +Lemma letin_iteT : f t -> letin (ite mf k1 k2) u t U = letin k1 u t U. +Proof. +move=> ftT. +rewrite !letinE/=. +apply: eq_measure_integral => V mV _. +by rewrite iteE ftT. +Qed. + +Lemma letin_iteF : ~~ f t -> letin (ite mf k1 k2) u t U = letin k2 u t U. +Proof. +move=> ftF. +rewrite !letinE/=. +apply: eq_measure_integral => V mV _. +by rewrite iteE (negbTE ftF). +Qed. + +End letin_ite. + +Section letinA. +Context d d' d1 d2 d3 (X : measurableType d) (Y : measurableType d') + (T1 : measurableType d1) (T2 : measurableType d2) (T3 : measurableType d3) + (R : realType). +Import Notations. +Variables (t : R.-sfker X ~> T1) + (u : R.-sfker [the measurableType _ of (X * T1)%type] ~> T2) + (v : R.-sfker [the measurableType _ of (X * T2)%type] ~> Y) + (v' : R.-sfker [the measurableType _ of (X * T1 * T2)%type] ~> Y) + (vv' : forall y, v =1 fun xz => v' (xz.1, y, xz.2)). + +Lemma letinA x A : measurable A -> + letin t (letin u v') x A + = + (letin (letin t u) v) x A. +Proof. +move=> mA. +rewrite !letinE. +under eq_integral do rewrite letinE. +rewrite integral_kcomp; [|by []|]. +- apply: eq_integral => y _. + apply: eq_integral => z _. + by rewrite (vv' y). +exact: (measurableT_comp (measurable_kernel v _ mA)). +Qed. + +End letinA. + +Section letinC. +Context d d1 d' (X : measurableType d) (Y : measurableType d1) + (Z : measurableType d') (R : realType). + +Import Notations. + +Variables (t : R.-sfker Z ~> X) + (t' : R.-sfker [the measurableType _ of (Z * Y)%type] ~> X) + (tt' : forall y, t =1 fun z => t' (z, y)) + (u : R.-sfker Z ~> Y) + (u' : R.-sfker [the measurableType _ of (Z * X)%type] ~> Y) + (uu' : forall x, u =1 fun z => u' (z, x)). + +Definition T z : set X -> \bar R := t z. +Let T0 z : (T z) set0 = 0. Proof. by []. Qed. +Let T_ge0 z x : 0 <= (T z) x. Proof. by []. Qed. +Let T_semi_sigma_additive z : semi_sigma_additive (T z). +Proof. exact: measure_semi_sigma_additive. Qed. +HB.instance Definition _ z := @isMeasure.Build _ R X (T z) (T0 z) (T_ge0 z) + (@T_semi_sigma_additive z). + +Let sfinT z : sfinite_measure (T z). Proof. exact: sfinite_kernel_measure. Qed. +HB.instance Definition _ z := @Measure_isSFinite_subdef.Build _ X R + (T z) (sfinT z). + +Definition U z : set Y -> \bar R := u z. +Let U0 z : (U z) set0 = 0. Proof. by []. Qed. +Let U_ge0 z x : 0 <= (U z) x. Proof. by []. Qed. +Let U_semi_sigma_additive z : semi_sigma_additive (U z). +Proof. exact: measure_semi_sigma_additive. Qed. +HB.instance Definition _ z := @isMeasure.Build _ R Y (U z) (U0 z) (U_ge0 z) + (@U_semi_sigma_additive z). + +Let sfinU z : sfinite_measure (U z). Proof. exact: sfinite_kernel_measure. Qed. +HB.instance Definition _ z := @Measure_isSFinite_subdef.Build _ Y R + (U z) (sfinU z). + +Lemma letinC z A : measurable A -> + letin t + (letin u' + (ret (measurable_fun_prod var2of3 var3of3))) z A = + letin u + (letin t' + (ret (measurable_fun_prod var3of3 var2of3))) z A. +Proof. +move=> mA. +rewrite !letinE. +under eq_integral. + move=> x _. + rewrite letinE -uu'. + under eq_integral do rewrite retE /=. + over. +rewrite (sfinite_Fubini + [the {sfinite_measure set X -> \bar R} of T z] + [the {sfinite_measure set Y -> \bar R} of U z] + (fun x => \d_(x.1, x.2) A ))//; last first. + apply/EFin_measurable_fun => /=; rewrite (_ : (fun x => _) = mindic R mA)//. + by apply/funext => -[]. +rewrite /=. +apply: eq_integral => y _. +by rewrite letinE/= -tt'; apply: eq_integral => // x _; rewrite retE. +Qed. + +End letinC. + +(* sample programs *) + +Section constants. +Variable R : realType. +Local Open Scope ring_scope. + +Lemma onem1S n : `1- (1 / n.+1%:R) = (n%:R / n.+1%:R)%:nng%:num :> R. +Proof. +by rewrite /onem/= -{1}(@divrr _ n.+1%:R) ?unitfE// -mulrBl -natr1 addrK. +Qed. + +Lemma p1S n : (1 / n.+1%:R)%:nng%:num <= 1 :> R. +Proof. by rewrite ler_pdivr_mulr//= mul1r ler1n. Qed. + +Lemma p12 : (1 / 2%:R)%:nng%:num <= 1 :> R. Proof. by rewrite p1S. Qed. + +Lemma p14 : (1 / 4%:R)%:nng%:num <= 1 :> R. Proof. by rewrite p1S. Qed. + +Lemma onem27 : `1- (2 / 7%:R) = (5%:R / 7%:R)%:nng%:num :> R. +Proof. by apply/eqP; rewrite subr_eq/= -mulrDl -natrD divrr// unitfE. Qed. + +Lemma p27 : (2 / 7%:R)%:nng%:num <= 1 :> R. +Proof. by rewrite /= lter_pdivr_mulr// mul1r ler_nat. Qed. + +End constants. +Arguments p12 {R}. +Arguments p14 {R}. +Arguments p27 {R}. + +Section poisson. +Variable R : realType. +Local Open Scope ring_scope. + +(* density function for Poisson *) +Definition poisson k r : R := r ^+ k / k`!%:R^-1 * expR (- r). + +Lemma poisson_ge0 k r : 0 <= r -> 0 <= poisson k r. +Proof. +move=> r0; rewrite /poisson mulr_ge0 ?expR_ge0//. +by rewrite mulr_ge0// exprn_ge0. +Qed. + +Lemma poisson_gt0 k r : 0 < r -> 0 < poisson k.+1 r. +Proof. +move=> r0; rewrite /poisson mulr_gt0 ?expR_gt0//. +by rewrite divr_gt0// ?exprn_gt0// invr_gt0 ltr0n fact_gt0. +Qed. + +Lemma mpoisson k : measurable_fun setT (poisson k). +Proof. +by apply: measurable_funM => /=; [exact: measurable_funM|exact: measurableT_comp]. +Qed. + +Definition poisson3 := poisson 4 3%:R. (* 0.168 *) +Definition poisson10 := poisson 4 10%:R. (* 0.019 *) + +End poisson. + +Section exponential. +Variable R : realType. +Local Open Scope ring_scope. + +(* density function for exponential *) +Definition exp_density x r : R := r * expR (- r * x). + +Lemma exp_density_gt0 x r : 0 < r -> 0 < exp_density x r. +Proof. by move=> r0; rewrite /exp_density mulr_gt0// expR_gt0. Qed. + +Lemma exp_density_ge0 x r : 0 <= r -> 0 <= exp_density x r. +Proof. by move=> r0; rewrite /exp_density mulr_ge0// expR_ge0. Qed. + +Lemma mexp_density x : measurable_fun setT (exp_density x). +Proof. +apply: measurable_funM => //=; apply: measurableT_comp => //. +exact: measurable_funM. +Qed. + +End exponential. + +Lemma letin_sample_bernoulli d d' (T : measurableType d) + (T' : measurableType d') (R : realType)(r : {nonneg R}) (r1 : (r%:num <= 1)%R) + (u : R.-sfker [the measurableType _ of (T * bool)%type] ~> T') x y : + letin (sample [the probability _ _ of bernoulli r1]) u x y = + r%:num%:E * u (x, true) y + (`1- (r%:num))%:E * u (x, false) y. +Proof. +rewrite letinE/=. +rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. +by rewrite !ge0_integral_mscale//= !integral_dirac//= indicT 2!mul1e. +Qed. + +Section sample_and_return. +Import Notations. +Context d (T : measurableType d) (R : realType). + +Definition sample_and_return : R.-sfker T ~> _ := + letin + (sample [the probability _ _ of bernoulli p27]) (* T -> B *) + (ret var2of2) (* T * B -> B *). + +Lemma sample_and_returnE t U : sample_and_return t U = + (2 / 7%:R)%:E * \d_true U + (5%:R / 7%:R)%:E * \d_false U. +Proof. +by rewrite /sample_and_return letin_sample_bernoulli !retE onem27. +Qed. + +End sample_and_return. + +(* trivial example *) +Section sample_and_branch. +Import Notations. +Context d (T : measurableType d) (R : realType). + +(* let x = sample (bernoulli (2/7)) in + let r = case x of {(1, _) => return (k3()), (2, _) => return (k10())} in + return r *) + +Definition sample_and_branch : R.-sfker T ~> mR R := + letin + (sample [the probability _ _ of bernoulli p27]) (* T -> B *) + (ite var2of2 (ret k3) (ret k10)). + +Lemma sample_and_branchE t U : sample_and_branch t U = + (2 / 7%:R)%:E * \d_(3%:R : R) U + + (5%:R / 7%:R)%:E * \d_(10%:R : R) U. +Proof. +by rewrite /sample_and_branch letin_sample_bernoulli/= !iteE !retE onem27. +Qed. + +End sample_and_branch. + +Section bernoulli_and. +Context d (T : measurableType d) (R : realType). +Import Notations. + +Definition mand (x y : T * mbool * mbool -> mbool) + (t : T * mbool * mbool) : mbool := x t && y t. + +Lemma measurable_fun_mand (x y : T * mbool * mbool -> mbool) : + measurable_fun setT x -> measurable_fun setT y -> + measurable_fun setT (mand x y). +Proof. +move=> /= mx my; apply: (@measurable_fun_bool _ _ _ _ true). +rewrite [X in measurable X](_ : _ = + (x @^-1` [set true]) `&` (y @^-1` [set true])); last first. + by rewrite /mand; apply/seteqP; split => z/= /andP. +apply: measurableI. +- by rewrite -[X in measurable X]setTI; exact: mx. +- by rewrite -[X in measurable X]setTI; exact: my. +Qed. + +Definition bernoulli_and : R.-sfker T ~> mbool := + (letin (sample [the probability _ _ of bernoulli p12]) + (letin (sample [the probability _ _ of bernoulli p12]) + (ret (measurable_fun_mand var2of3 var3of3)))). + +Lemma bernoulli_andE t U : + bernoulli_and t U = + sample [the probability _ _ of bernoulli p14] t U. +Proof. +rewrite /bernoulli_and 3!letin_sample_bernoulli/= /mand/= muleDr//= -muleDl//. +rewrite !muleA -addeA -muleDl// -!EFinM !onem1S/= -splitr mulr1. +have -> : (1 / 2 * (1 / 2) = 1 / 4%:R :> R)%R by rewrite mulf_div mulr1// -natrM. +rewrite /bernoulli/= measure_addE/= /mscale/= -!EFinM; congr( _ + (_ * _)%:E). +have -> : (1 / 2 = 2 / 4%:R :> R)%R. + by apply/eqP; rewrite eqr_div// ?pnatr_eq0// mul1r -natrM. +by rewrite onem1S// -mulrDl. +Qed. + +End bernoulli_and. + +Section staton_bus. +Import Notations. +Context d (T : measurableType d) (R : realType) (h : R -> R). +Hypothesis mh : measurable_fun setT h. +Definition kstaton_bus : R.-sfker T ~> mbool := + letin (sample [the probability _ _ of bernoulli p27]) + (letin + (letin (ite var2of2 (ret k3) (ret k10)) + (score (measurableT_comp mh var3of3))) + (ret var2of3)). + +Definition staton_bus := normalize kstaton_bus. + +End staton_bus. + +(* let x = sample (bernoulli (2/7)) in + let r = case x of {(1, _) => return (k3()), (2, _) => return (k10())} in + let _ = score (1/4! r^4 e^-r) in + return x *) +Section staton_bus_poisson. +Import Notations. +Context d (T : measurableType d) (R : realType). +Let poisson4 := @poisson R 4%N. +Let mpoisson4 := @mpoisson R 4%N. + +Definition kstaton_bus_poisson : R.-sfker (mR R) ~> mbool := + kstaton_bus _ mpoisson4. + +Let kstaton_bus_poissonE t U : kstaton_bus_poisson t U = + (2 / 7%:R)%:E * (poisson4 3%:R)%:E * \d_true U + + (5%:R / 7%:R)%:E * (poisson4 10%:R)%:E * \d_false U. +Proof. +rewrite /kstaton_bus. +rewrite letin_sample_bernoulli. +rewrite -!muleA; congr (_ * _ + _ * _). +- rewrite letin_kret//. + rewrite letin_iteT//. + rewrite letin_retk//. + by rewrite scoreE//= => r r0; exact: poisson_ge0. +- by rewrite onem27. + rewrite letin_kret//. + rewrite letin_iteF//. + rewrite letin_retk//. + by rewrite scoreE//= => r r0; exact: poisson_ge0. +Qed. + +(* true -> 2/7 * 0.168 = 2/7 * 3^4 e^-3 / 4! *) +(* false -> 5/7 * 0.019 = 5/7 * 10^4 e^-10 / 4! *) + +Lemma staton_busE P (t : R) U : + let N := ((2 / 7%:R) * poisson4 3%:R + + (5%:R / 7%:R) * poisson4 10%:R)%R in + staton_bus mpoisson4 P t U = + ((2 / 7%:R)%:E * (poisson4 3%:R)%:E * \d_true U + + (5%:R / 7%:R)%:E * (poisson4 10%:R)%:E * \d_false U) * N^-1%:E. +Proof. +rewrite /staton_bus normalizeE /= !kstaton_bus_poissonE !diracT !mule1 ifF //. +apply/negbTE; rewrite gt_eqF// lte_fin. +by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// poisson_gt0// ltr0n. +Qed. + +End staton_bus_poisson. + +(* let x = sample (bernoulli (2/7)) in + let r = case x of {(1, _) => return (k3()), (2, _) => return (k10())} in + let _ = score (r e^-(15/60 r)) in + return x *) +Section staton_bus_exponential. +Import Notations. +Context d (T : measurableType d) (R : realType). +Let exp1560 := @exp_density R (ratr (15%:Q / 60%:Q)). +Let mexp1560 := @mexp_density R (ratr (15%:Q / 60%:Q)). + +(* 15/60 = 0.25 *) + +Definition kstaton_bus_exponential : R.-sfker (mR R) ~> mbool := + kstaton_bus _ mexp1560. + +Let kstaton_bus_exponentialE t U : kstaton_bus_exponential t U = + (2 / 7%:R)%:E * (exp1560 3%:R)%:E * \d_true U + + (5%:R / 7%:R)%:E * (exp1560 10%:R)%:E * \d_false U. +Proof. +rewrite /kstaton_bus. +rewrite letin_sample_bernoulli. +rewrite -!muleA; congr (_ * _ + _ * _). +- rewrite letin_kret//. + rewrite letin_iteT//. + rewrite letin_retk//. + rewrite scoreE//= => r r0; exact: exp_density_ge0. +- by rewrite onem27. + rewrite letin_kret//. + rewrite letin_iteF//. + rewrite letin_retk//. + by rewrite scoreE//= => r r0; exact: exp_density_ge0. +Qed. + +(* true -> 5/7 * 0.019 = 5/7 * 10^4 e^-10 / 4! *) +(* false -> 2/7 * 0.168 = 2/7 * 3^4 e^-3 / 4! *) + +Lemma staton_bus_exponentialE P (t : R) U : + let N := ((2 / 7%:R) * exp1560 3%:R + + (5%:R / 7%:R) * exp1560 10%:R)%R in + staton_bus mexp1560 P t U = + ((2 / 7%:R)%:E * (exp1560 3%:R)%:E * \d_true U + + (5%:R / 7%:R)%:E * (exp1560 10%:R)%:E * \d_false U) * N^-1%:E. +Proof. +rewrite /staton_bus. +rewrite normalizeE /= !kstaton_bus_exponentialE !diracT !mule1 ifF //. +apply/negbTE; rewrite gt_eqF// lte_fin. +by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// exp_density_gt0 ?ltr0n. +Qed. + +End staton_bus_exponential. diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v new file mode 100644 index 0000000000..7b12eed1b6 --- /dev/null +++ b/theories/prob_lang_wip.v @@ -0,0 +1,148 @@ +From HB Require Import structures. +From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. +From mathcomp Require Import rat. +From mathcomp.classical Require Import mathcomp_extra boolp classical_sets. +From mathcomp.classical Require Import functions cardinality fsbigop. +Require Import signed reals ereal topology normedtype sequences esum measure. +Require Import lebesgue_measure numfun lebesgue_integral exp kernel trigo. +Require Import prob_lang. + +(******************************************************************************) +(* Semantics of a probabilistic programming language using s-finite kernels *) +(* (wip) *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. +Import Order.TTheory GRing.Theory Num.Def Num.ExtraDef Num.Theory. +Import numFieldTopology.Exports. + +Local Open Scope classical_set_scope. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. + +Section gauss. +Variable R : realType. +Local Open Scope ring_scope. + +(* density function for gauss *) +Definition gauss_density m s x : R := + (s * sqrtr (pi *+ 2))^-1 * expR (- ((x - m) / s) ^+ 2 / 2%:R). + +Lemma gauss_density_ge0 m s x : 0 <= s -> 0 <= gauss_density m s x. +Proof. by move=> s0; rewrite mulr_ge0 ?expR_ge0// invr_ge0 mulr_ge0. Qed. + +Lemma gauss_density_gt0 m s x : 0 < s -> 0 < gauss_density m s x. +Proof. +move=> s0; rewrite mulr_gt0 ?expR_gt0// invr_gt0 mulr_gt0//. +by rewrite sqrtr_gt0 pmulrn_rgt0// pi_gt0. +Qed. + +Definition gauss01_density : R -> R := gauss_density 0 1. + +Lemma gauss01_densityE x : + gauss01_density x = (sqrtr (pi *+ 2))^-1 * expR (- (x ^+ 2) / 2%:R). +Proof. by rewrite /gauss01_density /gauss_density mul1r subr0 divr1. Qed. + +Definition mgauss01 (V : set R) := + (\int[lebesgue_measure]_(x in V) (gauss01_density x)%:E)%E. + +Lemma measurable_fun_gauss_density m s : + measurable_fun setT (gauss_density m s). +Proof. +apply: measurable_funM => //=. +apply: measurableT_comp => //=. +apply: measurable_funM => //=. +apply: measurableT_comp => //=. +apply: measurableT_comp (measurable_exprn _) _ => /=. +apply: measurable_funM => //=. +exact: measurable_funD. +Qed. + +Let mgauss010 : mgauss01 set0 = 0%E. +Proof. by rewrite /mgauss01 integral_set0. Qed. + +Let mgauss01_ge0 A : (0 <= mgauss01 A)%E. +Proof. +by rewrite /mgauss01 integral_ge0//= => x _; rewrite lee_fin gauss_density_ge0. +Qed. + +Axiom integral_gauss01_density : + (\int[lebesgue_measure]_x (gauss01_density x)%:E = 1%E)%E. + +Let mgauss01_sigma_additive : semi_sigma_additive mgauss01. +Proof. +move=> /= F mF tF mUF. +rewrite /mgauss01/= integral_bigcup//=; last first. + apply/integrableP; split. + apply/EFin_measurable_fun. + exact: measurable_funS (measurable_fun_gauss_density 0 1). + rewrite (_ : (fun x => _) = (EFin \o gauss01_density)); last first. + by apply/funext => x; rewrite gee0_abs// lee_fin gauss_density_ge0. + apply: le_lt_trans. + apply: (@subset_integral _ _ _ _ _ setT) => //=. + apply/EFin_measurable_fun. + exact: measurable_fun_gauss_density. + by move=> ? _; rewrite lee_fin gauss_density_ge0. + by rewrite integral_gauss01_density// ltey. +apply: is_cvg_ereal_nneg_natsum_cond => n _ _. +by apply: integral_ge0 => /= x ?; rewrite lee_fin gauss_density_ge0. +Qed. + +HB.instance Definition _ := isMeasure.Build _ _ _ + mgauss01 mgauss010 mgauss01_ge0 mgauss01_sigma_additive. + +Let mgauss01_setT : mgauss01 [set: _] = 1%E. +Proof. by rewrite /mgauss01 integral_gauss01_density. Qed. + +HB.instance Definition _ := @Measure_isProbability.Build _ _ R mgauss01 mgauss01_setT. + +Definition gauss01 := [the probability _ _ of mgauss01]. + +End gauss. + +Section gauss_lebesgue. +Import Notations. +Context d (T : measurableType d) (R : realType). + +Let f1 (x : R) := (gauss01_density x) ^-1. + +Let mf1 : measurable_fun setT f1. +Proof. +apply: (measurable_comp (F := [set r : R | r != 0%R])) => //. +- exact: open_measurable. +- by move=> /= r [t _ <-]; rewrite gt_eqF// gauss_density_gt0. +- apply: open_continuous_measurable_fun => //. + by apply/in_setP => x /= x0; exact: inv_continuous. +- exact: measurable_fun_gauss_density. +Qed. + +Variable mu : {measure set mR R -> \bar R}. + +Definition staton_lebesgue : R.-sfker T ~> _ := + letin (sample (@gauss01 R)) + (letin + (score (measurableT_comp mf1 var2of2)) + (ret var2of3)). + +Lemma staton_lebesgueE x U : measurable U -> + staton_lebesgue x U = lebesgue_measure U. +Proof. +move=> mU; rewrite [in LHS]/staton_lebesgue/=. +rewrite [in LHS]letinE /=. +transitivity (\int[@mgauss01 R]_(y in U) (f1 y)%:E). + rewrite -[in RHS](setTI U) integral_setI_indic//=. + apply: eq_integral => /= r _. + rewrite letinE/= ge0_integral_mscale//= ger0_norm//; last first. + by rewrite invr_ge0// gauss_density_ge0. + by rewrite integral_dirac// indicT mul1e diracE indicE. +transitivity (\int[lebesgue_measure]_(x in U) (gauss01_density x * f1 x)%:E). + admit. +transitivity (\int[lebesgue_measure]_(x in U) (\1_U x)%:E). + apply: eq_integral => /= y yU. + by rewrite /f1 divrr ?indicE ?yU// unitfE gt_eqF// gauss_density_gt0. +by rewrite integral_indic//= setIid. +Abort. + +End gauss_lebesgue. From 8656f2db72b08e7305294e6e05e0175406df1613 Mon Sep 17 00:00:00 2001 From: AyumuSaito Date: Tue, 14 Feb 2023 22:02:38 +0900 Subject: [PATCH 02/48] syntax and denotational semantics Co-authored-by: Reynald Affeldt --- _CoqProject | 4 + coq-mathcomp-analysis.opam | 1 + theories/Make | 4 + theories/lang_syntax.v | 1119 +++++++++++++++++++++++++++++++ theories/lang_syntax_examples.v | 794 ++++++++++++++++++++++ theories/lang_syntax_toy.v | 551 +++++++++++++++ theories/lang_syntax_util.v | 83 +++ theories/prob_lang.v | 1042 ++++++++++++++++++++++++++-- theories/prob_lang_wip.v | 143 +++- 9 files changed, 3648 insertions(+), 93 deletions(-) create mode 100644 theories/lang_syntax.v create mode 100644 theories/lang_syntax_examples.v create mode 100644 theories/lang_syntax_toy.v create mode 100644 theories/lang_syntax_util.v diff --git a/_CoqProject b/_CoqProject index f318ea26ef..4c99e218df 100644 --- a/_CoqProject +++ b/_CoqProject @@ -138,3 +138,7 @@ analysis_stdlib/Rstruct_topology.v analysis_stdlib/showcase/uniform_bigO.v theories/prob_lang.v theories/prob_lang_wip.v +theories/lang_syntax_util.v +theories/lang_syntax_toy.v +theories/lang_syntax.v +theories/lang_syntax_examples.v diff --git a/coq-mathcomp-analysis.opam b/coq-mathcomp-analysis.opam index a0890eb504..b56eb46df9 100644 --- a/coq-mathcomp-analysis.opam +++ b/coq-mathcomp-analysis.opam @@ -19,6 +19,7 @@ depends: [ "coq-mathcomp-solvable" "coq-mathcomp-field" "coq-mathcomp-bigenough" { (>= "1.0.0") } + "coq-mathcomp-algebra-tactics" { (>= "1.1.1") } ] tags: [ diff --git a/theories/Make b/theories/Make index c3b9137f8f..43b739a23d 100644 --- a/theories/Make +++ b/theories/Make @@ -99,3 +99,7 @@ all_analysis.v showcase/summability.v prob_lang.v prob_lang_wip.v +lang_syntax_util.v +lang_syntax_toy.v +lang_syntax.v +lang_syntax_examples.v diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v new file mode 100644 index 0000000000..253cbe837b --- /dev/null +++ b/theories/lang_syntax.v @@ -0,0 +1,1119 @@ +Require Import String. +From HB Require Import structures. +From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. +From mathcomp Require Import interval_inference. +From mathcomp Require Import mathcomp_extra boolp classical_sets. +From mathcomp Require Import functions cardinality fsbigop. +From mathcomp Require Import reals ereal topology normedtype sequences esum. +From mathcomp Require Import measure lebesgue_measure numfun lebesgue_integral. +From mathcomp Require Import kernel prob_lang lang_syntax_util. +From mathcomp Require Import ring lra. + +(******************************************************************************) +(* Syntax and Evaluation for a Probabilistic Programming Language *) +(* *) +(* typ == syntax for types of data structures *) +(* measurable_of_typ t == the measurable type corresponding to type t *) +(* It is of type {d & measurableType d} *) +(* mtyp_disp t == the display corresponding to type t *) +(* mtyp t == the measurable type corresponding to type t *) +(* It is of type measurableType (mtyp_disp t) *) +(* measurable_of_seq s == the product space corresponding to the *) +(* list s : seq typ *) +(* It is of type {d & measurableType d} *) +(* acc_typ s n == function that access the nth element of s : seq typ *) +(* It is a function from projT2 (measurable_of_seq s) *) +(* to projT2 (measurable_of_typ (nth Unit s n)) *) +(* ctx == type of context *) +(* := seq (string * type) *) +(* mctx_disp g == the display corresponding to the context g *) +(* mctx g := the measurable type corresponding to the context g *) +(* It is formed of nested pairings of measurable *) +(* spaces. It is of type measurableType (mctx_disp g) *) +(* flag == a flag is either D (deterministic) or *) +(* P (probabilistic) *) +(* exp f g t == syntax of expressions with flag f of type t *) +(* context g *) +(* dval R g t == "deterministic value", i.e., *) +(* function from mctx g to mtyp t *) +(* pval R g t == "probabilistic value", i.e., *) +(* s-finite kernel, from mctx g to mtyp t *) +(* mkswap k == given a kernel k : (Y * X) ~> Z, *) +(* returns a kernel of type (X * Y) ~> Z *) +(* letin' := mkcomp \o mkswap *) +(* e -D> f ; mf == the evaluation of the deterministic expression e *) +(* leads to the deterministic value f *) +(* (mf is the proof that f is measurable) *) +(* e -P> k == the evaluation of the probabilistic function f *) +(* leads to the probabilistic value k *) +(* execD e == a dependent pair of a function corresponding to the *) +(* evaluation of e and a proof that this function is *) +(* measurable *) +(* execP e == a s-finite kernel corresponding to the evaluation *) +(* of the probabilistic expression e *) +(* *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import Order.TTheory GRing.Theory Num.Def Num.Theory. +Import numFieldTopology.Exports. + +Reserved Notation "f .; g" (at level 60, right associativity, + format "f .; '/ ' g"). +Reserved Notation "e -D> f ; mf" (at level 40). +Reserved Notation "e -P> k" (at level 40). + +Local Open Scope classical_set_scope. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. + +(* TODO: mv *) +Arguments measurable_fst {d1 d2 T1 T2}. +Arguments measurable_snd {d1 d2 T1 T2}. + +Section mswap. +Context d d' d3 (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). +Variable k : R.-ker Y * X ~> Z. + +Definition mswap xy U := k (swap xy) U. + +Let mswap0 xy : mswap xy set0 = 0. +Proof. done. Qed. + +Let mswap_ge0 x U : 0 <= mswap x U. +Proof. done. Qed. + +Let mswap_sigma_additive x : semi_sigma_additive (mswap x). +Proof. exact: measure_semi_sigma_additive. Qed. + +HB.instance Definition _ x := isMeasure.Build _ _ R + (mswap x) (mswap0 x) (mswap_ge0 x) (@mswap_sigma_additive x). + +Definition mkswap : _ -> {measure set Z -> \bar R} := + fun x => mswap x. + +Let measurable_fun_kswap U : + measurable U -> measurable_fun setT (mkswap ^~ U). +Proof. +move=> mU. +rewrite [X in measurable_fun _ X](_ : _ = k ^~ U \o @swap _ _)//. +apply measurableT_comp => //=; first exact: measurable_kernel. +exact: measurable_swap. +Qed. + +HB.instance Definition _ := isKernel.Build _ _ + (X * Y)%type Z R mkswap measurable_fun_kswap. + +End mswap. + +Section mswap_sfinite_kernel. +Variables (d d' d3 : _) (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). +Variable k : R.-sfker Y * X ~> Z. + +Let mkswap_sfinite : + exists2 k_ : (R.-ker X * Y ~> Z)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> mkswap k x U = kseries k_ x U. +Proof. +have [k_ /= kE] := sfinite_kernel k. +exists (fun n => mkswap (k_ n)). + move=> n. + have /measure_fam_uubP[M hM] := measure_uub (k_ n). + by exists M%:num => x/=; exact: hM. +move=> xy U mU. +by rewrite /mswap/= kE. +Qed. + +HB.instance Definition _ := + Kernel_isSFinite_subdef.Build _ _ _ Z R (mkswap k) mkswap_sfinite. + +End mswap_sfinite_kernel. + +Section kswap_finite_kernel_finite. +Context d d' d3 (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType) + (k : R.-fker Y * X ~> Z). + +Let mkswap_finite : measure_fam_uub (mkswap k). +Proof. +have /measure_fam_uubP[r hr] := measure_uub k. +apply/measure_fam_uubP; exists (PosNum [gt0 of r%:num%R]) => x /=. +exact: hr. +Qed. + +HB.instance Definition _ := + Kernel_isFinite.Build _ _ _ Z R (mkswap k) mkswap_finite. + +End kswap_finite_kernel_finite. + +Notation "l .; k" := (mkcomp l (mkswap k)) : ereal_scope. + +Section letin'. +Variables (d d' d3 : _) (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). + +Definition letin' (l : R.-sfker X ~> Y) (k : R.-sfker Y * X ~> Z) := + locked [the R.-sfker X ~> Z of l .; k]. + +Lemma letin'E (l : R.-sfker X ~> Y) (k : R.-sfker Y * X ~> Z) x U : + letin' l k x U = \int[l x]_y k (y, x) U. +Proof. by rewrite /letin'; unlock. Qed. + +Lemma letin'_letin (l : R.-sfker X ~> Y) (k : R.-sfker Y * X ~> Z) : + letin' l k = letin l (mkswap k). +Proof. by rewrite /letin'; unlock. Qed. + +End letin'. + +Section letin'C. +Import Notations. +Context d d1 d' (X : measurableType d) (Y : measurableType d1) + (Z : measurableType d') (R : realType). +Variables (t : R.-sfker Z ~> X) + (u' : R.-sfker X * Z ~> Y) + (u : R.-sfker Z ~> Y) + (t' : R.-sfker Y * Z ~> X) + (tt' : forall y, t =1 fun z => t' (y, z)) + (uu' : forall x, u =1 fun z => u' (x, z)). + +Definition T' z : set X -> \bar R := t z. +Let T0 z : (T' z) set0 = 0. Proof. by []. Qed. +Let T_ge0 z x : 0 <= (T' z) x. Proof. by []. Qed. +Let T_semi_sigma_additive z : semi_sigma_additive (T' z). +Proof. exact: measure_semi_sigma_additive. Qed. +HB.instance Definition _ z := @isMeasure.Build _ X R (T' z) (T0 z) (T_ge0 z) + (@T_semi_sigma_additive z). + +Let sfinT z : sfinite_measure (T' z). Proof. exact: sfinite_kernel_measure. Qed. +HB.instance Definition _ z := @isSFinite.Build _ X R + (T' z) (sfinT z). + +Definition U' z : set Y -> \bar R := u z. +Let U0 z : (U' z) set0 = 0. Proof. by []. Qed. +Let U_ge0 z x : 0 <= (U' z) x. Proof. by []. Qed. +Let U_semi_sigma_additive z : semi_sigma_additive (U' z). +Proof. exact: measure_semi_sigma_additive. Qed. +HB.instance Definition _ z := @isMeasure.Build _ Y R (U' z) (U0 z) (U_ge0 z) + (@U_semi_sigma_additive z). + +Let sfinU z : sfinite_measure (U' z). Proof. exact: sfinite_kernel_measure. Qed. +HB.instance Definition _ z := @isSFinite.Build _ Y R (U' z) (sfinU z). + +Lemma letin'C z A : measurable A -> + letin' t + (letin' u' + (ret (measurable_fun_prod macc1of3' macc0of3'))) z A = + letin' u + (letin' t' + (ret (measurable_fun_prod macc0of3' macc1of3'))) z A. +Proof. +move=> mA. +rewrite !letin'E. +under eq_integral. + move=> x _. + rewrite letin'E -uu'. + under eq_integral do rewrite retE /=. + over. +rewrite (sfinite_Fubini (T' z) (U' z) (fun x => \d_(x.1, x.2) A ))//; last first. + apply/EFin_measurable_fun => /=; rewrite (_ : (fun x => _) = mindic R mA)//. + by apply/funext => -[]. +rewrite /=. +apply: eq_integral => y _. +by rewrite letin'E/= -tt'; apply: eq_integral => // x _; rewrite retE. +Qed. + +End letin'C. +Arguments letin'C {d d1 d' X Y Z R} _ _ _ _. + +Section letin'A. +Context d d' d1 d2 d3 (X : measurableType d) (Y : measurableType d') + (T1 : measurableType d1) (T2 : measurableType d2) (T3 : measurableType d3) + (R : realType). +Import Notations. +Variables (t : R.-sfker X ~> T1) + (u : R.-sfker T1 * X ~> T2) + (v : R.-sfker T2 * X ~> Y) + (v' : R.-sfker T2 * (T1 * X) ~> Y) + (vv' : forall y, v =1 fun xz => v' (xz.1, (y, xz.2))). + +Lemma letin'A x A : measurable A -> + letin' t (letin' u v') x A + = + (letin' (letin' t u) v) x A. +Proof. +move=> mA. +rewrite !letin'E. +under eq_integral do rewrite letin'E. +rewrite letin'_letin/=. +rewrite integral_kcomp; [|by []|]. + apply: eq_integral => z _. + apply: eq_integral => y _. + by rewrite (vv' z). +exact: measurableT_comp (@measurable_kernel _ _ _ _ _ v _ mA) _. +Qed. + +End letin'A. + +Declare Scope lang_scope. +Delimit Scope lang_scope with P. + +Section syntax_of_types. +Import Notations. +Context {R : realType}. + +Inductive typ := +| Unit | Bool | Real +| Pair : typ -> typ -> typ +| Prob : typ -> typ. + +HB.instance Definition _ := gen_eqMixin typ. + +Fixpoint measurable_of_typ (t : typ) : {d & measurableType d} := + match t with + | Unit => existT _ _ munit + | Bool => existT _ _ mbool + | Real => existT _ _ [the measurableType _ of (mR R)] + | Pair A B => existT _ _ + [the measurableType (projT1 (measurable_of_typ A), + projT1 (measurable_of_typ B)).-prod%mdisp of + (projT2 (measurable_of_typ A) * + projT2 (measurable_of_typ B))%type] + | Prob A => existT _ _ (pprobability (projT2 (measurable_of_typ A)) R) + end. + +Definition mtyp_disp t : measure_display := projT1 (measurable_of_typ t). + +Definition mtyp t : measurableType (mtyp_disp t) := + projT2 (measurable_of_typ t). + +Definition measurable_of_seq (l : seq typ) : {d & measurableType d} := + iter_mprod (map measurable_of_typ l). + +End syntax_of_types. +Arguments measurable_of_typ {R}. +Arguments mtyp {R}. +Arguments measurable_of_seq {R}. + +Section accessor_functions. +Context {R : realType}. + +(* NB: almost the same as acc (map (@measurable_of_typ R) s) n l, + modulo commutativity of map and measurable_of_typ *) +Fixpoint acc_typ (s : seq typ) n : + projT2 (@measurable_of_seq R s) -> + projT2 (measurable_of_typ (nth Unit s n)) := + match s return + projT2 (measurable_of_seq s) -> projT2 (measurable_of_typ (nth Unit s n)) + with + | [::] => match n with | 0 => (fun=> tt) | m.+1 => (fun=> tt) end + | a :: l => match n with + | 0 => fst + | m.+1 => fun H => @acc_typ l m H.2 + end + end. + +(*Definition acc_typ : forall (s : seq typ) n, + projT2 (@measurable_of_seq R s) -> + projT2 (@measurable_of_typ R (nth Unit s n)). +fix H 1. +intros s n x. +destruct s as [|s]. + destruct n as [|n]. + exact tt. + exact tt. +destruct n as [|n]. + exact (fst x). +rewrite /=. +apply H. +exact: (snd x). +Show Proof. +Defined.*) + +Lemma measurable_acc_typ (s : seq typ) n : measurable_fun setT (@acc_typ s n). +Proof. +elim: s n => //= h t ih [|m]; first exact: measurable_fst. +by apply: (measurableT_comp (ih _)); exact: measurable_snd. +Qed. + +End accessor_functions. +Arguments acc_typ {R} s n. +Arguments measurable_acc_typ {R} s n. + +Section context. +Variables (R : realType). +Definition ctx := seq (string * typ). + +Definition mctx_disp (g : ctx) := projT1 (@measurable_of_seq R (map snd g)). + +Definition mctx (g : ctx) : measurableType (mctx_disp g) := + projT2 (@measurable_of_seq R (map snd g)). + +End context. +Arguments mctx {R}. + +Section syntax_of_expressions. +Context {R : realType}. + +Inductive flag := D | P. + +Inductive exp : flag -> ctx -> typ -> Type := +| exp_unit g : exp D g Unit +| exp_bool g : bool -> exp D g Bool +| exp_real g : R -> exp D g Real +| exp_pair g t1 t2 : exp D g t1 -> exp D g t2 -> exp D g (Pair t1 t2) +| exp_proj1 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t1 +| exp_proj2 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t2 +| exp_var g str t : t = lookup Unit g str -> exp D g t +| exp_bernoulli g (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : + exp D g (Prob Bool) +| exp_poisson g : nat -> exp D g Real -> exp D g Real +| exp_normalize g t : exp P g t -> exp D g (Prob t) +| exp_letin g t1 t2 str : exp P g t1 -> exp P ((str, t1) :: g) t2 -> + exp P g t2 +| exp_sample g t : exp D g (Prob t) -> exp P g t +| exp_score g : exp D g Real -> exp P g Unit +| exp_return g t : exp D g t -> exp P g t +| exp_if z g t : exp D g Bool -> exp z g t -> exp z g t -> exp z g t +| exp_weak z g h t x : exp z (g ++ h) t -> + x.1 \notin dom (g ++ h) -> exp z (g ++ x :: h) t. +Arguments exp_var {g} _ {t}. + +Definition exp_var' (str : string) (t : typ) (g : find str t) := + @exp_var (untag (ctx_of g)) str t (ctx_prf g). +Arguments exp_var' str {t} g. + +Lemma exp_var'E str t (f : find str t) H : + exp_var' str f = exp_var str H :> (@exp _ _ _). +Proof. by rewrite /exp_var'; congr exp_var. Qed. + +End syntax_of_expressions. +Arguments exp {R}. +Arguments exp_unit {R g}. +Arguments exp_bool {R g}. +Arguments exp_real {R g}. +Arguments exp_pair {R g} & {t1 t2}. +Arguments exp_var {R g} _ {t} H. +Arguments exp_bernoulli {R g}. +Arguments exp_poisson {R g}. +Arguments exp_normalize {R g _}. +Arguments exp_letin {R g} & {_ _}. +Arguments exp_sample {R g t}. +Arguments exp_score {R g}. +Arguments exp_return {R g} & {_}. +Arguments exp_if {R z g t}. +Arguments exp_weak {R} z g h {t} x. +Arguments exp_var' {R} str {t} g. + +Declare Custom Entry expr. +Notation "[ e ]" := e (e custom expr at level 5) : lang_scope. +Notation "'TT'" := (exp_unit) (in custom expr at level 1) : lang_scope. +Notation "b ':B'" := (@exp_bool _ _ b%bool) + (in custom expr at level 1) : lang_scope. +Notation "r ':R'" := (@exp_real _ _ r%R) + (in custom expr at level 1, format "r :R") : lang_scope. +Notation "'return' e" := (@exp_return _ _ _ e) + (in custom expr at level 2) : lang_scope. +(*Notation "% str" := (@exp_var _ _ str%string _ erefl) + (in custom expr at level 1, format "% str") : lang_scope.*) +(* Notation "% str H" := (@exp_var _ _ str%string _ H) + (in custom expr at level 1, format "% str H") : lang_scope. *) +Notation "# str" := (@exp_var' _ str%string _ _) + (in custom expr at level 1, format "# str"). +Notation "e :+ str" := (exp_weak _ [::] _ (str, _) e erefl) + (in custom expr at level 1) : lang_scope. +Notation "( e1 , e2 )" := (exp_pair e1 e2) + (in custom expr at level 1) : lang_scope. +Notation "\pi_1 e" := (exp_proj1 e) + (in custom expr at level 1) : lang_scope. +Notation "\pi_2 e" := (exp_proj2 e) + (in custom expr at level 1) : lang_scope. +Notation "'let' x ':=' e 'in' f" := (exp_letin x e f) + (in custom expr at level 3, + x constr, + f custom expr at level 3, + left associativity) : lang_scope. +Notation "{ c }" := c (in custom expr, c constr) : lang_scope. +Notation "x" := x + (in custom expr at level 0, x ident) : lang_scope. +Notation "'Sample' e" := (exp_sample e) + (in custom expr at level 2) : lang_scope. +Notation "'Score' e" := (exp_score e) + (in custom expr at level 2) : lang_scope. +Notation "'Normalize' e" := (exp_normalize e) + (in custom expr at level 0) : lang_scope. +Notation "'if' e1 'then' e2 'else' e3" := (exp_if e1 e2 e3) + (in custom expr at level 1) : lang_scope. + +Local Open Scope lang_scope. +Example three_letin {R : realType} x y z : @exp R P [::] _ := + [let x := return {1}:R in + let y := return #x in + let z := return #y in return #z]. + +Section free_vars. +Context {R : realType}. + +Fixpoint free_vars k g t (e : @exp R k g t) : seq string := + match e with + | exp_unit _ => [::] + | exp_bool _ _ => [::] + | exp_real _ _ => [::] + | exp_pair _ _ _ e1 e2 => free_vars e1 ++ free_vars e2 + | exp_proj1 _ _ _ e => free_vars e + | exp_proj2 _ _ _ e => free_vars e + | exp_var _ x _ _ => [:: x] + | exp_bernoulli _ _ _ => [::] + | exp_poisson _ _ e => free_vars e + | exp_normalize _ _ e => free_vars e + | exp_letin _ _ _ x e1 e2 => free_vars e1 ++ rem x (free_vars e2) + | exp_sample _ _ _ => [::] + | exp_score _ e => free_vars e + | exp_return _ _ e => free_vars e + | exp_if _ _ _ e1 e2 e3 => free_vars e1 ++ free_vars e2 ++ free_vars e3 + | exp_weak _ _ _ _ x e _ => rem x.1 (free_vars e) + end. + +End free_vars. + +Definition dval R g t := @mctx R g -> @mtyp R t. +Definition pval R g t := R.-sfker @mctx R g ~> @mtyp R t. + +Section weak. +Context {R : realType}. +Implicit Types (g h : ctx) (x : string * typ). + +Fixpoint mctx_strong g h x (f : @mctx R (g ++ x :: h)) : @mctx R (g ++ h) := + match g as g0 return mctx (g0 ++ x :: h) -> mctx (g0 ++ h) with + | [::] => fun f0 : mctx ([::] ++ x :: h) => let (a, b) := f0 in (fun=> id) a b + | a :: t => uncurry (fun a b => (a, @mctx_strong t h x b)) + end f. + +Definition weak g h x t (f : dval R (g ++ h) t) : dval R (g ++ x :: h) t := + f \o @mctx_strong g h x. + +Lemma measurable_fun_mctx_strong g h x : + measurable_fun setT (@mctx_strong g h x). +Proof. +elim: g h x => [h x|x g ih h x0]; first exact: measurable_snd. +apply/prod_measurable_funP; split. +- rewrite [X in measurable_fun _ X](_ : _ = fst)//. + by apply/funext => -[]. +- rewrite [X in measurable_fun _ X](_ : _ = @mctx_strong g h x0 \o snd). + apply: measurableT_comp; last exact: measurable_snd. + exact: ih. + by apply/funext => -[]. +Qed. + +Lemma measurable_weak g h x t (f : dval R (g ++ h) t) : + measurable_fun setT f -> measurable_fun setT (@weak g h x t f). +Proof. +move=> mf; apply: measurableT_comp; first exact: mf. +exact: measurable_fun_mctx_strong. +Qed. + +Definition kweak g h x t (f : pval R (g ++ h) t) + : @mctx R (g ++ x :: h) -> {measure set @mtyp R t -> \bar R} := + f \o @mctx_strong g h x. + +Section kernel_weak. +Context g h x t (f : pval R (g ++ h) t). + +Let mf U : measurable U -> measurable_fun setT (@kweak g h x t f ^~ U). +Proof. +move=> mU. +rewrite (_ : kweak _ ^~ U = f ^~ U \o @mctx_strong g h x)//. +apply: measurableT_comp => //; first exact: measurable_kernel. +exact: measurable_fun_mctx_strong. +Qed. + +HB.instance Definition _ := isKernel.Build _ _ _ _ _ (@kweak g h x t f) mf. +End kernel_weak. + +Section sfkernel_weak. +Context g h (x : string * typ) t (f : pval R (g ++ h) t). + +Let sf : exists2 s : (R.-ker @mctx R (g ++ x :: h) ~> @mtyp R t)^nat, + forall n, measure_fam_uub (s n) & + forall z U, measurable U -> (@kweak g h x t f) z U = kseries s z U . +Proof. +have [s hs] := sfinite_kernel f. +exists (fun n => @kweak g h x t (s n)). + by move=> n; have [M hM] := measure_uub (s n); exists M => x0; exact: hM. +by move=> z U mU; by rewrite /kweak/= hs. +Qed. + +HB.instance Definition _ := + Kernel_isSFinite_subdef.Build _ _ _ _ _ (@kweak g h x t f) sf. + +End sfkernel_weak. + +Section fkernel_weak. +Context g h x t (f : R.-fker @mctx R (g ++ h) ~> @mtyp R t). + +Let uub : measure_fam_uub (@kweak g h x t f). +Proof. by have [M hM] := measure_uub f; exists M => x0; exact: hM. Qed. + +HB.instance Definition _ := @Kernel_isFinite.Build _ _ _ _ _ + (@kweak g h x t f) uub. +End fkernel_weak. + +End weak. +Arguments weak {R} g h x {t}. +Arguments measurable_weak {R} g h x {t}. +Arguments kweak {R} g h x {t}. + +Section eval. +Context {R : realType}. +Implicit Type (g : ctx) (str : string). +Local Open Scope lang_scope. + +Inductive evalD : forall g t, exp D g t -> + forall f : dval R g t, measurable_fun setT f -> Prop := +| eval_unit g : ([TT] : exp D g _) -D> cst tt ; ktt + +| eval_bool g b : ([b:B] : exp D g _) -D> cst b ; kb b + +| eval_real g r : ([r:R] : exp D g _) -D> cst r ; kr r + +| eval_pair g t1 (e1 : exp D g t1) f1 mf1 t2 (e2 : exp D g t2) f2 mf2 : + e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> + [(e1, e2)] -D> fun x => (f1 x, f2 x) ; measurable_fun_prod mf1 mf2 + +| eval_proj1 g t1 t2 (e : exp D g (Pair t1 t2)) f mf : + e -D> f ; mf -> + [\pi_1 e] -D> fst \o f ; measurableT_comp measurable_fst mf + +| eval_proj2 g t1 t2 (e : exp D g (Pair t1 t2)) f mf : + e -D> f ; mf -> + [\pi_2 e] -D> snd \o f ; measurableT_comp measurable_snd mf + +(* | eval_var g str : let i := index str (dom g) in + [% str] -D> acc_typ (map snd g) i ; measurable_acc_typ (map snd g) i *) + +| eval_var g x H : let i := index x (dom g) in + exp_var x H -D> acc_typ (map snd g) i ; measurable_acc_typ (map snd g) i + +| eval_bernoulli g (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : + (exp_bernoulli r r1 : exp D g _) -D> cst (bernoulli r1) ; + measurable_cst _ + +| eval_poisson g n (e : exp D g _) f mf : + e -D> f ; mf -> + exp_poisson n e -D> poisson n \o f ; + measurableT_comp (measurable_poisson n) mf + +| eval_normalize g t (e : exp P g t) k : + e -P> k -> + [Normalize e] -D> normalize k point ; measurable_fun_mnormalize k + +| evalD_if g t e f mf (e1 : exp D g t) f1 mf1 e2 f2 mf2 : + e -D> f ; mf -> e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> + [if e then e1 else e2] -D> fun x => if f x then f1 x else f2 x ; + measurable_fun_ifT mf mf1 mf2 + +| evalD_weak g h t e x (H : x.1 \notin dom (g ++ h)) f mf : + e -D> f ; mf -> + (exp_weak _ g h x e H : exp _ _ t) -D> weak g h x f ; + measurable_weak g h x f mf + +where "e -D> v ; mv" := (@evalD _ _ e v mv) + +with evalP : forall g t, exp P g t -> pval R g t -> Prop := + +| eval_letin g t1 t2 str (e1 : exp _ g t1) (e2 : exp _ _ t2) k1 k2 : + e1 -P> k1 -> e2 -P> k2 -> + [let str := e1 in e2] -P> letin' k1 k2 + +| eval_sample g t (e : exp _ _ (Prob t)) + (p : mctx g -> pprobability (mtyp t) R) mp : + e -D> p ; mp -> [Sample e] -P> sample p mp + +| eval_score g (e : exp _ g _) f mf : + e -D> f ; mf -> [Score e] -P> kscore mf + +| eval_return g t (e : exp D g t) f mf : + e -D> f ; mf -> [return e] -P> ret mf + +| evalP_if g t e f mf (e1 : exp P g t) k1 e2 k2 : + e -D> f ; mf -> e1 -P> k1 -> e2 -P> k2 -> + [if e then e1 else e2] -P> ite mf k1 k2 + +| evalP_weak g h t (e : exp P (g ++ h) t) x + (H : x.1 \notin dom (g ++ h)) f : + e -P> f -> + exp_weak _ g h x e H -P> kweak g h x f + +where "e -P> v" := (@evalP _ _ e v). + +End eval. + +Notation "e -D> v ; mv" := (@evalD _ _ _ e v mv) : lang_scope. +Notation "e -P> v" := (@evalP _ _ _ e v) : lang_scope. + +Scheme evalD_mut_ind := Induction for evalD Sort Prop +with evalP_mut_ind := Induction for evalP Sort Prop. + +(* properties of the evaluation relation *) +Section eval_prop. +Variables (R : realType). +Local Open Scope lang_scope. + +Lemma evalD_uniq g t (e : exp D g t) (u v : dval R g t) mu mv : + e -D> u ; mu -> e -D> v ; mv -> u = v. +Proof. +move=> hu. +apply: (@evalD_mut_ind R + (fun g t (e : exp D g t) f mf (h1 : e -D> f; mf) => + forall v mv, e -D> v; mv -> f = v) + (fun g t (e : exp P g t) u (h1 : e -P> u) => + forall v, e -P> v -> u = v)); last exact: hu. +all: (rewrite {g t e u v mu mv hu}). +- move=> g {}v {}mv. + inversion 1; subst g0. + by inj_ex H3. +- move=> g b {}v {}mv. + inversion 1; subst g0 b0. + by inj_ex H3. +- move=> g r {}v {}mv. + inversion 1; subst g0 r0. + by inj_ex H3. +- move=> g t1 e1 f1 mf1 t2 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. + simple inversion 1 => //; subst g0. + case: H3 => ? ?; subst t0 t3. + inj_ex H4; case: H4 => He1 He2. + inj_ex He1; subst e0. + inj_ex He2; subst e3. + inj_ex H5; subst v. + by move=> /IH1 <- /IH2 <-. +- move=> g t1 t2 e f mf H ih v mv. + inversion 1; subst g0 t3 t0. + inj_ex H11; subst v. + clear H9. + inj_ex H7; subst e1. + by rewrite (ih _ _ H4). +- move=> g t1 t2 e f mf H ih v mv. + inversion 1; subst g0 t3 t0. + inj_ex H11; subst v. + clear H9. + inj_ex H7; subst e1. + by rewrite (ih _ _ H4). +(* - move=> g str n {}v {}mv. + inversion 1; subst g0. + inj_ex H6; rewrite -H6. + by inj_ex H7. + inj_ex H8; rewrite -H8. + by inj_ex H9. *) +- move=> g str H n {}v {}mv. + inversion 1; subst g0. + (* inj_ex H7; rewrite -H7. + by inj_ex H8. *) + inj_ex H9; rewrite -H9. + by inj_ex H10. +- move=> g r r1 {}v {}mv. + inversion 1; subst g0 r0. + inj_ex H3; subst v. + by have -> : r1 = r3 by []. +- move=> g n e0 f mf ev IH {}v {}mv. + inversion 1; subst g0 n0. + inj_ex H2; subst e0. + inj_ex H4; subst v. + by rewrite (IH _ _ H3). +- move=> g t e0 k ev IH {}v {}mv. + inversion 1; subst g0 t0. + inj_ex H2; subst e0. + inj_ex H4; subst v. + by rewrite (IH _ H3). +- move=> g t e f mf e1 f1 mf1 e2 f2 mf2 ev ih ev1 ih1 ev2 ih2 v m. + inversion 1; subst g0 t0. + inj_ex H2; subst e0. + inj_ex H6; subst e5. + inj_ex H7; subst e6. + inj_ex H9; subst v. + clear H11. + have ? := ih1 _ _ H12; subst f6. + have ? := ih2 _ _ H13; subst f7. + by rewrite (ih _ _ H5). +- move=> g h t e x H f mf ef ih {}v {}mv. + inversion 1; subst t0 g0 h0 x0. + inj_ex H12; subst e1. + inj_ex H14; subst v. + clear H16. + by rewrite (ih _ _ H5). +- move=> g t1 t2 x e1 e2 k1 k2 ev1 IH1 ev2 IH2 k. + inversion 1; subst g0 t0 t3 x. + inj_ex H7; subst k. + inj_ex H6; subst e5. + inj_ex H5; subst e4. + by rewrite (IH1 _ H4) (IH2 _ H8). +- move=> g t e p mp ev IH k. + inversion 1; subst g0. + inj_ex H5; subst t0. + inj_ex H5; subst e1. + inj_ex H7; subst k. + have ? := IH _ _ H3; subst p1. + by have -> : mp = mp1 by []. +- move=> g e f mf ev IH k. + inversion 1; subst g0. + inj_ex H0; subst e0. + inj_ex H4; subst k. + have ? := IH _ _ H2; subst f1. + by have -> : mf = mf0 by []. +- move=> g t e0 f mf ev IH k. + inversion 1; subst g0 t0. + inj_ex H5; subst e1. + inj_ex H7; subst k. + have ? := IH _ _ H3; subst f1. + by have -> : mf = mf1 by []. +- move=> g t e f mf e1 k1 e2 k2 ev ih ev1 ih1 ev2 ih2 k. + inversion 1; subst g0 t0. + inj_ex H0; subst e0. + inj_ex H1; subst e3. + inj_ex H5; subst k. + inj_ex H2; subst e4. + have ? := ih _ _ H6; subst f1. + have -> : mf = mf0 by []. + by rewrite (ih1 _ H7) (ih2 _ H8). +- move=> g h t e x xgh k ek ih. + inversion 1; subst x0 g0 h0 t0. + inj_ex H13; rewrite -H13. + inj_ex H11; subst e1. + by rewrite (ih _ H4). +Qed. + +Lemma evalP_uniq g t (e : exp P g t) (u v : pval R g t) : + e -P> u -> e -P> v -> u = v. +Proof. +move=> eu. +apply: (@evalP_mut_ind R + (fun g t (e : exp D g t) f mf (h : e -D> f; mf) => + forall v mv, e -D> v; mv -> f = v) + (fun g t (e : exp P g t) u (h : e -P> u) => + forall v, e -P> v -> u = v)); last exact: eu. +all: rewrite {g t e u v eu}. +- move=> g {}v {}mv. + inversion 1; subst g0. + by inj_ex H3. +- move=> g b {}v {}mv. + inversion 1; subst g0 b0. + by inj_ex H3. +- move=> g r {}v {}mv. + inversion 1; subst g0 r0. + by inj_ex H3. +- move=> g t1 e1 f1 mf1 t2 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. + simple inversion 1 => //; subst g0. + case: H3 => ? ?; subst t0 t3. + inj_ex H4; case: H4 => He1 He2. + inj_ex He1; subst e0. + inj_ex He2; subst e3. + inj_ex H5; subst v. + move=> e1f0 e2f3. + by rewrite (IH1 _ _ e1f0) (IH2 _ _ e2f3). +- move=> g t1 t2 e f mf H ih v mv. + inversion 1; subst g0 t3 t0. + inj_ex H11; subst v. + clear H9. + inj_ex H7; subst e1. + by rewrite (ih _ _ H4). +- move=> g t1 t2 e f mf H ih v mv. + inversion 1; subst g0 t3 t0. + inj_ex H11; subst v. + clear H9. + inj_ex H7; subst e1. + by rewrite (ih _ _ H4). +(* - move=> g str n {}v {}mv. + inversion 1; subst g0. + inj_ex H6; rewrite -H6. + by inj_ex H7. + inj_ex H8; rewrite -H8. + by inj_ex H9. *) +- move=> g str H n {}v {}mv. + inversion 1; subst g0. + (* inj_ex H7; rewrite -H7. + by inj_ex H8. *) + inj_ex H9; rewrite -H9. + by inj_ex H10. +- move=> g r r1 {}v {}mv. + inversion 1; subst g0 r0. + inj_ex H3; subst v. + by have -> : r1 = r3 by []. +- move=> g n e f mf ev IH {}v {}mv. + inversion 1; subst g0 n0. + inj_ex H2; subst e0. + inj_ex H4; subst v. + inj_ex H5; subst mv. + by rewrite (IH _ _ H3). +- move=> g t e k ev IH {}v {}mv. + inversion 1; subst g0 t0. + inj_ex H2; subst e0. + inj_ex H4; subst v. + inj_ex H5; subst mv. + by rewrite (IH _ H3). +- move=> g t e f mf e1 f1 mf1 e2 f2 mf2 ef ih ef1 ih1 ef2 ih2 {}v {}mv. + inversion 1; subst g0 t0. + inj_ex H2; subst e0. + inj_ex H6; subst e5. + inj_ex H7; subst e6. + inj_ex H9; subst v. + clear H11. + have ? := ih1 _ _ H12; subst f6. + have ? := ih2 _ _ H13; subst f7. + by rewrite (ih _ _ H5). +- move=> g h t e x H f mf ef ih {}v {}mv. + inversion 1; subst x0 g0 h0 t0. + inj_ex H12; subst e1. + inj_ex H14; subst v. + clear H16. + by rewrite (ih _ _ H5). +- move=> g t1 t2 x e1 e2 k1 k2 ev1 IH1 ev2 IH2 k. + inversion 1; subst g0 x t3 t0. + inj_ex H7; subst k. + inj_ex H5; subst e4. + inj_ex H6; subst e5. + by rewrite (IH1 _ H4) (IH2 _ H8). +- move=> g t e p mp ep IH v. + inversion 1; subst g0 t0. + inj_ex H7; subst v. + inj_ex H5; subst e1. + have ? := IH _ _ H3; subst p1. + by have -> : mp = mp1 by []. +- move=> g e f mf ev IH k. + inversion 1; subst g0. + inj_ex H0; subst e0. + inj_ex H4; subst k. + have ? := IH _ _ H2; subst f1. + by have -> : mf = mf0 by []. +- move=> g t e f mf ev IH k. + inversion 1; subst g0 t0. + inj_ex H7; subst k. + inj_ex H5; subst e1. + have ? := IH _ _ H3; subst f1. + by have -> : mf = mf1 by []. +- move=> g t e f mf e1 k1 e2 k2 ev ih ev1 ih1 ev2 ih2 k. + inversion 1; subst g0 t0. + inj_ex H0; subst e0. + inj_ex H1; subst e3. + inj_ex H5; subst k. + inj_ex H2; subst e4. + have ? := ih _ _ H6; subst f1. + have -> : mf0 = mf by []. + by rewrite (ih1 _ H7) (ih2 _ H8). +- move=> g h t e x xgh k ek ih. + inversion 1; subst x0 g0 h0 t0. + inj_ex H13; rewrite -H13. + inj_ex H11; subst e1. + by rewrite (ih _ H4). +Qed. + +Lemma eval_total z g t (e : @exp R z g t) : + (match z with + | D => fun e => exists f mf, e -D> f ; mf + | P => fun e => exists k, e -P> k + end) e. +Proof. +elim: e. +all: rewrite {z g t}. +- by do 2 eexists; exact: eval_unit. +- by do 2 eexists; exact: eval_bool. +- by do 2 eexists; exact: eval_real. +- move=> g t1 t2 e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. + by exists (fun x => (f1 x, f2 x)); eexists; exact: eval_pair. +- move=> g t1 t2 e [f [mf H]]. + by exists (fst \o f); eexists; exact: eval_proj1. +- move=> g t1 t2 e [f [mf H]]. + by exists (snd \o f); eexists; exact: eval_proj2. +- by move=> g x t tE; subst t; eexists; eexists; exact: eval_var. +- by move=> r r1; eexists; eexists; exact: eval_bernoulli. +- move=> g h e [f [mf H]]. + by exists (poisson h \o f); eexists; exact: eval_poisson. +- move=> g t e [k ek]. + by exists (normalize k point); eexists; exact: eval_normalize. +- move=> g t1 t2 x e1 [k1 ev1] e2 [k2 ev2]. + by exists (letin' k1 k2); exact: eval_letin. +- move=> g t e [f [/= mf ef]]. + by eexists; exact: (@eval_sample _ _ _ _ _ mf). +- move=> g e [f [mf f_mf]]. + by exists (kscore mf); exact: eval_score. +- by move=> g t e [f [mf f_mf]]; exists (ret mf); exact: eval_return. +- case. + + move=> g t e1 [f [mf H1]] e2 [f2 [mf2 H2]] e3 [f3 [mf3 H3]]. + by exists (fun g => if f g then f2 g else f3 g), + (measurable_fun_ifT mf mf2 mf3); exact: evalD_if. + + move=> g t e1 [f [mf H1]] e2 [k2 H2] e3 [k3 H3]. + by exists (ite mf k2 k3); exact: evalP_if. +- case=> [g h t x e [f [mf ef]] xgh|g h st x e [k ek] xgh]. + + by exists (weak _ _ _ f), (measurable_weak _ _ _ _ mf); exact/evalD_weak. + + by exists (kweak _ _ _ k); exact: evalP_weak. +Qed. + +Lemma evalD_total g t (e : @exp R D g t) : exists f mf, e -D> f ; mf. +Proof. exact: (eval_total e). Qed. + +Lemma evalP_total g t (e : @exp R P g t) : exists k, e -P> k. +Proof. exact: (eval_total e). Qed. + +End eval_prop. + +Section execution_functions. +Local Open Scope lang_scope. +Context {R : realType}. +Implicit Type g : ctx. + +Definition execD g t (e : exp D g t) : + {f : dval R g t & measurable_fun setT f} := + let: exist _ H := cid (evalD_total e) in + existT _ _ (projT1 (cid H)). + +Lemma eq_execD g t (p1 p2 : @exp R D g t) : + projT1 (execD p1) = projT1 (execD p2) -> execD p1 = execD p2. +Proof. +rewrite /execD /=. +case: cid => /= f1 [mf1 ev1]. +case: cid => /= f2 [mf2 ev2] f12. +subst f2. +have ? : mf1 = mf2 by []. +subst mf2. +congr existT. +rewrite /sval. +case: cid => mf1' ev1'. +have ? : mf1 = mf1' by []. +subst mf1'. +case: cid => mf2' ev2'. +have ? : mf1 = mf2' by []. +by subst mf2'. +Qed. + +Definition execP g t (e : exp P g t) : pval R g t := + projT1 (cid (evalP_total e)). + +Lemma execD_evalD g t e x mx: + @execD g t e = existT _ x mx <-> e -D> x ; mx. +Proof. +rewrite /execD; split. + case: cid => x' [mx' H] [?]; subst x'. + have ? : mx = mx' by []. + by subst mx'. +case: cid => f' [mf' f'mf']/=. +move/evalD_uniq => /(_ _ _ f'mf') => ?; subst f'. +by case: cid => //= ? ?; congr existT. +Qed. + +Lemma evalD_execD g t (e : exp D g t) : + e -D> projT1 (execD e); projT2 (execD e). +Proof. +by rewrite /execD; case: cid => // x [mx xmx]/=; case: cid. +Qed. + +Lemma execP_evalP g t (e : exp P g t) x : + execP e = x <-> e -P> x. +Proof. +rewrite /execP; split; first by move=> <-; case: cid. +case: cid => // x0 Hx0. +by move/evalP_uniq => /(_ _ Hx0) ?; subst x. +Qed. + +Lemma evalP_execP g t (e : exp P g t) : e -P> execP e. +Proof. by rewrite /execP; case: cid. Qed. + +Lemma execD_unit g : @execD g _ [TT] = existT _ (cst tt) ktt. +Proof. exact/execD_evalD/eval_unit. Qed. + +Lemma execD_bool g b : @execD g _ [b:B] = existT _ (cst b) (kb b). +Proof. exact/execD_evalD/eval_bool. Qed. + +Lemma execD_real g r : @execD g _ [r:R] = existT _ (cst r) (kr r). +Proof. exact/execD_evalD/eval_real. Qed. + +Lemma execD_pair g t1 t2 (e1 : exp D g t1) (e2 : exp D g t2) : + let f1 := projT1 (execD e1) in let f2 := projT1 (execD e2) in + let mf1 := projT2 (execD e1) in let mf2 := projT2 (execD e2) in + execD [(e1, e2)] = + @existT _ _ (fun z => (f1 z, f2 z)) + (@measurable_fun_prod _ _ _ (mctx g) (mtyp t1) (mtyp t2) + f1 f2 mf1 mf2). +Proof. +by move=> f1 f2 mf1 mf2; apply/execD_evalD/eval_pair; exact: evalD_execD. +Qed. + +Lemma execD_proj1 g t1 t2 (e : exp D g (Pair t1 t2)) : + let f := projT1 (execD e) in + let mf := projT2 (execD e) in + execD [\pi_1 e] = @existT _ _ (fst \o f) + (measurableT_comp measurable_fst mf). +Proof. +by move=> f mf; apply/execD_evalD/eval_proj1; exact: evalD_execD. +Qed. + +Lemma execD_proj2 g t1 t2 (e : exp D g (Pair t1 t2)) : + let f := projT1 (execD e) in let mf := projT2 (execD e) in + execD [\pi_2 e] = @existT _ _ (snd \o f) + (measurableT_comp measurable_snd mf). +Proof. +by move=> f mf; apply/execD_evalD/eval_proj2; exact: evalD_execD. +Qed. + +Lemma execD_var_erefl g str : let i := index str (dom g) in + @execD g _ (exp_var str erefl) = existT _ (acc_typ (map snd g) i) + (measurable_acc_typ (map snd g) i). +Proof. by move=> i; apply/execD_evalD; exact: eval_var. Qed. + +Lemma execD_var g x (H : nth Unit (map snd g) (index x (dom g)) = lookup Unit g x) : + let i := index x (dom g) in + @execD g _ (exp_var x H) = existT _ (acc_typ (map snd g) i) + (measurable_acc_typ (map snd g) i). +Proof. by move=> i; apply/execD_evalD; exact: eval_var. Qed. + +Lemma execD_bernoulli g r (r1 : (r%:num <= 1)%R) : + @execD g _ (exp_bernoulli r r1) = + existT _ (cst [the probability _ _ of bernoulli r1]) (measurable_cst _). +Proof. exact/execD_evalD/eval_bernoulli. Qed. + +Lemma execD_normalize g t (e : exp P g t) : + @execD g _ [Normalize e] = + existT _ (normalize (execP e) point : _ -> pprobability _ _) + (measurable_fun_mnormalize (execP e)). +Proof. exact/execD_evalD/eval_normalize/evalP_execP. Qed. + +Lemma execD_poisson g n (e : exp D g Real) : + execD (exp_poisson n e) = + existT _ (poisson n \o (projT1 (execD e))) + (measurableT_comp (measurable_poisson n) (projT2 (execD e))). +Proof. exact/execD_evalD/eval_poisson/evalD_execD. Qed. + +Lemma execP_if g st e1 e2 e3 : + @execP g st [if e1 then e2 else e3] = + ite (projT2 (execD e1)) (execP e2) (execP e3). +Proof. +by apply/execP_evalP/evalP_if; [apply: evalD_execD| exact: evalP_execP..]. +Qed. + +Lemma execP_letin g x t1 t2 (e1 : exp P g t1) (e2 : exp P ((x, t1) :: g) t2) : + execP [let x := e1 in e2] = letin' (execP e1) (execP e2) :> (R.-sfker _ ~> _). +Proof. by apply/execP_evalP/eval_letin; exact: evalP_execP. Qed. + +Lemma execP_sample g t (e : @exp R D g (Prob t)) : + let x := execD e in + execP [Sample e] = sample (projT1 x) (projT2 x). +Proof. exact/execP_evalP/eval_sample/evalD_execD. Qed. + +Lemma execP_score g (e : exp D g Real) : + execP [Score e] = score (projT2 (execD e)). +Proof. exact/execP_evalP/eval_score/evalD_execD. Qed. + +Lemma execP_return g t (e : exp D g t) : + execP [return e] = ret (projT2 (execD e)). +Proof. exact/execP_evalP/eval_return/evalD_execD. Qed. + +Lemma execP_weak g h x t (e : exp P (g ++ h) t) + (xl : x.1 \notin dom (g ++ h)) : + execP (exp_weak P g h _ e xl) = kweak _ _ _ (execP e). +Proof. exact/execP_evalP/evalP_weak/evalP_execP. Qed. + +End execution_functions. +Arguments execD_var_erefl {R g} str. +Arguments execP_weak {R} g h x {t} e. +Arguments exp_var'E {R} str. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v new file mode 100644 index 0000000000..87424ac899 --- /dev/null +++ b/theories/lang_syntax_examples.v @@ -0,0 +1,794 @@ +Require Import String. +From HB Require Import structures. +From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. +From mathcomp Require Import interval_inference. +From mathcomp Require Import mathcomp_extra boolp classical_sets. +From mathcomp Require Import functions cardinality fsbigop. +From mathcomp Require Import reals ereal topology normedtype sequences esum. +From mathcomp Require Import measure lebesgue_measure numfun lebesgue_integral. +From mathcomp Require Import kernel prob_lang lang_syntax_util lang_syntax. +From mathcomp Require Import ring lra. + +(******************************************************************************) +(* Examples using the Probabilistic Programming Language of lang_syntax.v *) +(* *) +(* sample_pair_syntax := normalize ( *) +(* let x := sample (bernoulli 1/2) in *) +(* let y := sample (bernoulli 1/3) in *) +(* return (x, y)) *) +(* *) +(* bernoulli13_score := normalize ( *) +(* let x := sample (bernoulli 1/3) in *) +(* let _ := if x then score (1/3) else score (2/3) in *) +(* return x) *) +(* *) +(* bernoulli12_score := normalize ( *) +(* let x := sample (bernoulli 1/2) in *) +(* let _ := if x then score (1/3) else score (2/3) in *) +(* return x) *) +(* *) +(* hard_constraint := let x := Score {0}:R in return TT *) +(* *) +(* associativity of let-in expressions *) +(* *) +(* staton_bus_syntax == example from [Staton, ESOP 2017] *) +(* *) +(* staton_busA_syntax == same as staton_bus_syntax module associativity of *) +(* let-in expression *) +(* *) +(* commutativity of let-in expressions *) +(* *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import Order.TTheory GRing.Theory Num.Def Num.Theory. +Import numFieldTopology.Exports. + +Local Open Scope classical_set_scope. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. +Local Open Scope string_scope. + +(* letin' versions of rewriting laws *) +Lemma letin'_sample_bernoulli d d' (T : measurableType d) + (T' : measurableType d') (R : realType)(r : {nonneg R}) (r1 : (r%:num <= 1)%R) + (u : R.-sfker bool * T ~> T') x y : + letin' (sample_cst (bernoulli r1)) u x y = + r%:num%:E * u (true, x) y + (`1- (r%:num))%:E * u (false, x) y. +Proof. +rewrite letin'E/=. +rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. +by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT 2!mul1e. +Qed. + +Section letin'_return. +Context d d' d3 (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). + +Lemma letin'_kret (k : R.-sfker X ~> Y) + (f : Y * X -> Z) (mf : measurable_fun setT f) x U : + measurable U -> + letin' k (ret mf) x U = k x (curry f ^~ x @^-1` U). +Proof. +move=> mU; rewrite letin'E. +under eq_integral do rewrite retE. +rewrite integral_indic ?setIT// -[X in measurable X]setTI. +exact: (measurableT_comp mf). +Qed. + +Lemma letin'_retk (f : X -> Y) (mf : measurable_fun setT f) + (k : R.-sfker Y * X ~> Z) x U : + measurable U -> letin' (ret mf) k x U = k (f x, x) U. +Proof. +move=> mU; rewrite letin'E retE integral_dirac ?diracT ?mul1e//. +exact: (measurableT_comp (measurable_kernel k _ mU)). +Qed. + +End letin'_return. + +Section letin'_ite. +Context d d2 d3 (T : measurableType d) (T2 : measurableType d2) + (Z : measurableType d3) (R : realType). +Variables (k1 k2 : R.-sfker T ~> Z) + (u : R.-sfker Z * T ~> T2) + (f : T -> bool) (mf : measurable_fun setT f) + (t : T) (U : set T2). + +Lemma letin'_iteT : f t -> letin' (ite mf k1 k2) u t U = letin' k1 u t U. +Proof. +move=> ftT. +rewrite !letin'E/=. +apply: eq_measure_integral => V mV _. +by rewrite iteE ftT. +Qed. + +Lemma letin'_iteF : ~~ f t -> letin' (ite mf k1 k2) u t U = letin' k2 u t U. +Proof. +move=> ftF. +rewrite !letin'E/=. +apply: eq_measure_integral => V mV _. +by rewrite iteE (negbTE ftF). +Qed. + +End letin'_ite. +(* /letin' versions of rewriting laws *) + +Local Open Scope lang_scope. + +Lemma execP_letinL {R : realType} g t1 t2 x (e1 : @exp R P g t1) (e1' : exp P g t1) + (e2 : exp P ((x, t1) :: g) t2) : + forall U, measurable U -> + execP e1 = execP e1' -> + execP [let x := e1 in e2] ^~ U = execP [let x := e1' in e2] ^~ U. +Proof. +by move=> U mU e1e1'; rewrite !execP_letin e1e1'. +Qed. + +Lemma execP_letinR {R : realType} g t1 t2 x (e1 : @exp R P g t1) + (e2 : exp P _ t2) (e2' : exp P ((x, t1) :: g) t2) : + forall U, measurable U -> + execP e2 = execP e2' -> + execP [let x := e1 in e2] ^~ U = execP [let x := e1 in e2'] ^~ U. +Proof. +by move=> U mU e1e1'; rewrite !execP_letin e1e1'. +Qed. + +Local Close Scope lang_scope. + +(* simple tests to check bidirectional hints *) +Module tests_bidi. +Section tests_bidi. +Local Open Scope lang_scope. +Import Notations. +Context (R : realType). + +Definition v1 x : @exp R P [::] _ := [ + let x := return {1}:R in + return #x]. + +Definition v2 (a b : string) + (a := "a") (b := "b") + (* (H : infer (b != a)) *) + : @exp R P [::] _ := [ + let a := return {1}:R in + let b := return {true}:B in + (* let c := return {3}:R in + let d := return {4}:R in *) + return (#a, #b)]. + +Definition v3 (a b c d : string) (H1 : infer (b != a)) (H2 : infer (c != a)) + (H3 : infer (c != b)) (H4 : infer (a != b)) (H5 : infer (a != c)) + (H6 : infer (b != c)) : @exp R P [::] _ := [ + let a := return {1}:R in + let b := return {2}:R in + let c := return {3}:R in + (* let d := return {4}:R in *) + return (#b, #a)]. + +End tests_bidi. +End tests_bidi. + +Section trivial_example. +Local Open Scope lang_scope. +Import Notations. +Context {R : realType}. + +Lemma exec_normalize_return g x r : + projT1 (@execD _ g _ [Normalize return r:R]) x = + @dirac _ (measurableTypeR R) r _ :> probability _ R. + (* NB: \d_r notation? *) +Proof. +by rewrite execD_normalize execP_return execD_real//=; exact: normalize_kdirac. +Qed. + +End trivial_example. + +Section sample_pair. +Local Open Scope lang_scope. +Local Open Scope ring_scope. +Import Notations. +Context {R : realType}. + +Definition sample_pair_syntax0 : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in + let "y" := Sample {exp_bernoulli (1 / 3%:R)%:nng (p1S 2)} in + return (#{"x"}, #{"y"})]. + +Definition sample_pair_syntax : exp _ [::] _ := + [Normalize {sample_pair_syntax0}]. + +Lemma exec_sample_pair0 (A : set (bool * bool)) : + @execP R [::] _ sample_pair_syntax0 tt A = + ((1 / 2)%:E * + ((1 / 3)%:E * ((true, true) \in A)%:R%:E + + (1 - 1 / 3)%:E * ((true, false) \in A)%:R%:E) + + (1 - 1 / 2)%:E * + ((1 / 3)%:E * ((false, true) \in A)%:R%:E + + (1 - 1 / 3)%:E * ((false, false) \in A)%:R%:E))%E. +Proof. +rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. +rewrite execD_pair !exp_var'E (execD_var_erefl "x") (execD_var_erefl "y") /=. +rewrite letin'E ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !diracE !in_setT/= !mul1e. +rewrite !letin'E !ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. +by rewrite !integral_dirac//= !diracT !mul1e !diracE. +Qed. + +Lemma exec_sample_pair0_TandT : + @execP R [::] _ sample_pair_syntax0 tt [set (true, true)] = (1 / 6)%:E. +Proof. +rewrite exec_sample_pair0 mem_set//; do 3 rewrite memNset//=. +by rewrite /= !mule0 mule1 !add0e mule0 adde0; congr (_%:E); lra. +Qed. + +Lemma exec_sample_pair0_TandF : + @execP R [::] _ sample_pair_syntax0 tt [set (true, false)] = (1 / 3)%:E. +Proof. +rewrite exec_sample_pair0 memNset// mem_set//; do 2 rewrite memNset//. +by rewrite /= !mule0 mule1 !add0e mule0 adde0; congr (_%:E); lra. +Qed. + +Lemma exec_sample_pair0_TandT' : + @execP R [::] _ sample_pair_syntax0 tt [set p | p.1 && p.2] = (1 / 6)%:E. +Proof. +rewrite exec_sample_pair0 mem_set//; do 3 rewrite memNset//=. +by rewrite /= !mule0 mule1 !add0e mule0 adde0; congr (_%:E); lra. +Qed. + +Lemma exec_sample_pair_TorT : + (projT1 (execD sample_pair_syntax)) tt [set p | p.1 || p.2] = (2 / 3)%:E. +Proof. +rewrite execD_normalize normalizeE/= exec_sample_pair0. +do 4 rewrite mem_set//=. +rewrite eqe ifF; last by apply/negbTE/negP => /orP[/eqP|//]; lra. +rewrite exec_sample_pair0; do 3 rewrite mem_set//; rewrite memNset//=. +by rewrite !mule1; congr (_%:E); field. +Qed. + +End sample_pair. + +Section bernoulli_examples. +Local Open Scope ring_scope. +Local Open Scope lang_scope. +Import Notations. +Context {R : realType}. + +Definition bernoulli13_score := [Normalize + let "x" := Sample {@exp_bernoulli R [::] (1 / 3%:R)%:nng (p1S 2)} in + let "_" := if #{"x"} then Score {(1 / 3)}:R else Score {(2 / 3)}:R in + return #{"x"}]. + +Lemma exec_bernoulli13_score : + execD bernoulli13_score = execD (exp_bernoulli (1 / 5%:R)%:nng (p1S 4)). +Proof. +apply: eq_execD. +rewrite execD_bernoulli/= /bernoulli13_score execD_normalize 2!execP_letin. +rewrite execP_sample/= execD_bernoulli/= execP_if /= exp_var'E. +rewrite (execD_var_erefl "x")/= !execP_return/= 2!execP_score 2!execD_real/=. +apply: funext=> g; apply: eq_probability => U. +rewrite normalizeE !letin'E/=. +under eq_integral. + move=> x _. + rewrite !letin'E. + under eq_integral do rewrite retE /=. + over. +rewrite !ge0_integral_measure_add //=; last by move=> b _; rewrite integral_ge0. +rewrite !ge0_integral_mscale //=; last 2 first. + by move=> b _; rewrite integral_ge0. + by move=> b _; rewrite integral_ge0. +rewrite !integral_dirac// !diracT !mul1e. +rewrite iteE/= !ge0_integral_mscale//=. +rewrite ger0_norm//. +rewrite !integral_indic//= !iteE/= /mscale/=. +rewrite setTI !diracT !mule1. +rewrite ger0_norm//. +rewrite -EFinD/= eqe ifF; last first. + by apply/negbTE/negP => /orP[/eqP|//]; rewrite /onem; lra. +rewrite !letin'E/= !iteE/=. +rewrite !ge0_integral_mscale//=. +rewrite ger0_norm//. +rewrite !integral_dirac//= !diracT/= !mul1e ger0_norm//. +rewrite exp_var'E (execD_var_erefl "x")/=. +rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. +rewrite indicT/= !mulr1. +by rewrite muleDl//; congr (_ + _)%E; + rewrite -!EFinM; + congr (_%:E); + rewrite indicE /onem; case: (_ \in _); field. +Qed. + +Definition bernoulli12_score := [Normalize + let "x" := Sample {@exp_bernoulli R [::] (1 / 2)%:nng (p1S 1)} in + let "r" := if #{"x"} then Score {(1 / 3)}:R else Score {(2 / 3)}:R in + return #{"x"}]. + +Lemma exec_bernoulli12_score : + execD bernoulli12_score = execD (exp_bernoulli (1 / 3%:R)%:nng (p1S 2)). +Proof. +apply: eq_execD. +rewrite execD_bernoulli/= /bernoulli12_score execD_normalize 2!execP_letin. +rewrite execP_sample/= execD_bernoulli/= execP_if /= exp_var'E. +rewrite (execD_var_erefl "x")/= !execP_return/= 2!execP_score 2!execD_real/=. +apply: funext=> g; apply: eq_probability => U. +rewrite normalizeE !letin'E/=. +under eq_integral. + move=> x _. + rewrite !letin'E. + under eq_integral do rewrite retE /=. + over. +rewrite !ge0_integral_measure_add //=; last by move=> b _; rewrite integral_ge0. +rewrite !ge0_integral_mscale //=; last 2 first. + by move=> b _; rewrite integral_ge0. + by move=> b _; rewrite integral_ge0. +rewrite !integral_dirac// !diracT !mul1e. +rewrite iteE/= !ge0_integral_mscale//=. +rewrite ger0_norm//. +rewrite !integral_indic//= !iteE/= /mscale/=. +rewrite setTI diracE !in_setT !mule1. +rewrite ger0_norm//. +rewrite -EFinD/= eqe ifF; last first. + apply/negbTE/negP => /orP[/eqP|//]. + by rewrite /onem; lra. +rewrite !letin'E/= !iteE/=. +rewrite !ge0_integral_mscale//=. +rewrite ger0_norm//. +rewrite !integral_dirac//= !diracT /= !mul1e ger0_norm//. +rewrite !indicT/= mulr1. +rewrite exp_var'E (execD_var_erefl "x")/=. +rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. +by rewrite muleDl//; congr (_ + _)%E; + rewrite -!EFinM; + congr (_%:E); + rewrite indicE /onem; case: (_ \in _); field. +Qed. + +(* https://dl.acm.org/doi/pdf/10.1145/2933575.2935313 (Sect. 4) *) +Definition bernoulli14_score := [Normalize + let "x" := Sample {@exp_bernoulli R [::] (1 / 4%:R)%:nng (p1S 3)} in + let "r" := if #{"x"} then Score {5}:R else Score {2}:R in + return #{"x"}]. + +Let p511 : ((5%:R / 11%:R)%:nng%:num <= (1 : R)). +Proof. by rewrite /=; lra. Qed. + +Lemma exec_bernoulli14_score : + execD bernoulli14_score = execD (exp_bernoulli (5%:R / 11%:R)%:nng p511). +Proof. +apply: eq_execD. +rewrite execD_bernoulli/= execD_normalize 2!execP_letin. +rewrite execP_sample/= execD_bernoulli/= execP_if /= !exp_var'E. +rewrite !execP_return/= 2!execP_score 2!execD_real/=. +rewrite !(execD_var_erefl "x")/=. +apply: funext=> g; apply: eq_probability => U. +rewrite normalizeE !letin'E/=. +under eq_integral. + move=> x _. + rewrite !letin'E. + under eq_integral do rewrite retE /=. + over. +rewrite !ge0_integral_measure_add //=; last by move=> b _; rewrite integral_ge0. +rewrite !ge0_integral_mscale //=; last 2 first. + by move=> b _; rewrite integral_ge0. + by move=> b _; rewrite integral_ge0. +rewrite !integral_dirac// !diracT !mul1e. +rewrite iteE/= !ge0_integral_mscale//=. +rewrite ger0_norm//. +rewrite !integral_dirac//= !diracT !mule1. +rewrite iteE/= !ge0_integral_mscale//=. +rewrite ger0_norm//. +rewrite !integral_dirac//= !diracT !mule1. +rewrite -EFinD/= eqe ifF; last first. + apply/negbTE/negP => /orP[/eqP|//]. + by rewrite /onem; lra. +rewrite !letin'E/= !iteE/=. +rewrite !ge0_integral_mscale//=. +rewrite ger0_norm//. +rewrite !integral_dirac//= !diracT /= !mul1e ger0_norm//. +rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. +rewrite !indicT !mulr1. +by rewrite muleDl//; congr (_ + _)%E; + rewrite -!EFinM; + congr (_%:E); + by rewrite indicE /onem; case: (_ \in _); field. +Qed. + +End bernoulli_examples. + +Section hard_constraint'. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). + +Definition fail' := + letin' (score (@measurable_cst _ _ X _ setT (0%R : R))) + (ret (@measurable_cst _ _ _ Y setT point)). + +Lemma fail'E x U : fail' x U = 0. +Proof. by rewrite /fail' letin'E ge0_integral_mscale//= normr0 mul0e. Qed. + +End hard_constraint'. +Arguments fail' {d d' X Y R}. + +(* hard constraints to express score below 1 *) +Lemma score_fail' d (X : measurableType d) {R : realType} + (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : + score (kr r%:num) = + letin' (sample_cst (bernoulli r1) : R.-pker X ~> _) + (ite macc0of2 (ret ktt) fail'). +Proof. +apply/eq_sfkernel => x U. +rewrite letin'E/= /sample; unlock. +rewrite ge0_integral_measure_add//= ge0_integral_mscale//= ge0_integral_mscale//=. +rewrite integral_dirac//= integral_dirac//= !diracT/= !mul1e. +by rewrite /mscale/= iteE//= iteE//= fail'E mule0 adde0 ger0_norm. +Qed. + +Section hard_constraint. +Local Open Scope ring_scope. +Local Open Scope lang_scope. +Import Notations. +Context {R : realType} {str : string}. + +Definition hard_constraint g : @exp R _ g _ := + [let str := Score {0}:R in return TT]. + +Lemma exec_hard_constraint g mg U : + execP (hard_constraint g) mg U = fail' (false, tt) U. +Proof. +rewrite execP_letin execP_score execD_real execP_return execD_unit/=. +rewrite letin'E integral_indic//= /mscale/= normr0 mul0e. +by rewrite /fail' letin'E/= ge0_integral_mscale//= normr0 mul0e. +Qed. + +Lemma exec_score_fail (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : + execP (g := [::]) [Score {r%:num}:R] = + execP [let str := Sample {exp_bernoulli r r1} in + if #str then return TT else {hard_constraint _}]. +Proof. +rewrite execP_score execD_real /= score_fail'. +rewrite execP_letin execP_sample/= execD_bernoulli execP_if execP_return. +rewrite execD_unit/= exp_var'E /=. + apply/ctx_prf_head. +move=> h. +apply: eq_sfkernel=> /= -[] U. +rewrite 2!letin'E/=. +apply: eq_integral => b _. +rewrite 2!iteE//=. +case: b => //=. +- suff : projT1 (@execD R _ _ (exp_var str h)) (true, tt) = true by move=> ->. + set g := [:: (str, Bool)]. + have /= := @execD_var R [:: (str, Bool)] str. + by rewrite eqxx => /(_ h) ->. +- have -> : projT1 (@execD R _ _ (exp_var str h)) (false, tt) = false. + set g := [:: (str, Bool)]. + have /= := @execD_var R [:: (str, Bool)] str. + by rewrite eqxx /= => /(_ h) ->. + by rewrite (@exec_hard_constraint [:: (str, Bool)]). +Qed. + +End hard_constraint. + +Section letinA. +Local Open Scope lang_scope. +Variable R : realType. + +Lemma letinA g x y t1 t2 t3 (xyg : x \notin dom ((y, t2) :: g)) + (e1 : @exp R P g t1) + (e2 : exp P [:: (x, t1) & g] t2) + (e3 : exp P [:: (y, t2) & g] t3) : + forall U, measurable U -> + execP [let x := e1 in + let y := e2 in + {@exp_weak _ _ [:: (y, t2)] _ _ (x, t1) e3 xyg}] ^~ U = + execP [let y := + let x := e1 in e2 in + e3] ^~ U. +Proof. +move=> U mU; apply/funext=> z1. +rewrite !execP_letin. +rewrite (execP_weak [:: (y, t2)]). +apply: letin'A => //= z2 z3. +rewrite /kweak /mctx_strong /=. +by destruct z3. +Qed. + +Example letinA12 : forall U, measurable U -> + @execP R [::] _ [let "y" := return {1}:R in + let "x" := return {2}:R in + return #{"x"}] ^~ U = + @execP R [::] _ [let "x" := + let "y" := return {1}:R in return {2}:R in + return #{"x"}] ^~ U. +Proof. +move=> U mU. +rewrite !execP_letin !execP_return !execD_real. +apply: funext=> x. +rewrite !exp_var'E /= !(execD_var_erefl "x")/=. +exact: letin'A. +Qed. + +End letinA. + +Section staton_bus. +Local Open Scope ring_scope. +Local Open Scope lang_scope. +Import Notations. +Context {R : realType}. + +Definition staton_bus_syntax0 : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in + let "r" := if #{"x"} then return {3}:R else return {10}:R in + let "_" := Score {exp_poisson 4 [#{"r"}]} in + return #{"x"}]. + +Definition staton_bus_syntax := [Normalize {staton_bus_syntax0}]. + +Let sample_bern : R.-sfker munit ~> mbool := sample_cst (bernoulli p27). + +Let ite_3_10 : R.-sfker mbool * munit ~> measurableTypeR R := + ite macc0of2 (@ret _ _ _ (measurableTypeR R) R _ (kr 3)) (@ret _ _ _ (measurableTypeR R) R _ (kr 10)). + +Let score_poisson4 : R.-sfker measurableTypeR R * (mbool * munit) ~> munit := + score (measurableT_comp (measurable_poisson 4) (@macc0of2 _ _ (measurableTypeR R) _)). + +Let kstaton_bus' := + letin' sample_bern + (letin' ite_3_10 + (letin' score_poisson4 (ret macc2of4'))). + +Lemma eval_staton_bus0 : staton_bus_syntax0 -P> kstaton_bus'. +Proof. +apply: eval_letin; first by apply: eval_sample; exact: eval_bernoulli. +apply: eval_letin. + apply/evalP_if; [|exact/eval_return/eval_real..]. + rewrite exp_var'E. + by apply/execD_evalD; rewrite (execD_var_erefl "x")/=; congr existT. +apply: eval_letin. + apply/eval_score/eval_poisson. + rewrite exp_var'E. + by apply/execD_evalD; rewrite (execD_var_erefl "r")/=; congr existT. +apply/eval_return/execD_evalD. +by rewrite exp_var'E (execD_var_erefl "x")/=; congr existT. +Qed. + +Lemma exec_staton_bus0' : execP staton_bus_syntax0 = kstaton_bus'. +Proof. +rewrite 3!execP_letin execP_sample/= execD_bernoulli. +rewrite /kstaton_bus'; congr letin'. +rewrite !execP_if !execP_return !execD_real/=. +rewrite exp_var'E (execD_var_erefl "x")/=. +have -> : measurable_acc_typ [:: Bool] 0 = macc0of2 by []. +congr letin'. +rewrite execP_score execD_poisson/=. +rewrite exp_var'E (execD_var_erefl "r")/=. +have -> : measurable_acc_typ [:: Real; Bool] 0 = macc0of2 by []. +congr letin'. +by rewrite exp_var'E (execD_var_erefl "x") /=; congr ret. +Qed. + +Lemma exec_staton_bus : execD staton_bus_syntax = + existT _ (normalize kstaton_bus' point) (measurable_fun_mnormalize _). +Proof. by rewrite execD_normalize exec_staton_bus0'. Qed. + +Let poisson4 := @poisson R 4%N. + +Let staton_bus_probability U := + ((2 / 7)%:E * (poisson4 3)%:E * \d_true U + + (5 / 7)%:E * (poisson4 10)%:E * \d_false U)%E. + +Lemma exec_staton_bus0 (U : set bool) : + execP staton_bus_syntax0 tt U = staton_bus_probability U. +Proof. +rewrite exec_staton_bus0' /staton_bus_probability /kstaton_bus'. +rewrite letin'_sample_bernoulli. +rewrite -!muleA; congr (_ * _ + _ * _)%E. +- rewrite letin'_iteT//. + rewrite letin'_retk//. + rewrite letin'_kret//. + rewrite /score_poisson4. + by rewrite /score/= /mscale/= ger0_norm//= poisson_ge0. +- by rewrite onem27. +- rewrite letin'_iteF//. + rewrite letin'_retk//. + rewrite letin'_kret//. + rewrite /score_poisson4. + by rewrite /score/= /mscale/= ger0_norm//= poisson_ge0. +Qed. + +End staton_bus. + +(* same as staton_bus module associativity of letin *) +Section staton_busA. +Local Open Scope ring_scope. +Local Open Scope lang_scope. +Import Notations. +Context {R : realType}. + +Definition staton_busA_syntax0 : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in + let "_" := + let "r" := if #{"x"} then return {3}:R else return {10}:R in + Score {exp_poisson 4 [#{"r"}]} in + return #{"x"}]. + +Definition staton_busA_syntax : exp _ [::] _ := + [Normalize {staton_busA_syntax0}]. + +Let sample_bern : R.-sfker munit ~> mbool := sample_cst (bernoulli p27). + +Let ite_3_10 : R.-sfker mbool * munit ~> measurableTypeR R := + ite macc0of2 (@ret _ _ _ (measurableTypeR R) R _ (kr 3)) (@ret _ _ _ (measurableTypeR R) R _ (kr 10)). + +Let score_poisson4 : R.-sfker measurableTypeR R * (mbool * munit) ~> munit := + score (measurableT_comp (measurable_poisson 4) (@macc0of3' _ _ _ (measurableTypeR R) _ _)). + +(* same as kstaton_bus _ (measurable_poisson 4) but expressed with letin' + instead of letin *) +Let kstaton_busA' := + letin' sample_bern + (letin' + (letin' ite_3_10 + score_poisson4) + (ret macc1of3')). + +(*Lemma kstaton_busA'E : kstaton_busA' = kstaton_bus _ (measurable_poisson 4). +Proof. +apply/eq_sfkernel => -[] U. +rewrite /kstaton_busA' /kstaton_bus. +rewrite letin'_letin. +rewrite /sample_bern. +congr (letin _ _ tt U). +rewrite 2!letin'_letin/=. +Abort.*) + +Lemma eval_staton_busA0 : staton_busA_syntax0 -P> kstaton_busA'. +Proof. +apply: eval_letin; first by apply: eval_sample; exact: eval_bernoulli. +apply: eval_letin. + apply: eval_letin. + apply/evalP_if; [|exact/eval_return/eval_real..]. + rewrite exp_var'E. + by apply/execD_evalD; rewrite (execD_var_erefl "x")/=; congr existT. + apply/eval_score/eval_poisson. + rewrite exp_var'E. + by apply/execD_evalD; rewrite (execD_var_erefl "r")/=; congr existT. +apply/eval_return. +by apply/execD_evalD; rewrite exp_var'E (execD_var_erefl "x")/=; congr existT. +Qed. + +Lemma exec_staton_busA0' : execP staton_busA_syntax0 = kstaton_busA'. +Proof. +rewrite 3!execP_letin execP_sample/= execD_bernoulli. +rewrite /kstaton_busA'; congr letin'. +rewrite !execP_if !execP_return !execD_real/=. +rewrite exp_var'E (execD_var_erefl "x")/=. +have -> : measurable_acc_typ [:: Bool] 0 = macc0of2 by []. +congr letin'. + rewrite execP_score execD_poisson/=. + rewrite exp_var'E (execD_var_erefl "r")/=. + by have -> : measurable_acc_typ [:: Real; Bool] 0 = macc0of3' by []. +by rewrite exp_var'E (execD_var_erefl "x") /=; congr ret. +Qed. + +Lemma exec_statonA_bus : execD staton_busA_syntax = + existT _ (normalize kstaton_busA' point) (measurable_fun_mnormalize _). +Proof. by rewrite execD_normalize exec_staton_busA0'. Qed. + +(* equivalence between staton_bus and staton_busA *) +Lemma staton_bus_staton_busA : + execP staton_bus_syntax0 = @execP R _ _ staton_busA_syntax0. +Proof. +rewrite /staton_bus_syntax0 /staton_busA_syntax0. +rewrite execP_letin. +rewrite [in RHS]execP_letin. +congr (letin' _). +set e1 := exp_if _ _ _. +set e2 := exp_score _. +set e3 := (exp_return _ in RHS). +pose f := @found _ Unit "x" Bool [::]. +have r_f : "r" \notin [seq i.1 | i <- ("_", Unit) :: untag (ctx_of f)] by []. +have H := @letinA _ _ _ _ _ _ + (lookup Unit (("_", Unit) :: untag (ctx_of f)) "x") + r_f e1 e2 e3. +apply/eq_sfkernel => /= x U. +have mU : + (@mtyp_disp R (lookup Unit (("_", Unit) :: untag (ctx_of f)) "x")).-measurable U. + by []. +move: H => /(_ U mU) /(congr1 (fun f => f x)) <-. +set e3' := exp_return _. +set e3_weak := exp_weak _ _ _ _. +rewrite !execP_letin. +suff: execP e3' = execP (e3_weak e3 r_f) by move=> <-. +rewrite execP_return/= exp_var'E (execD_var_erefl "x") /= /e3_weak. +rewrite (@execP_weak R [:: ("_", Unit)] _ ("r", Real) _ e3 r_f). +rewrite execP_return exp_var'E/= (execD_var_erefl "x") //=. +by apply/eq_sfkernel => /= -[[] [a [b []]]] U0. +Qed. + +Let poisson4 := @poisson R 4%N. + +Lemma exec_staton_busA0 U : execP staton_busA_syntax0 tt U = + ((2 / 7%:R)%:E * (poisson4 3%:R)%:E * \d_true U + + (5%:R / 7%:R)%:E * (poisson4 10%:R)%:E * \d_false U)%E. +Proof. by rewrite -staton_bus_staton_busA exec_staton_bus0. Qed. + +End staton_busA. + +Section letinC. +Local Open Scope lang_scope. +Variable (R : realType). + +Require Import Classical_Prop. + +Lemma letinC g t1 t2 (e1 : @exp R P g t1) (e2 : exp P g t2) + (x y : string) + (xy : infer (x != y)) (yx : infer (y != x)) + (xg : x \notin dom g) (yg : y \notin dom g) : + forall U, measurable U -> + execP [ + let x := e1 in + let y := {exp_weak _ [::] _ (x, t1) e2 xg} in + return (#x, #y)] ^~ U = + execP [ + let y := e2 in + let x := {exp_weak _ [::] _ (y, t2) e1 yg} in + return (#x, #y)] ^~ U. +Proof. +move=> U mU; apply/funext => z. +rewrite 4!execP_letin. +rewrite 2!(execP_weak [::] g). +rewrite 2!execP_return/=. +rewrite 2!execD_pair/=. +rewrite !exp_var'E. +- exact/(ctx_prf_tail _ yx)/ctx_prf_head. +- exact/ctx_prf_head. +- exact/ctx_prf_head. +- exact/(ctx_prf_tail _ xy)/ctx_prf_head. +- move=> h1 h2 h3 h4. + set g1 := [:: (y, t2), (x, t1) & g]. + set g2 := [:: (x, t1), (y, t2) & g]. + have /= := @execD_var R g1 x. + rewrite (negbTE yx) eqxx => /(_ h4) ->. + have /= := @execD_var R g2 x. + rewrite (negbTE yx) eqxx => /(_ h2) ->. + have /= := @execD_var R g1 y. + rewrite eqxx => /(_ h3) ->. + have /= := @execD_var R g2 y. + rewrite (negbTE xy) eqxx => /(_ h1) -> /=. + have -> : measurable_acc_typ [:: t2, t1 & map snd g] 0 = macc0of3' by []. + have -> : measurable_acc_typ [:: t2, t1 & map snd g] 1 = macc1of3' by []. + rewrite (letin'C _ _ (execP e2) + [the R.-sfker _ ~> _ of @kweak _ [::] _ (y, t2) _ (execP e1)]); + [ |by [] | by [] |by []]. + have -> : measurable_acc_typ [:: t1, t2 & map snd g] 0 = macc0of3' by []. + by have -> : measurable_acc_typ [:: t1, t2 & map snd g] 1 = macc1of3' by []. +Qed. + +Example letinC_ground_variables g t1 t2 (e1 : @exp R P g t1) (e2 : exp P g t2) + (x := "x") (y := "y") + (xg : x \notin dom g) (yg : y \notin dom g) : + forall U, measurable U -> + execP [ + let x := e1 in + let y := {exp_weak _ [::] _ (x, t1) e2 xg} in + return (#x, #y)] ^~ U = + execP [ + let y := e2 in + let x := {exp_weak _ [::] _ (y, t2) e1 yg} in + return (#x, #y)] ^~ U. +Proof. by move=> U mU; rewrite letinC. Qed. + +Example letinC_ground (g := [:: ("a", Unit); ("b", Bool)]) t1 t2 + (e1 : @exp R P g t1) + (e2 : exp P g t2) : + forall U, measurable U -> + execP [let "x" := e1 in + let "y" := e2 :+ {"x"} in + return (#{"x"}, #{"y"})] ^~ U = + execP [let "y" := e2 in + let "x" := e1 :+ {"y"} in + return (#{"x"}, #{"y"})] ^~ U. +Proof. move=> U mU; exact: letinC. Qed. + +End letinC. diff --git a/theories/lang_syntax_toy.v b/theories/lang_syntax_toy.v new file mode 100644 index 0000000000..e96777393f --- /dev/null +++ b/theories/lang_syntax_toy.v @@ -0,0 +1,551 @@ +Require Import String Classical. +From HB Require Import structures. +From mathcomp Require Import all_ssreflect ssralg interval_inference. +From mathcomp Require Import mathcomp_extra boolp. +From mathcomp Require Import reals topology normedtype. +From mathcomp Require Import lang_syntax_util. + +(******************************************************************************) +(* Intrinsically-typed concrete syntax for a toy language *) +(* *) +(* The main module provided by this file is "lang_intrinsic_tysc" which *) +(* provides an example of intrinsically-typed concrete syntax for a toy *) +(* language (a simplification of the syntax/evaluation formalized in *) +(* lang_syntax.v). Other modules provide even more simplified language for *) +(* pedagogical purposes. *) +(* *) +(* lang_extrinsic == non-intrinsic definition of expression *) +(* lang_intrinsic_ty == intrinsically-typed syntax *) +(* lang_intrinsic_sc == intrinsically-scoped syntax *) +(* lang_intrinsic_tysc == intrinsically-typed/scoped syntax *) +(* *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Set Printing Implicit Defensive. + +Import numFieldTopology.Exports. + +Local Open Scope classical_set_scope. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. +Local Open Scope string_scope. + +Section type. +Variables (R : realType). + +Inductive typ := Real | Unit. + +HB.instance Definition _ := gen_eqMixin typ. + +Definition iter_pair (l : list Type) : Type := + foldr (fun x y => (x * y)%type) unit l. + +Definition Type_of_typ (t : typ) : Type := + match t with + | Real => R + | Unit => unit + end. + +Definition ctx := seq (string * typ). + +Definition Type_of_ctx (g : ctx) := iter_pair (map (Type_of_typ \o snd) g). + +Goal Type_of_ctx [:: ("x", Real); ("y", Real)] = (R * (R * unit))%type. +Proof. by []. Qed. + +End type. + +Module lang_extrinsic. +Section lang_extrinsic. +Variable R : realType. +Implicit Types str : string. + +Inductive exp : Type := +| exp_unit : exp +| exp_real : R -> exp +| exp_var (g : ctx) t str : t = lookup Unit g str -> exp +| exp_add : exp -> exp -> exp +| exp_letin str : exp -> exp -> exp. +Arguments exp_var {g t}. + +Fail Example letin_once : exp := + exp_letin "x" (exp_real 1) (exp_var "x" erefl). +Example letin_once : exp := + exp_letin "x" (exp_real 1) (@exp_var [:: ("x", Real)] Real "x" erefl). + +End lang_extrinsic. +End lang_extrinsic. + +Module lang_intrinsic_ty. +Section lang_intrinsic_ty. +Variable R : realType. +Implicit Types str : string. + +Inductive exp : typ -> Type := +| exp_unit : exp Unit +| exp_real : R -> exp Real +| exp_var g t str : t = lookup Unit g str -> exp t +| exp_add : exp Real -> exp Real -> exp Real +| exp_letin t u : string -> exp t -> exp u -> exp u. +Arguments exp_var {g t}. + +Fail Example letin_once : exp Real := + exp_letin "x" (exp_real 1) (exp_var "x" erefl). +Example letin_once : exp Real := + exp_letin "x" (exp_real 1) (@exp_var [:: ("x", Real)] _ "x" erefl). + +End lang_intrinsic_ty. +End lang_intrinsic_ty. + +Module lang_intrinsic_sc. +Section lang_intrinsic_sc. +Variable R : realType. +Implicit Types str : string. + +Inductive exp : ctx -> Type := +| exp_unit g : exp g +| exp_real g : R -> exp g +| exp_var g t str : t = lookup Unit g str -> exp g +| exp_add g : exp g -> exp g -> exp g +| exp_letin g t str : exp g -> exp ((str, t) :: g) -> exp g. +Arguments exp_real {g}. +Arguments exp_var {g t}. +Arguments exp_letin {g t}. + +Declare Custom Entry expr. + +Notation "[ e ]" := e (e custom expr at level 5). +Notation "{ x }" := x (in custom expr, x constr). +Notation "x ':R'" := (exp_real x) (in custom expr at level 1). +Notation "x" := x (in custom expr at level 0, x ident). +Notation "$ x" := (exp_var x erefl) (in custom expr at level 1). +Notation "x + y" := (exp_add x y) + (in custom expr at level 2, left associativity). +Notation "'let' x ':=' e1 'in' e2" := (exp_letin x e1 e2) + (in custom expr at level 3, x constr, + e1 custom expr at level 2, e2 custom expr at level 3, + left associativity). + +Fail Example letin_once : exp [::] := + [let "x" := {1%R}:R in ${"x"}]. +Example letin_once : exp [::] := + [let "x" := {1%R}:R in {@exp_var [:: ("x", Real)] _ "x" erefl}]. + +Fixpoint acc (g : ctx) (i : nat) : + Type_of_ctx R g -> @Type_of_typ R (nth Unit (map snd g) i) := + match g return Type_of_ctx R g -> Type_of_typ R (nth Unit (map snd g) i) with + | [::] => match i with | O => id | j.+1 => id end + | _ :: _ => match i with + | O => fst + | j.+1 => fun H => acc j H.2 + end + end. +Arguments acc : clear implicits. + +Inductive eval : forall g (t : typ), exp g -> (Type_of_ctx R g -> Type_of_typ R t) -> Prop := +| eval_real g c : @eval g Real [c:R] (fun=> c) +| eval_plus g (e1 e2 : exp g) (v1 v2 : R) : + @eval g Real e1 (fun=> v1) -> + @eval g Real e2 (fun=> v2) -> + @eval g Real [e1 + e2] (fun=> v1 + v2) +| eval_var (g : ctx) str i : + i = index str (map fst g) -> eval [$ str] (acc g i). + +Goal @eval [::] Real [{1}:R] (fun=> 1). +Proof. exact: eval_real. Qed. +Goal @eval [::] Real [{1}:R + {2}:R] (fun=> 3). +Proof. exact/eval_plus/eval_real/eval_real. Qed. +Goal @eval [:: ("x", Real)] _ [$ {"x"}] (acc [:: ("x", Real)] 0). +Proof. exact: eval_var. Qed. + +End lang_intrinsic_sc. +End lang_intrinsic_sc. + +Module lang_intrinsic_tysc. +Section lang_intrinsic_tysc. +Variable R : realType. +Implicit Types str : string. + +Inductive typ := Real | Unit | Pair : typ -> typ -> typ. + +HB.instance Definition _ := gen_eqMixin typ. + +Fixpoint mtyp (t : typ) : Type := + match t with + | Real => R + | Unit => unit + | Pair t1 t2 => (mtyp t1 * mtyp t2) + end. + +Definition ctx := seq (string * typ). + +Definition Type_of_ctx (g : ctx) := iter_pair (map (mtyp \o snd) g). + +Goal Type_of_ctx [:: ("x", Real); ("y", Real)] = (R * (R * unit))%type. +Proof. by []. Qed. + +Inductive exp : ctx -> typ -> Type := +| exp_unit g : exp g Unit +| exp_real g : R -> exp g Real +| exp_var g t str : t = lookup Unit g str -> exp g t +| exp_add g : exp g Real -> exp g Real -> exp g Real +| exp_pair g t1 t2 : exp g t1 -> exp g t2 -> exp g (Pair t1 t2) +| exp_letin g t1 t2 x : exp g t1 -> exp ((x, t1) :: g) t2 -> exp g t2. + +Definition exp_var' str (t : typ) (g : find str t) := + @exp_var (untag (ctx_of g)) t str (ctx_prf g). + +Section no_bidirectional_hints. + +Arguments exp_unit {g}. +Arguments exp_real {g}. +Arguments exp_var {g t}. +Arguments exp_add {g}. +Arguments exp_pair {g t1 t2}. +Arguments exp_letin {g t1 t2}. +Arguments exp_var' str {t} g. + +Fail Example letin_add : exp [::] _ := + exp_letin "x" (exp_real 1) + (exp_letin "y" (exp_real 2) + (exp_add (exp_var "x" erefl) + (exp_var "y" erefl))). +Example letin_add : exp [::] _ := + exp_letin "x" (exp_real 1) + (exp_letin "y" (exp_real 2) + (exp_add (@exp_var [:: ("y", Real); ("x", Real)] _ "x" erefl) + (exp_var "y" erefl))). +Reset letin_add. + +Declare Custom Entry expr. + +Notation "[ e ]" := e (e custom expr at level 5). +Notation "{ x }" := x (in custom expr, x constr). +Notation "x ':R'" := (exp_real x) (in custom expr at level 1). +Notation "x" := x (in custom expr at level 0, x ident). +Notation "$ x" := (exp_var x erefl) (in custom expr at level 1). +Notation "# x" := (exp_var' x%string _) (in custom expr at level 1). +Notation "e1 + e2" := (exp_add e1 e2) + (in custom expr at level 2, + (* e1 custom expr at level 1, e2 custom expr at level 2, *) + left associativity). +Notation "( e1 , e2 )" := (exp_pair e1 e2) + (in custom expr at level 1). +Notation "'let' x ':=' e1 'in' e2" := (exp_letin x e1 e2) + (in custom expr at level 3, x constr, + e1 custom expr at level 2, e2 custom expr at level 3, + left associativity). + +Fail Definition let3_add_erefl (a b c : string) + (ba : infer (b != a)) (ca : infer (c != a)) (cb : infer (c != b)) + (ab : infer (a != b)) (ac : infer (a != c)) (bc : infer (b != c)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + let c := {3}:R in + $a + $b]. +(* The term "[$ a]" has type "exp ?g2 (lookup Unit ?g2 a)" while it is expected to have type "exp ?g2 Real". *) + +Definition let3_pair_erefl (a b c : string) + (ba : infer (b != a)) (ca : infer (c != a)) (cb : infer (c != b)) + (ab : infer (a != b)) (ac : infer (a != c)) (bc : infer (b != c)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + let c := {3}:R in + ($a, $b)]. + +Fail Definition let3_add (a b c : string) + (ba : infer (b != a)) (ca : infer (c != a)) (cb : infer (c != b)) + (ab : infer (a != b)) (ac : infer (a != c)) (bc : infer (b != c)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + let c := {3}:R in + #a + #b]. +(* The term "[# a + # b]" has type + "exp (untag (ctx_of (recurse (str':=b) Real ?f))) Real" +while it is expected to have type "exp ((c, Real) :: ?g1) ?u1" +(cannot unify "(b, Real)" and "(c, Real)"). *) + +Fail Definition let3_pair (a b c : string) + (ba : infer (b != a)) (ca : infer (c != a)) (cb : infer (c != b)) + (ab : infer (a != b)) (ac : infer (a != c)) (bc : infer (b != c)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + let c := {3}:R in + (#a, #b)]. +(* The term "[# a + # b]" has type "exp (untag (ctx_of (recurse (str':=b) Real ?f))) Real" while it is expected to have type + "exp ((c, Real) :: ?g1) ?u1" (cannot unify "(b, Real)" and "(c, Real)"). *) + +End no_bidirectional_hints. + +Section with_bidirectional_hints. + +Arguments exp_unit {g}. +Arguments exp_real {g}. +Arguments exp_var {g t}. +Arguments exp_add {g} &. +Arguments exp_pair {g} & {t1 t2}. +Arguments exp_letin {g} & {t1 t2}. +Arguments exp_var' str {t} g. + +Declare Custom Entry expr. + +Notation "[ e ]" := e (e custom expr at level 5). +Notation "{ x }" := x (in custom expr, x constr). +Notation "x ':R'" := (exp_real x) (in custom expr at level 1). +Notation "x" := x (in custom expr at level 0, x ident). +Notation "$ x" := (exp_var x%string erefl) (in custom expr at level 1). +Notation "# x" := (exp_var' x%string _) (in custom expr at level 1). +Notation "e1 + e2" := (exp_add e1 e2) + (in custom expr at level 2, + left associativity). +Notation "( e1 , e2 )" := (exp_pair e1 e2) + (in custom expr at level 1). +Notation "'let' x ':=' e1 'in' e2" := (exp_letin x e1 e2) + (in custom expr at level 3, x constr, + e1 custom expr at level 2, e2 custom expr at level 3, + left associativity). + +Fail Definition let2_add_erefl_bidi (a b : string) + (ba : infer (b != a)) (ab : infer (a != b)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + $a + $b]. + +Definition let2_add_erefl_bidi (a b : string) + (ba : infer (b != a)) (ab : infer (a != b)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + #a + #b]. + +Fail Definition let3_add_erefl_bidi (a b c d : string) + (ba : infer (b != a)) (ca : infer (c != a)) (cb : infer (c != b)) + (ab : infer (a != b)) (ac : infer (a != c)) (bc : infer (b != c)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + let c := {3}:R in + $a + $b]. +(* The term "[$ a]" has type "exp [:: (c, Real); (b, Real); (a, Real)] (lookup Unit [:: (c, Real); (b, Real); (a, Real)] a)" +while it is expected to have type "exp [:: (c, Real); (b, Real); (a, Real)] Real" +(cannot unify "lookup Unit [:: (c, Real); (b, Real); (a, Real)] a" and "Real"). *) + +Definition let3_pair_erefl_bidi (a b c : string) + (ba : infer (b != a)) (ca : infer (c != a)) (cb : infer (c != b)) + (ab : infer (a != b)) (ac : infer (a != c)) (bc : infer (b != c)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + let c := {3}:R in + ($a, $b)]. + +Check let3_pair_erefl_bidi. +(* exp [::] (Pair (lookup Unit [:: (c, Real); (b, Real); (a, Real)] a) (lookup Unit [:: (c, Real); (b, Real); (a, Real)] b)) *) + +Definition let3_add_bidi (a b c : string) + (ba : infer (b != a)) (ca : infer (c != a)) (cb : infer (c != b)) + (ab : infer (a != b)) (ac : infer (a != c)) (bc : infer (b != c)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + let c := {3}:R in + #a + #b]. + +Definition let3_pair_bidi (a b c : string) + (ba : infer (b != a)) (ca : infer (c != a)) (cb : infer (c != b)) + (ab : infer (a != b)) (ac : infer (a != c)) (bc : infer (b != c)) + : exp [::] _ := [ + let a := {1}:R in + let b := {2}:R in + let c := {3}:R in + (#a , #b)]. + +Check let3_pair_bidi. +(* exp [::] (Pair Real Real) *) + +Example e0 : exp [::] _ := exp_real 1. +Example letin1 : exp [::] _ := + exp_letin "x" (exp_real 1) (exp_var "x" erefl). +Example letin2 : exp [::] _ := + exp_letin "x" (exp_real 1) + (exp_letin "y" (exp_real 2) + (exp_var "x" erefl)). + +Example letin_add : exp [::] _ := + exp_letin "x" (exp_real 1) + (exp_letin "y" (exp_real 2) + (exp_add (exp_var "x" erefl) + (exp_var "y" erefl))). +Reset letin_add. +Fail Example letin_add (x y : string) + (xy : infer (x != y)) (yx : infer (y != x)) : exp [::] _ := + exp_letin x (exp_real 1) + (exp_letin y (exp_real 2) + (exp_add (exp_var x erefl) (exp_var y erefl))). +Example letin_add (x y : string) + (xy : infer (x != y)) (yx : infer (y != x)) : exp [::] _ := + exp_letin x (exp_real 1) + (exp_letin y (exp_real 2) + (exp_add (exp_var' x _) (exp_var' y _))). +Reset letin_add. + +Example letin_add_custom : exp [::] _ := + [let "x" := {1}:R in + let "y" := {2}:R in + #{"x"} + #{"y"}]. + +Section eval. + +Fixpoint acc (g : ctx) (i : nat) : + Type_of_ctx g -> mtyp (nth Unit (map snd g) i) := + match g return Type_of_ctx g -> mtyp (nth Unit (map snd g) i) with + | [::] => match i with | O => id | j.+1 => id end + | _ :: _ => match i with + | O => fst + | j.+1 => fun H => acc j H.2 + end + end. +Arguments acc : clear implicits. + +Reserved Notation "e '-e->' v" (at level 40). + +Inductive eval : forall g t, exp g t -> (Type_of_ctx g -> mtyp t) -> Prop := +| eval_tt g : (exp_unit : exp g _) -e-> (fun=> tt) +| eval_real g c : (exp_real c : exp g _) -e-> (fun=> c) +| eval_plus g (e1 e2 : exp g Real) v1 v2 : + e1 -e-> v1 -> + e2 -e-> v2 -> + [e1 + e2] -e-> fun x => v1 x + v2 x +| eval_var g str : + let i := index str (map fst g) in + exp_var str erefl -e-> acc g i +| eval_pair g t1 t2 e1 e2 v1 v2 : + e1 -e-> v1 -> + e2 -e-> v2 -> + @exp_pair g t1 t2 e1 e2 -e-> fun x => (v1 x, v2 x) +| eval_letin g t t' str (e1 : exp g t) (e2 : exp ((str, t) :: g) t') v1 v2 : + e1 -e-> v1 -> + e2 -e-> v2 -> + exp_letin str e1 e2 -e-> (fun a => v2 (v1 a, a)) +where "e '-e->' v" := (@eval _ _ e v). + +Lemma eval_uniq g t (e : exp g t) u v : + e -e-> u -> e -e-> v -> u = v. +Proof. +move=> hu. +apply: (@eval_ind + (fun g t (e : exp g t) (u : Type_of_ctx g -> mtyp t) => + forall v, e -e-> v -> u = v)); last exact: hu. +all: (rewrite {g t e u v hu}). +- move=> g v. + inversion 1. + by inj_ex H3. +- move=> g c v. + inversion 1. + by inj_ex H3. +- move=> g e1 e2 v1 v2 ev1 IH1 ev2 IH2 v. + inversion 1. + inj_ex H0; inj_ex H1; subst. + inj_ex H5; subst. + by rewrite (IH1 _ H3) (IH2 _ H4). +- move=> g x i v. + inversion 1. + by inj_ex H6; subst. +- move=> g t1 t2 e1 e2 v1 v2 ev1 IH1 ev2 IH2 v. + inversion 1. + inj_ex H3; inj_ex H4; subst. + inj_ex H5; subst. + by rewrite (IH1 _ H6) (IH2 _ H7). +- move=> g t t' x0 e0 e1 v1 v2 ev1 IH1 ev2 IH2 v. + inversion 1. + inj_ex H5; subst. + inj_ex H6; subst. + inj_ex H7; subst. + by rewrite (IH1 _ H4) (IH2 _ H8). +Qed. + +Lemma eval_total g t (e : exp g t) : exists v, e -e-> v. +Proof. +elim: e. +- by eexists; exact: eval_tt. +- by eexists; exact: eval_real. +- move=> {}g {}t x e; subst t. + by eexists; exact: eval_var. +- move=> {}g e1 [v1] IH1 e2 [v2] IH2. + by eexists; exact: (eval_plus IH1 IH2). +- move=> {}g t1 t2 e1 [v1] IH1 e2 [v2] IH2. + by eexists; exact: (eval_pair IH1 IH2). +- move=> {}g {}t u x e1 [v1] IH1 e2 [v2] IH2. + by eexists; exact: (eval_letin IH1 IH2). +Qed. + +Definition exec g t (e : exp g t) : Type_of_ctx g -> mtyp t := + proj1_sig (cid (@eval_total g t e)). + +Lemma exec_eval g t (e : exp g t) v : exec e = v <-> e -e-> v. +Proof. +split. + by move=> <-; rewrite /exec; case: cid. +move=> ev. +by rewrite /exec; case: cid => f H/=; apply: eval_uniq; eauto. +Qed. + +Lemma eval_exec g t (e : exp g t) : e -e-> exec e. +Proof. by rewrite /exec; case: cid. Qed. + +Lemma exec_real g r : @exec g Real (exp_real r) = (fun=> r). +Proof. exact/exec_eval/eval_real. Qed. + +Lemma exec_var g str t H : + exec (@exp_var _ t str H) = + eq_rect _ (fun a => Type_of_ctx g -> mtyp a) + (acc g (index str (map fst g))) + _ (esym H). +Proof. +subst t. +rewrite {1}/exec. +case: cid => f H. +inversion H; subst g0 str0. +by inj_ex H6; subst f. +Qed. + +Lemma exp_var'E str t (f : find str t) H : exp_var' str f = exp_var str H. +Proof. by rewrite /exp_var'; congr exp_var. Qed. + +Lemma exec_letin g x t1 t2 (e1 : exp g t1) (e2 : exp ((x, t1) :: g) t2) : + exec [let x := e1 in e2] = (fun a => (exec e2) ((exec e1) a, a)). +Proof. by apply/exec_eval/eval_letin; exact: eval_exec. Qed. + +Goal ([{1}:R] : exp [::] _) -e-> (fun=> 1). +Proof. exact: eval_real. Qed. +Goal @eval [::] _ [{1}:R + {2}:R] (fun=> 3). +Proof. exact/eval_plus/eval_real/eval_real. Qed. +Goal @eval [:: ("x", Real)] _ (exp_var "x" erefl) (@acc [:: ("x", Real)] 0). +Proof. exact: eval_var. Qed. +Goal @eval [::] _ [let "x" := {1}:R in #{"x"}] (fun=> 1). +Proof. +apply/exec_eval; rewrite exec_letin/=. +apply/funext => t/=. +by rewrite exp_var'E exec_real/= exec_var/=. +Qed. + +Goal exec (g := [::]) [let "x" := {1}:R in #{"x"}] = (fun=> 1). +Proof. +rewrite exec_letin//=. +apply/funext => x. +by rewrite exp_var'E exec_var/= exec_real. +Qed. + +End eval. + +End with_bidirectional_hints. + +End lang_intrinsic_tysc. +End lang_intrinsic_tysc. diff --git a/theories/lang_syntax_util.v b/theories/lang_syntax_util.v new file mode 100644 index 0000000000..91096193f0 --- /dev/null +++ b/theories/lang_syntax_util.v @@ -0,0 +1,83 @@ +Require Import String. +From HB Require Import structures. +Require Import Classical_Prop. (* NB: to compile with Coq 8.17 *) +From mathcomp Require Import all_ssreflect. +From mathcomp Require Import interval_inference. + +(******************************************************************************) +(* Shared by lang_syntax_*.v files *) +(******************************************************************************) + +HB.instance Definition _ := hasDecEq.Build string eqb_spec. + +Ltac inj_ex H := revert H; + match goal with + | |- existT ?P ?l (existT ?Q ?t (existT ?R ?u (existT ?T ?v ?v1))) = + existT ?P ?l (existT ?Q ?t (existT ?R ?u (existT ?T ?v ?v2))) -> _ => + (intro H; do 4 apply Classical_Prop.EqdepTheory.inj_pair2 in H) + | |- existT ?P ?l (existT ?Q ?t (existT ?R ?u ?v1)) = + existT ?P ?l (existT ?Q ?t (existT ?R ?u ?v2)) -> _ => + (intro H; do 3 apply Classical_Prop.EqdepTheory.inj_pair2 in H) + | |- existT ?P ?l (existT ?Q ?t ?v1) = + existT ?P ?l (existT ?Q ?t ?v2) -> _ => + (intro H; do 2 apply Classical_Prop.EqdepTheory.inj_pair2 in H) + | |- existT ?P ?l (existT ?Q ?t ?v1) = + existT ?P ?l (existT ?Q ?t' ?v2) -> _ => + (intro H; do 2 apply Classical_Prop.EqdepTheory.inj_pair2 in H) + | |- existT ?P ?l ?v1 = + existT ?P ?l ?v2 -> _ => + (intro H; apply Classical_Prop.EqdepTheory.inj_pair2 in H) + | |- existT ?P ?l ?v1 = + existT ?P ?l' ?v2 -> _ => + (intro H; apply Classical_Prop.EqdepTheory.inj_pair2 in H) +end. + +Set Implicit Arguments. +Unset Strict Implicit. +Set Printing Implicit Defensive. + +Class infer (P : Prop) := Infer : P. +#[global] Hint Mode infer ! : typeclass_instances. +#[global] Hint Extern 0 (infer _) => (exact) : typeclass_instances. + +Section tagged_context. +Context {T : eqType} {t0 : T}. +Let ctx := seq (string * T). +Implicit Types (str : string) (g : ctx) (t : T). + +Definition dom g := map fst g. + +Definition lookup g str := nth t0 (map snd g) (index str (dom g)). + +Structure tagged_ctx := Tag {untag : ctx}. + +Structure find str t := Find { + ctx_of : tagged_ctx ; + #[canonical=no] ctx_prf : t = lookup (untag ctx_of) str}. + +Lemma ctx_prf_head str t g : t = lookup ((str, t) :: g) str. +Proof. by rewrite /lookup /= !eqxx. Qed. + +Lemma ctx_prf_tail str t g str' t' : + str' != str -> + t = lookup g str -> + t = lookup ((str', t') :: g) str. +Proof. +move=> str'str tg /=; rewrite /lookup/=. +by case: ifPn => //=; rewrite (negbTE str'str). +Qed. + +Definition recurse_tag g := Tag g. +Canonical found_tag g := recurse_tag g. + +Canonical found str t g : find str t := + @Find str t (found_tag ((str, t) :: g)) + (@ctx_prf_head str t g). + +Canonical recurse str t str' t' {H : infer (str' != str)} + (g : find str t) : find str t := + @Find str t (recurse_tag ((str', t') :: untag (ctx_of g))) + (@ctx_prf_tail str t (untag (ctx_of g)) str' t' H (ctx_prf g)). + +End tagged_context. +Arguments lookup {T} t0 g str. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index bffc1621c9..c9119b43ab 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -1,25 +1,25 @@ (* mathcomp analysis (c) 2022 Inria and AIST. License: CeCILL-C. *) From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. -From mathcomp Require Import rat. -From mathcomp.classical Require Import mathcomp_extra boolp classical_sets. -From mathcomp.classical Require Import functions cardinality fsbigop. -Require Import reals ereal signed topology normedtype sequences esum measure. -Require Import lebesgue_measure numfun lebesgue_integral exp kernel. +From mathcomp Require Import interval_inference archimedean rat. +From mathcomp Require Import mathcomp_extra boolp classical_sets. +From mathcomp Require Import functions cardinality fsbigop. +From mathcomp Require Import reals ereal topology normedtype sequences esum measure. +From mathcomp Require Import lebesgue_measure numfun lebesgue_integral exp kernel. (******************************************************************************) (* Semantics of a probabilistic programming language using s-finite kernels *) (* *) (* bernoulli r1 == Bernoulli probability with r1 a proof that *) (* r : {nonneg R} is smaller than 1 *) -(* *) -(* sample P == sample according to the probability P *) +(* uniform_probability a b ab0 == uniform probability over the interval [a,b] *) +(* sample_cst P == sample according to the probability P *) (* letin l k == execute l, augment the context, and execute k *) (* ret mf == access the context with f and return the result *) (* score mf == observe t from d, where f is the density of d and *) (* t occurs in f *) (* e.g., score (r e^(-r * t)) = observe t from exp(r) *) -(* pnormalize k P == normalize the kernel k into a probability kernel, *) +(* normalize k P == normalize the kernel k into a probability kernel, *) (* P is a default probability in case normalization is *) (* not possible *) (* ite mf k1 k2 == access the context with the boolean function f and *) @@ -41,9 +41,6 @@ Local Open Scope ring_scope. Local Open Scope ereal_scope. (* TODO: PR *) -Lemma onem1' (R : numDomainType) (p : R) : (p + `1- p = 1)%R. -Proof. by rewrite /onem addrCA subrr addr0. Qed. - Lemma onem_nonneg_proof (R : numDomainType) (p : {nonneg R}) : (p%:num <= 1 -> 0 <= `1-(p%:num))%R. Proof. by rewrite /onem/= subr_ge0. Qed. @@ -53,6 +50,39 @@ Definition onem_nonneg (R : numDomainType) (p : {nonneg R}) NngNum (onem_nonneg_proof p1). (* /TODO: PR *) +Lemma invr_nonneg_proof (R : numDomainType) (p : {nonneg R}) : + (0 <= (p%:num)^-1)%R. +Proof. by rewrite invr_ge0. Qed. + +Definition invr_nonneg (R : numDomainType) (p : {nonneg R}) := + NngNum (invr_nonneg_proof p). + +(* TODO: move *) +Lemma eq_probability R d (Y : measurableType d) (m1 m2 : probability Y R) : + (m1 =1 m2 :> (set Y -> \bar R)) -> m1 = m2. +Proof. +move: m1 m2 => [m1 +] [m2 +] /= m1m2. +move/funext : m1m2 => <- -[[c11 c12] [m01] [sf1] [sig1] [fin1] [sub1] [p1]] + [[c21 c22] [m02] [sf2] [sig2] [fin2] [sub2] [p2]]. +have ? : c11 = c21 by []. +subst c21. +have ? : c12 = c22 by []. +subst c22. +have ? : m01 = m02 by []. +subst m02. +have ? : sf1 = sf2 by []. +subst sf2. +have ? : sig1 = sig2 by []. +subst sig2. +have ? : fin1 = fin2 by []. +subst fin2. +have ? : sub1 = sub2 by []. +subst sub2. +have ? : p1 = p2 by []. +subst p2. +by f_equal. +Qed. + Section bernoulli. Variables (R : realType) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). Local Open Scope ring_scope. @@ -69,13 +99,46 @@ Local Close Scope ring_scope. Let bernoulli_setT : bernoulli [set: _] = 1. Proof. rewrite /bernoulli/= /measure_add/= /msum 2!big_ord_recr/= big_ord0 add0e/=. -by rewrite /mscale/= !diracT !mule1 -EFinD onem1'. +by rewrite /mscale/= !diracT !mule1 -EFinD add_onemK. Qed. -HB.instance Definition _ := @Measure_isProbability.Build _ _ R bernoulli bernoulli_setT. +HB.instance Definition _ := + @Measure_isProbability.Build _ _ R bernoulli bernoulli_setT. End bernoulli. +Lemma integral_bernoulli {R : realType} + (p : {nonneg R}) (p1 : (p%:num <= 1)%R) (f : bool -> set bool -> _) U : + (forall x y, 0 <= f x y) -> + \int[bernoulli p1]_y (f y ) U = + p%:num%:E * f true U + (`1-(p%:num))%:E * f false U. +Proof. +move=> f0. +rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. +by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT !mul1e. +Qed. + +Section uniform_probability. +Context {R : realType}. +Variables (a b : R) (ab0 : (0 < b - a)%R). + +Definition uniform_probability := mscale (invr_nonneg (NngNum (ltW ab0))) + (mrestr lebesgue_measure (measurable_itv `[a, b])). + +HB.instance Definition _ := Measure.on uniform_probability. + +Let uniform_probability_setT : uniform_probability [set: _] = 1. +Proof. +rewrite /uniform_probability /mscale/= /mrestr/=. +rewrite setTI lebesgue_measure_itv/= lte_fin. +by rewrite -subr_gt0 ab0 -EFinD -EFinM mulVf// gt_eqF// subr_gt0. +Qed. + +HB.instance Definition _ := @Measure_isProbability.Build _ _ R + uniform_probability uniform_probability_setT. + +End uniform_probability. + Section mscore. Context d (T : measurableType d) (R : realType). Variable f : T -> R. @@ -221,9 +284,9 @@ pose floor_f := widen_ord (leq_addl n `|floor `|f t| |.+1) (Ordinal (ltnSn `|floor `|f t| |)). rewrite big_mkord (bigD1 floor_f)//= ifT; last first. rewrite lee_fin lte_fin; apply/andP; split. - by rewrite natr_absz (@ger0_norm _ (floor `|f t|)) ?floor_ge0 ?floor_le. + by rewrite natr_absz (@ger0_norm _ (floor `|f t|)) ?floor_ge0// floor_le_tmp. rewrite -addn1 natrD natr_absz. - by rewrite (@ger0_norm _ (floor `|f t|)) ?floor_ge0 ?lt_succ_floor. + by rewrite (@ger0_norm _ (floor `|f t|)) ?floor_ge0// intrD1 floorD1_gt. rewrite big1 ?adde0//= => j jk. rewrite ifF// lte_fin lee_fin. move: jk; rewrite neq_ltn/= => /orP[|] jr. @@ -233,7 +296,7 @@ move: jk; rewrite neq_ltn/= => /orP[|] jr. by rewrite -[leRHS](@ger0_norm _ (floor `|f t|)) ?floor_ge0. - suff : (`|f t| < j%:R)%R by rewrite ltNge => /negbTE ->. move: jr; rewrite -ltz_nat -(@ltr_int R) (@gez0_abs (floor `|f t|)) ?floor_ge0//. - by rewrite ltr_int -floor_lt_int. + by rewrite ltr_int floor_lt_int. Qed. HB.instance Definition _ := @@ -283,7 +346,7 @@ by rewrite /mseries eseries0. Qed. #[export] -HB.instance Definition _ t := @Kernel_isSFinite_subdef.Build _ _ _ _ _ +HB.instance Definition _ := @Kernel_isSFinite_subdef.Build _ _ _ _ _ (kiteT k) sfinite_kiteT. End sfkiteT. @@ -298,7 +361,7 @@ by rewrite /= /mzero. Qed. #[export] -HB.instance Definition _ t := Kernel_isFinite.Build _ _ _ _ _ +HB.instance Definition _ := Kernel_isFinite.Build _ _ _ _ _ (kiteT k) kiteT_uub. End fkiteT. @@ -366,6 +429,7 @@ Section ite. Context d d' (T : measurableType d) (T' : measurableType d') (R : realType). Variables (f : T -> bool) (u1 u2 : R.-sfker T ~> T'). +(* NB: not used? *) Definition mite (mf : measurable_fun setT f) : T -> set T' -> \bar R := fun t => if f t then u1 t else u2 t. @@ -405,9 +469,12 @@ Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). Definition ret (f : X -> Y) (mf : measurable_fun setT f) : R.-pker X ~> Y := [the R.-pker _ ~> _ of kdirac mf]. -Definition sample (P : pprobability Y R) : R.-pker X ~> Y := +Definition sample_cst (P : pprobability Y R) : R.-pker X ~> Y := [the R.-pker _ ~> _ of kprobability (measurable_cst P)]. +Definition sample (P : X -> pprobability Y R) (mP : measurable_fun setT P) : R.-pker X ~> Y := + [the R.-pker _ ~> _ of kprobability mP]. + Definition normalize (k : R.-sfker X ~> Y) P : X -> probability Y R := fun x => [the probability _ _ of mnormalize (k x) P]. @@ -417,6 +484,7 @@ Definition ite (f : X -> bool) (mf : measurable_fun setT f) End insn2. Arguments ret {d d' X Y R f} mf. +Arguments sample_cst {d d' X Y R}. Arguments sample {d d' X Y R}. Section insn2_lemmas. @@ -426,7 +494,10 @@ Lemma retE (f : X -> Y) (mf : measurable_fun setT f) x : ret mf x = \d_(f x) :> (_ -> \bar R). Proof. by []. Qed. -Lemma sampleE (P : probability Y R) (x : X) : sample P x = P. +Lemma sample_cstE (P : probability Y R) (x : X) : sample_cst P x = P. +Proof. by []. Qed. + +Lemma sampleE (P : X -> pprobability Y R) (mP : measurable_fun setT P) (x : X) : sample P mP x = P x. Proof. by []. Qed. Lemma normalizeE (f : R.-sfker X ~> Y) P x U : @@ -442,7 +513,7 @@ Proof. apply/eq_measure/funext => U. rewrite /ite; unlock => /=. rewrite /kcomp/= integral_dirac//=. -rewrite indicT mul1e. +rewrite diracT mul1e. rewrite -/(measure_add (ITE.kiteT k1 (x, f x)) (ITE.kiteF k2 (x, f x))). rewrite measure_addE. rewrite /ITE.kiteT /ITE.kiteF/=. @@ -451,6 +522,15 @@ Qed. End insn2_lemmas. +Lemma normalize_kdirac (R : realType) + d (T : measurableType d) d' (T' : measurableType d') (x : T) (r : T') P : + normalize (kdirac (measurable_cst r)) P x = \d_r :> probability T' R. +Proof. +apply: eq_probability => U. +rewrite normalizeE /= diracE in_setT/=. +by rewrite onee_eq0/= indicE in_setT/= -div1r divr1 mule1. +Qed. + Section insn3. Context d d' d3 (X : measurableType d) (Y : measurableType d') (Z : measurableType d3) (R : realType). @@ -493,7 +573,7 @@ Lemma letin_retk x U : measurable U -> letin (ret mf) k x U = k (x, f x) U. Proof. -move=> mU; rewrite letinE retE integral_dirac ?indicT ?mul1e//. +move=> mU; rewrite letinE retE integral_dirac ?diracT ?mul1e//. exact: (measurableT_comp (measurable_kernel k _ mU)). Qed. @@ -523,17 +603,17 @@ Arguments fail {d d' X Y R}. Module Notations. -Notation var1of2 := (@measurable_fst _ _ _ _). +(*Notation var1of2 := (@measurable_fst _ _ _ _). Notation var2of2 := (@measurable_snd _ _ _ _). Notation var1of3 := (measurableT_comp (@measurable_fst _ _ _ _) (@measurable_fst _ _ _ _)). Notation var2of3 := (measurableT_comp (@measurable_snd _ _ _ _) (@measurable_fst _ _ _ _)). -Notation var3of3 := (@measurable_snd _ _ _ _). +Notation var3of3 := (@measurable_snd _ _ _ _).*) -Notation mR := Real_sort__canonical__measure_Measurable. -Notation munit := Datatypes_unit__canonical__measure_Measurable. -Notation mbool := Datatypes_bool__canonical__measure_Measurable. +Notation mR := measurableTypeR. +Notation munit := (unit : measurableType _). +Notation mbool := (bool : measurableType _). End Notations. @@ -544,12 +624,249 @@ Definition kr (r : R) := @measurable_cst _ _ T _ setT r. Definition k3 : measurable_fun _ _ := kr 3%:R. Definition k10 : measurable_fun _ _ := kr 10%:R. Definition ktt := @measurable_cst _ _ T _ setT tt. +Definition kb (b : bool) := @measurable_cst _ _ T _ setT b. End cst_fun. Arguments kr {d T R}. Arguments k3 {d T R}. Arguments k10 {d T R}. Arguments ktt {d T}. +Arguments kb {d T}. + +Section iter_mprod. +Import Notations. + +Fixpoint iter_mprod (l : list {d & measurableType d}) + : {d & measurableType d} := + match l with + | [::] => existT measurableType _ munit + | h :: t => let t' := iter_mprod t in + existT _ _ [the measurableType (projT1 h, projT1 t').-prod of + (projT2 h * projT2 t')%type] + end. + +End iter_mprod. + +Section acc. +Import Notations. +Context {R : realType}. + +Fixpoint acc (l : seq {d & measurableType d}) n : + projT2 (iter_mprod l) -> projT2 (nth (existT _ _ munit) l n) := + match l return + projT2 (iter_mprod l) -> projT2 (nth (existT _ _ munit) l n) + with + | [::] => match n with | O => id | m.+1 => id end + | _ :: _ => match n with + | O => fst + | m.+1 => fun H => acc m H.2 + end + end. + +Lemma measurable_acc (l : seq {d & measurableType d}) n : + measurable_fun setT (@acc l n). +Proof. +by elim: l n => //= h t ih [|m] //; exact: (measurableT_comp (ih _)). +Qed. +End acc. +Arguments acc : clear implicits. +Arguments measurable_acc : clear implicits. + +Section rpair_pairA. +Context d0 d1 d2 (T0 : measurableType d0) (T1 : measurableType d1) + (T2 : measurableType d2). + +Definition rpair d (T : measurableType d) t : + ([the measurableType _ of T0] -> [the measurableType _ of T0 * T])%type := + fun x => (x, t). + +Lemma mrpair d (T : measurableType d) t : measurable_fun setT (@rpair _ T t). +Proof. exact: measurable_fun_prod. Qed. + +Definition pairA : ([the measurableType _ of T0 * T1 * T2] -> + [the measurableType _ of T0 * (T1 * T2)])%type := + fun x => (x.1.1, (x.1.2, x.2)). + +Definition mpairA : measurable_fun setT pairA. +Proof. +apply: measurable_fun_prod => /=; first exact: measurableT_comp. +by apply: measurable_fun_prod => //=; exact: measurableT_comp. +Qed. + +Definition pairAi : ([the measurableType _ of T0 * (T1 * T2)] -> + [the measurableType _ of T0 * T1 * T2])%type := + fun x => (x.1, x.2.1, x.2.2). + +Definition mpairAi : measurable_fun setT pairAi. +Proof. +apply: measurable_fun_prod => //=; last exact: measurableT_comp. +apply: measurable_fun_prod => //=; exact: measurableT_comp. +Qed. + +End rpair_pairA. +Arguments rpair {d0 T0 d} T. +#[global] Hint Extern 0 (measurable_fun _ (rpair _ _)) => + solve [apply: mrpair] : core. +Arguments pairA {d0 d1 d2 T0 T1 T2}. +#[global] Hint Extern 0 (measurable_fun _ pairA) => + solve [apply: mpairA] : core. +Arguments pairAi {d0 d1 d2 T0 T1 T2}. +#[global] Hint Extern 0 (measurable_fun _ pairAi) => + solve [apply: mpairAi] : core. + +Section rpair_pairA_comp. +Import Notations. +Context d0 d1 d2 d3 (T0 : measurableType d0) (T1 : measurableType d1) + (T2 : measurableType d2) (T3 : measurableType d3) (R : realType). + +Definition pairAr d (T : measurableType d) t : + ([the measurableType _ of T0 * T1] -> + [the measurableType _ of T0 * (T1 * T)])%type := + pairA \o rpair T t. +Arguments pairAr {d} T. + +Lemma mpairAr d (T : measurableType d) t : measurable_fun setT (pairAr T t). +Proof. exact: measurableT_comp. Qed. + +Definition pairAAr : ([the measurableType _ of T0 * T1 * T2] -> + [the measurableType _ of T0 * (T1 * (T2 * munit))])%type := + pairA \o pairA \o rpair munit tt. + +Lemma mpairAAr : measurable_fun setT pairAAr. +Proof. by do 2 apply: measurableT_comp => //. Qed. + +Definition pairAAAr : ([the measurableType _ of T0 * T1 * T2 * T3] -> + [the measurableType _ of T0 * (T1 * (T2 * (T3 * munit)))])%type := + pairA \o pairA \o pairA \o rpair munit tt. + +Lemma mpairAAAr : measurable_fun setT pairAAAr. +Proof. by do 3 apply: measurableT_comp => //. Qed. + +Definition pairAArAi : ([the measurableType _ of T0 * (T1 * T2)] -> + [the measurableType _ of T0 * (T1 * (T2 * munit))])%type := + pairAAr \o pairAi. + +Lemma mpairAArAi : measurable_fun setT pairAArAi. +Proof. by apply: measurableT_comp => //=; exact: mpairAAr. Qed. + +Definition pairAAArAAi : ([the measurableType _ of T3 * (T0 * (T1 * T2))] -> + [the measurableType _ of T3 * (T0 * (T1 * (T2 * munit)))])%type := + pairA \o pairA \o pairA \o rpair munit tt \o pairAi \o pairAi. + +Lemma mpairAAARAAAi : measurable_fun setT pairAAArAAi. +Proof. by do 5 apply: measurableT_comp => //=. Qed. + +End rpair_pairA_comp. +Arguments pairAr {d0 d1 T0 T1 d} T. +Arguments pairAAr {d0 d1 d2 T0 T1 T2}. +Arguments pairAAAr {d0 d1 d2 d3 T0 T1 T2 T3}. +Arguments pairAArAi {d0 d1 d2 T0 T1 T2}. +Arguments pairAAArAAi {d0 d1 d2 d3 T0 T1 T2 T3}. + +Section accessor_functions. +Import Notations. +Context d0 d1 d2 d3 (T0 : measurableType d0) (T1 : measurableType d1) + (T2 : measurableType d2) (T3 : measurableType d3) (R : realType). + +Definition Of2 := [:: existT _ _ T0; existT _ _ T1]. + +Definition acc0of2 : [the measurableType _ of (T0 * T1)%type] -> T0 := + @acc Of2 0 \o pairAr munit tt. + +Lemma macc0of2 : measurable_fun setT acc0of2. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of2 0)|exact: mpairAr]. +Qed. + +Definition acc1of2 : [the measurableType _ of (T0 * T1)%type] -> T1 := + acc Of2 1 \o pairAr munit tt. + +Lemma macc1of2 : measurable_fun setT acc1of2. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of2 1)|exact: mpairAr]. +Qed. + +Definition Of3 := [:: existT _ _ T0; existT _ _ T1; existT _ d2 T2]. + +Definition acc1of3 : [the measurableType _ of (T0 * T1 * T2)%type] -> T1 := + acc Of3 1 \o pairAAr. + +Lemma macc1of3 : measurable_fun setT acc1of3. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of3 1)|exact: mpairAAr]. +Qed. + +Definition acc2of3 : [the measurableType _ of (T0 * T1 * T2)%type] -> T2 := + acc Of3 2 \o pairAAr. + +Lemma macc2of3 : measurable_fun setT acc2of3. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of3 2)|exact: mpairAAr]. +Qed. + +Definition acc0of3' : [the measurableType _ of (T0 * (T1 * T2))%type] -> T0 := + acc Of3 0 \o pairAArAi. + +Lemma macc0of3' : measurable_fun setT acc0of3'. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of3 0)|exact: mpairAArAi]. +Qed. + +Definition acc1of3' : [the measurableType _ of (T0 * (T1 * T2))%type] -> T1 := + acc Of3 1 \o pairAArAi. + +Lemma macc1of3' : measurable_fun setT acc1of3'. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of3 1)|exact: mpairAArAi]. +Qed. + +Definition acc2of3' : [the measurableType _ of (T0 * (T1 * T2))%type] -> T2 := + acc Of3 2 \o pairAArAi. + +Lemma macc2of3' : measurable_fun setT acc2of3'. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of3 2)|exact: mpairAArAi]. +Qed. + +Definition Of4 := + [:: existT _ _ T0; existT _ _ T1; existT _ d2 T2; existT _ d3 T3]. + +Definition acc1of4 : [the measurableType _ of (T0 * T1 * T2 * T3)%type] -> T1 := + acc Of4 1 \o pairAAAr. + +Lemma macc1of4 : measurable_fun setT acc1of4. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of4 1)|exact: mpairAAAr]. +Qed. + +Definition acc2of4' : + [the measurableType _ of (T0 * (T1 * (T2 * T3)))%type] -> T2 := + acc Of4 2 \o pairAAArAAi. + +Lemma macc2of4' : measurable_fun setT acc2of4'. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of4 2)|exact: mpairAAARAAAi]. +Qed. + +Definition acc3of4 : [the measurableType _ of (T0 * T1 * T2 * T3)%type] -> T3 := + acc Of4 3 \o pairAAAr. + +Lemma macc3of4 : measurable_fun setT acc3of4. +Proof. +by apply: measurableT_comp; [exact: (measurable_acc Of4 3)|exact: mpairAAAr]. +Qed. + +End accessor_functions. +Arguments macc0of2 {d0 d1 _ _}. +Arguments macc1of2 {d0 d1 _ _}. +Arguments macc0of3' {d0 d1 d2 _ _ _}. +Arguments macc1of3 {d0 d1 d2 _ _ _}. +Arguments macc1of3' {d0 d1 d2 _ _ _}. +Arguments macc2of3 {d0 d1 d2 _ _ _}. +Arguments macc2of3' {d0 d1 d2 _ _ _}. +Arguments macc1of4 {d0 d1 d2 d3 _ _ _ _}. +Arguments macc2of4' {d0 d1 d2 d3 _ _ _ _}. +Arguments macc3of4 {d0 d1 d2 d3 _ _ _ _}. Section insn1_lemmas. Import Notations. @@ -561,16 +878,18 @@ Let kcomp_scoreE d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) (score mf \; g) r U = `|f r|%:E * g (r, tt) U. Proof. rewrite /= /kcomp /kscore /= ge0_integral_mscale//=. -by rewrite integral_dirac// indicT mul1e. +by rewrite integral_dirac// diracT mul1e. Qed. Lemma scoreE d' (T' : measurableType d') (x : T * T') (U : set T') (f : R -> R) (r : R) (r0 : (0 <= r)%R) (f0 : (forall r, 0 <= r -> 0 <= f r)%R) (mf : measurable_fun setT f) : - score (measurableT_comp mf var2of2) + score (measurableT_comp mf (@macc1of2 _ _ _ _)) (x, r) (curry (snd \o fst) x @^-1` U) = (f r)%:E * \d_x.2 U. -Proof. by rewrite /score/= /mscale/= ger0_norm// f0. Qed. +Proof. +by rewrite /score/= /mscale/= ger0_norm//= f0. +Qed. Lemma score_score (f : R -> R) (g : R * unit -> R) (mf : measurable_fun setT f) @@ -583,18 +902,16 @@ rewrite {1}/letin; unlock. by rewrite kcomp_scoreE/= /mscale/= diracE normrM muleA EFinM. Qed. -Import Notations. - (* hard constraints to express score below 1 *) Lemma score_fail (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : score (kr r%:num) = - letin (sample [the probability _ _ of bernoulli r1] : R.-pker T ~> _) - (ite var2of2 (ret ktt) fail). + letin (sample_cst (bernoulli r1) : R.-pker T ~> _) + (ite (@macc1of2 _ _ _ _) (ret ktt) fail). Proof. apply/eq_sfkernel => x U. rewrite letinE/= /sample; unlock. -rewrite integral_measure_add//= ge0_integral_mscale//= ge0_integral_mscale//=. -rewrite integral_dirac//= integral_dirac//= !indicT/= !mul1e. +rewrite ge0_integral_measure_add//= ge0_integral_mscale//= ge0_integral_mscale//=. +rewrite integral_dirac//= integral_dirac//= !diracT/= !mul1e. by rewrite /mscale/= iteE//= iteE//= failE mule0 adde0 ger0_norm. Qed. @@ -603,7 +920,8 @@ End insn1_lemmas. Section letin_ite. Context d d2 d3 (T : measurableType d) (T2 : measurableType d2) (Z : measurableType d3) (R : realType). -Variables (k1 k2 : R.-sfker T ~> Z) (u : R.-sfker [the measurableType _ of (T * Z)%type] ~> T2) +Variables (k1 k2 : R.-sfker T ~> Z) + (u : R.-sfker [the measurableType _ of (T * Z)%type] ~> T2) (f : T -> bool) (mf : measurable_fun setT f) (t : T) (U : set T2). @@ -671,11 +989,11 @@ Let T0 z : (T z) set0 = 0. Proof. by []. Qed. Let T_ge0 z x : 0 <= (T z) x. Proof. by []. Qed. Let T_semi_sigma_additive z : semi_sigma_additive (T z). Proof. exact: measure_semi_sigma_additive. Qed. -HB.instance Definition _ z := @isMeasure.Build _ R X (T z) (T0 z) (T_ge0 z) +HB.instance Definition _ z := @isMeasure.Build _ X R (T z) (T0 z) (T_ge0 z) (@T_semi_sigma_additive z). Let sfinT z : sfinite_measure (T z). Proof. exact: sfinite_kernel_measure. Qed. -HB.instance Definition _ z := @Measure_isSFinite_subdef.Build _ X R +HB.instance Definition _ z := @isSFinite.Build _ X R (T z) (sfinT z). Definition U z : set Y -> \bar R := u z. @@ -683,20 +1001,20 @@ Let U0 z : (U z) set0 = 0. Proof. by []. Qed. Let U_ge0 z x : 0 <= (U z) x. Proof. by []. Qed. Let U_semi_sigma_additive z : semi_sigma_additive (U z). Proof. exact: measure_semi_sigma_additive. Qed. -HB.instance Definition _ z := @isMeasure.Build _ R Y (U z) (U0 z) (U_ge0 z) +HB.instance Definition _ z := @isMeasure.Build _ Y R (U z) (U0 z) (U_ge0 z) (@U_semi_sigma_additive z). Let sfinU z : sfinite_measure (U z). Proof. exact: sfinite_kernel_measure. Qed. -HB.instance Definition _ z := @Measure_isSFinite_subdef.Build _ Y R +HB.instance Definition _ z := @isSFinite.Build _ Y R (U z) (sfinU z). Lemma letinC z A : measurable A -> letin t (letin u' - (ret (measurable_fun_prod var2of3 var3of3))) z A = + (ret (measurable_fun_prod macc1of3 macc2of3))) z A = letin u (letin t' - (ret (measurable_fun_prod var3of3 var2of3))) z A. + (ret (measurable_fun_prod macc2of3 macc1of3))) z A. Proof. move=> mA. rewrite !letinE. @@ -730,7 +1048,7 @@ by rewrite /onem/= -{1}(@divrr _ n.+1%:R) ?unitfE// -mulrBl -natr1 addrK. Qed. Lemma p1S n : (1 / n.+1%:R)%:nng%:num <= 1 :> R. -Proof. by rewrite ler_pdivr_mulr//= mul1r ler1n. Qed. +Proof. by rewrite ler_pdivrMr//= mul1r ler1n. Qed. Lemma p12 : (1 / 2%:R)%:nng%:num <= 1 :> R. Proof. by rewrite p1S. Qed. @@ -740,35 +1058,42 @@ Lemma onem27 : `1- (2 / 7%:R) = (5%:R / 7%:R)%:nng%:num :> R. Proof. by apply/eqP; rewrite subr_eq/= -mulrDl -natrD divrr// unitfE. Qed. Lemma p27 : (2 / 7%:R)%:nng%:num <= 1 :> R. -Proof. by rewrite /= lter_pdivr_mulr// mul1r ler_nat. Qed. +Proof. by rewrite /= lter_pdivrMr// mul1r ler_nat. Qed. End constants. Arguments p12 {R}. Arguments p14 {R}. Arguments p27 {R}. +Arguments p1S {R}. Section poisson. Variable R : realType. Local Open Scope ring_scope. (* density function for Poisson *) -Definition poisson k r : R := r ^+ k / k`!%:R^-1 * expR (- r). +Definition poisson k r : R := + if r > 0 then r ^+ k / k`!%:R^-1 * expR (- r) else 1%:R. -Lemma poisson_ge0 k r : 0 <= r -> 0 <= poisson k r. +Lemma poisson_ge0 k r : 0 <= poisson k r. Proof. -move=> r0; rewrite /poisson mulr_ge0 ?expR_ge0//. -by rewrite mulr_ge0// exprn_ge0. +rewrite /poisson; case: ifPn => r0//. +by rewrite mulr_ge0 ?expR_ge0// mulr_ge0// exprn_ge0 ?ltW. Qed. Lemma poisson_gt0 k r : 0 < r -> 0 < poisson k.+1 r. Proof. -move=> r0; rewrite /poisson mulr_gt0 ?expR_gt0//. +move=> r0; rewrite /poisson r0 mulr_gt0 ?expR_gt0//. by rewrite divr_gt0// ?exprn_gt0// invr_gt0 ltr0n fact_gt0. Qed. -Lemma mpoisson k : measurable_fun setT (poisson k). +Lemma measurable_poisson k : measurable_fun setT (poisson k). Proof. -by apply: measurable_funM => /=; [exact: measurable_funM|exact: measurableT_comp]. +rewrite /poisson; apply: measurable_fun_if => //. + apply: (measurable_fun_bool true) => /=. + rewrite setTI (_ : _ @^-1` _ = `]0, +oo[%classic)//. + by apply/seteqP; split => x /=; rewrite in_itv/= andbT. +by apply: measurable_funM => /=; + [exact: measurable_funM|exact: measurableT_comp]. Qed. Definition poisson3 := poisson 4 3%:R. (* 0.168 *) @@ -800,12 +1125,12 @@ End exponential. Lemma letin_sample_bernoulli d d' (T : measurableType d) (T' : measurableType d') (R : realType)(r : {nonneg R}) (r1 : (r%:num <= 1)%R) (u : R.-sfker [the measurableType _ of (T * bool)%type] ~> T') x y : - letin (sample [the probability _ _ of bernoulli r1]) u x y = + letin (sample_cst (bernoulli r1)) u x y = r%:num%:E * u (x, true) y + (`1- (r%:num))%:E * u (x, false) y. Proof. rewrite letinE/=. rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. -by rewrite !ge0_integral_mscale//= !integral_dirac//= indicT 2!mul1e. +by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT 2!mul1e. Qed. Section sample_and_return. @@ -814,8 +1139,8 @@ Context d (T : measurableType d) (R : realType). Definition sample_and_return : R.-sfker T ~> _ := letin - (sample [the probability _ _ of bernoulli p27]) (* T -> B *) - (ret var2of2) (* T * B -> B *). + (sample_cst [the probability _ _ of bernoulli p27]) (* T -> B *) + (ret macc1of2) (* T * B -> B *). Lemma sample_and_returnE t U : sample_and_return t U = (2 / 7%:R)%:E * \d_true U + (5%:R / 7%:R)%:E * \d_false U. @@ -834,16 +1159,16 @@ Context d (T : measurableType d) (R : realType). let r = case x of {(1, _) => return (k3()), (2, _) => return (k10())} in return r *) -Definition sample_and_branch : R.-sfker T ~> mR R := +Definition sample_and_branch : R.-sfker T ~> R := letin - (sample [the probability _ _ of bernoulli p27]) (* T -> B *) - (ite var2of2 (ret k3) (ret k10)). + (sample_cst [the probability _ _ of bernoulli p27]) (* T -> B *) + (ite macc1of2 (ret k3) (ret k10)). Lemma sample_and_branchE t U : sample_and_branch t U = (2 / 7%:R)%:E * \d_(3%:R : R) U + (5%:R / 7%:R)%:E * \d_(10%:R : R) U. Proof. -by rewrite /sample_and_branch letin_sample_bernoulli/= !iteE !retE onem27. +by rewrite /sample_and_branch letin_sample_bernoulli/= !iteE !retE// onem27. Qed. End sample_and_branch. @@ -859,8 +1184,8 @@ Lemma measurable_fun_mand (x y : T * mbool * mbool -> mbool) : measurable_fun setT x -> measurable_fun setT y -> measurable_fun setT (mand x y). Proof. -move=> /= mx my; apply: (@measurable_fun_bool _ _ _ _ true). -rewrite [X in measurable X](_ : _ = +move=> /= mx my; apply: (measurable_fun_bool true) => /=. +rewrite setTI [X in measurable X](_ : _ = (x @^-1` [set true]) `&` (y @^-1` [set true])); last first. by rewrite /mand; apply/seteqP; split => z/= /andP. apply: measurableI. @@ -869,13 +1194,13 @@ apply: measurableI. Qed. Definition bernoulli_and : R.-sfker T ~> mbool := - (letin (sample [the probability _ _ of bernoulli p12]) - (letin (sample [the probability _ _ of bernoulli p12]) - (ret (measurable_fun_mand var2of3 var3of3)))). + (letin (sample_cst [the probability _ _ of bernoulli p12]) + (letin (sample_cst [the probability _ _ of bernoulli p12]) + (ret (measurable_fun_mand macc1of3 macc2of3)))). Lemma bernoulli_andE t U : bernoulli_and t U = - sample [the probability _ _ of bernoulli p14] t U. + sample_cst (bernoulli p14) t U. Proof. rewrite /bernoulli_and 3!letin_sample_bernoulli/= /mand/= muleDr//= -muleDl//. rewrite !muleA -addeA -muleDl// -!EFinM !onem1S/= -splitr mulr1. @@ -893,11 +1218,11 @@ Import Notations. Context d (T : measurableType d) (R : realType) (h : R -> R). Hypothesis mh : measurable_fun setT h. Definition kstaton_bus : R.-sfker T ~> mbool := - letin (sample [the probability _ _ of bernoulli p27]) + letin (sample_cst [the probability _ _ of bernoulli p27]) (letin - (letin (ite var2of2 (ret k3) (ret k10)) - (score (measurableT_comp mh var3of3))) - (ret var2of3)). + (letin (ite macc1of2 (ret k3) (ret k10)) + (score (measurableT_comp mh macc2of3))) + (ret macc1of3)). Definition staton_bus := normalize kstaton_bus. @@ -911,9 +1236,9 @@ Section staton_bus_poisson. Import Notations. Context d (T : measurableType d) (R : realType). Let poisson4 := @poisson R 4%N. -Let mpoisson4 := @mpoisson R 4%N. +Let mpoisson4 := @measurable_poisson R 4%N. -Definition kstaton_bus_poisson : R.-sfker (mR R) ~> mbool := +Definition kstaton_bus_poisson : R.-sfker R ~> mbool := kstaton_bus _ mpoisson4. Let kstaton_bus_poissonE t U : kstaton_bus_poisson t U = @@ -963,7 +1288,7 @@ Let mexp1560 := @mexp_density R (ratr (15%:Q / 60%:Q)). (* 15/60 = 0.25 *) -Definition kstaton_bus_exponential : R.-sfker (mR R) ~> mbool := +Definition kstaton_bus_exponential : R.-sfker R ~> mbool := kstaton_bus _ mexp1560. Let kstaton_bus_exponentialE t U : kstaton_bus_exponential t U = @@ -1001,3 +1326,574 @@ by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// exp_density_gt0 ?ltr0n. Qed. End staton_bus_exponential. + +(* X + Y is a measurableType if X and Y are *) +HB.instance Definition _ (X Y : pointedType) := + isPointed.Build (X + Y)%type (@inl X Y point). + +Section measurable_sum. +Context d d' (X : measurableType d) (Y : measurableType d'). + +Definition measurable_sum : set (set (X + Y)) := setT. + +Let sum0 : measurable_sum set0. Proof. by []. Qed. + +Let sumC A : measurable_sum A -> measurable_sum (~` A). Proof. by []. Qed. + +Let sumU (F : (set (X + Y))^nat) : (forall i, measurable_sum (F i)) -> + measurable_sum (\bigcup_i F i). +Proof. by []. Qed. + +HB.instance Definition _ := @isMeasurable.Build default_measure_display (X + Y)%type + measurable_sum sum0 sumC sumU. + +End measurable_sum. + +Lemma measurable_fun_sum dA dB d' (A : measurableType dA) (B : measurableType dB) + (Y : measurableType d') (f : A -> Y) (g : B -> Y) : + measurable_fun setT f -> measurable_fun setT g -> + measurable_fun setT (fun tb : A + B => + match tb with inl a => f a | inr b => g b end). +Proof. +move=> mx my/= _ Z mZ /=; rewrite setTI /=. +rewrite (_ : _ @^-1` Z = inl @` (f @^-1` Z) `|` inr @` (g @^-1` Z)). + exact: measurableU. +apply/seteqP; split. + by move=> [a Zxa|b Zxb]/=; [left; exists a|right; exists b]. +by move=> z [/= [a Zxa <-//=]|]/= [b Zyb <-//=]. +Qed. + +(* TODO: measurable_fun_if_pair -> measurable_fun_if_pair_bool? *) +Lemma measurable_fun_if_pair_nat d d' (X : measurableType d) + (Y : measurableType d') (f g : X -> Y) (n : nat) : + measurable_fun setT f -> measurable_fun setT g -> + measurable_fun setT (fun xn => if xn.2 == n then f xn.1 else g xn.1). +Proof. +move=> mx my; apply: measurable_fun_ifT => //=. +- have h : measurable_fun [set: nat] (fun t => t == n) by []. + exact: (@measurableT_comp _ _ _ _ _ _ _ _ _ h). +- exact: measurableT_comp. +- exact: measurableT_comp. +Qed. + +Module CASE_NAT. +Section case_nat. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). + +Section case_nat_ker. +Variable k : R.-ker X ~> Y. + +Definition case_nat_ j : X * nat -> {measure set Y -> \bar R} := + fun xn => if xn.2 == j then k xn.1 else [the measure _ _ of mzero]. + +Let measurable_fun_case_nat_ m U : measurable U -> + measurable_fun setT (case_nat_ m ^~ U). +Proof. +move=> mU; rewrite /case_nat_ (_ : (fun _ => _) = + (fun x => if x.2 == m then k x.1 U else mzero U)) /=; last first. + by apply/funext => -[t b]/=; case: ifPn. +apply: (@measurable_fun_if_pair_nat _ _ _ _ (k ^~ U) (fun=> mzero U)) => //. +exact/measurable_kernel. +Qed. + +#[export] +HB.instance Definition _ j := isKernel.Build _ _ _ _ _ + (case_nat_ j) (measurable_fun_case_nat_ j). +End case_nat_ker. + +Section sfcase_nat. +Variable k : R.-sfker X ~> Y. + +Let sfcase_nat_ j : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> case_nat_ k j x U = mseries (k_ ^~ x) 0 U. +Proof. +have [k_ hk /=] := sfinite_kernel k. +exists (fun n => [the _.-ker _ ~> _ of case_nat_ (k_ n) j]) => /=. + move=> n; have /measure_fam_uubP[r k_r] := measure_uub (k_ n). + exists r%:num => /= -[x [|n']]; rewrite /case_nat_//= /mzero//. + by case: ifPn => //= ?; rewrite /mzero. + by case: ifPn => // ?; rewrite /= /mzero. +move=> [x b] U mU; rewrite /case_nat_; case: ifPn => hb; first by rewrite hk. +by rewrite /mseries eseries0. +Qed. + +#[export] +HB.instance Definition _ j := @Kernel_isSFinite_subdef.Build _ _ _ _ _ + (case_nat_ k j) (sfcase_nat_ j). +End sfcase_nat. + +Section fkcase_nat. +Variable k : R.-fker X ~> Y. + +Let case_nat_uub (m : nat) : measure_fam_uub (case_nat_ k m). +Proof. +have /measure_fam_uubP[M hM] := measure_uub k. +exists M%:num => /= -[]; rewrite /case_nat_ => t [|m']/=. + by case: ifPn => //= ?; rewrite /mzero//=. +by case: ifPn => //= ?; rewrite /mzero//=. +Qed. + +#[export] +HB.instance Definition _ j := Kernel_isFinite.Build _ _ _ _ _ + (case_nat_ k j) (case_nat_uub j). +End fkcase_nat. + +End case_nat. +End CASE_NAT. + +Import CASE_NAT. + +Section case_nat. +Context d d' (T : measurableType d) (T' : measurableType d') (R : realType). + +Import CASE_NAT. + +(* case analysis on the nat datatype *) +Definition case_nat (t : R.-sfker T ~> nat) (u_ : (R.-sfker T ~> T')^nat) + : R.-sfker T ~> T' := + [the R.-sfker T ~> nat of t] \; + [the R.-sfker T * nat ~> T' of + kseries (fun n => [the R.-sfker T * nat ~> T' of case_nat_ (u_ n) n])]. + +End case_nat. + +Definition measure_sum_display : + (measure_display * measure_display) -> measure_display. +Proof. exact. Qed. + +Definition image_classes d1 d2 + (T1 : measurableType d1) (T2 : measurableType d2) (T : Type) + (f1 : T1 -> T) (f2 : T2 -> T) := + <>. + +Section sum_salgebra_instance. +Context d1 d2 (T1 : measurableType d1) (T2 : measurableType d2). +Let f1 : T1 -> T1 + T2 := @inl T1 T2. +Let f2 : T2 -> T1 + T2 := @inr T1 T2. + +Lemma sum_salgebra_set0 : image_classes f1 f2 (set0 : set (T1 + T2)). +Proof. exact: sigma_algebra0. Qed. + +Lemma sum_salgebra_setC A : image_classes f1 f2 A -> + image_classes f1 f2 (~` A). +Proof. exact: sigma_algebraC. Qed. + +Lemma sum_salgebra_bigcup (F : _^nat) : (forall i, image_classes f1 f2 (F i)) -> + image_classes f1 f2 (\bigcup_i (F i)). +Proof. exact: sigma_algebra_bigcup. Qed. + +HB.instance Definition sum_salgebra_mixin := + @isMeasurable.Build (measure_sum_display (d1, d2)) + (T1 + T2)%type (image_classes f1 f2) + (sum_salgebra_set0) (sum_salgebra_setC) (sum_salgebra_bigcup). + +End sum_salgebra_instance. +Reserved Notation "p .-sum" (at level 1, format "p .-sum"). +Reserved Notation "p .-sum.-measurable" + (at level 2, format "p .-sum.-measurable"). +Notation "p .-sum" := (measure_sum_display p) : measure_display_scope. +Notation "p .-sum.-measurable" := + ((p.-sum).-measurable : set (set (_ + _))) : + classical_set_scope. + +Module CASE_SUM. + +Section case_suml. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). +Let A : measurableType _ := unit. + +Section kcase_suml. +Variable k : R.-ker X ~> Y. + +Definition case_suml (a : A) : X * A -> {measure set Y -> \bar R} := + fun xa => k xa.1. + +Let measurable_fun_case_suml a U : measurable U -> + measurable_fun setT (case_suml a ^~ U). +Proof. +move=> /= mU; rewrite /case_suml. +have h := measurable_kernel k _ mU. +rewrite (_ : (fun x : X * unit => k x.1 U) = (fun x : X => k x U) \o fst) //. +by apply: measurableT_comp => //. +Qed. + +#[export] +HB.instance Definition _ a := isKernel.Build _ _ _ _ _ + (case_suml a) (measurable_fun_case_suml a). +End kcase_suml. + +Section sfkcase_suml. +Variable k : R.-sfker X ~> Y. + +Let sfinite_case_suml (a : A) : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> case_suml k a x U = mseries (k_ ^~ x) 0 U. +Proof. +have [k_ hk /=] := sfinite_kernel k. +exists (fun n => [the _.-ker _ ~> _ of case_suml (k_ n) a]) => /=. + move=> n; have /measure_fam_uubP[r k_r] := measure_uub (k_ n). + by exists r%:num => /= -[x []]; rewrite /case_suml//= /mzero//. +move=> [x b] U mU; rewrite /case_suml /=. +by rewrite /mseries hk. +Qed. + +#[export] +HB.instance Definition _ (a : A) := @Kernel_isSFinite_subdef.Build _ _ _ _ _ + (case_suml k a) (sfinite_case_suml a). +End sfkcase_suml. + +Section fkcase_suml. +Variable k : R.-fker X ~> Y. + +Let case_suml_uub (a : A) : measure_fam_uub (case_suml k a). +Proof. +have /measure_fam_uubP[M hM] := measure_uub k. +by exists M%:num => /= -[]; rewrite /case_suml. +Qed. + +#[export] +HB.instance Definition _ a := Kernel_isFinite.Build _ _ _ _ _ + (case_suml k a) (case_suml_uub a). +End fkcase_suml. + +End case_suml. + +Section case_sumr. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). +Let B : measurableType _ := bool. + +Section kcase_sumr. +Variable k : R.-ker X ~> Y. + +Definition case_sumr (b : B) : X * B -> {measure set Y -> \bar R} := + fun xa => if xa.2 == b then k xa.1 else [the measure _ _ of mzero]. + +Let measurable_fun_case_sumr b U : measurable U -> + measurable_fun setT (case_sumr b ^~ U). +Proof. +move=> /= mU; rewrite /case_suml. +have h := measurable_kernel k _ mU. +rewrite /case_sumr. +rewrite (_ : (fun x : X * bool => case_sumr b x U) = + (fun x : X * bool => (if x.2 == b then k x.1 U else [the {measure set Y -> \bar R} of mzero] U))); last first. + apply/funext => x. + rewrite /case_sumr. + by case: ifPn. +apply: measurable_fun_ifT => //=. + rewrite (_ : (fun t : X * bool => t.2 == b) = (fun t : bool => t == b) \o snd)//. + apply: measurableT_comp => //. +rewrite (_ : (fun t : X * bool => k t.1 U) = (fun t : X => k t U) \o fst)//. +by apply: measurableT_comp => //. +Qed. + +#[export] +HB.instance Definition _ b := isKernel.Build _ _ _ _ _ + (case_sumr b) (measurable_fun_case_sumr b). +End kcase_sumr. + +Section sfkcase_sumr. +Variable k : R.-sfker X ~> Y. + +Let sfinite_case_sumr b : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> case_sumr k b x U = mseries (k_ ^~ x) 0 U. +Proof. +have [k_ hk /=] := sfinite_kernel k. +exists (fun n => [the _.-ker _ ~> _ of case_sumr (k_ n) b]) => /=. + move=> n; have /measure_fam_uubP[r k_r] := measure_uub (k_ n). + by exists r%:num => /= -[x []]; rewrite /case_sumr//=; case: ifPn => // _; + rewrite /= (le_lt_trans _ (k_r x))// /mzero//. +move=> [x [|]] U mU; rewrite /case_sumr /=; case: b => //=; rewrite ?hk//; +by rewrite /mseries/= eseries0. +Qed. + +#[export] +HB.instance Definition _ b := @Kernel_isSFinite_subdef.Build _ _ _ _ _ + (case_sumr k b) (sfinite_case_sumr b). +End sfkcase_sumr. + +Section fkcase_sumr. +Variable k : R.-fker X ~> Y. + +Let case_sumr_uub b : measure_fam_uub (case_sumr k b). +Proof. +have /measure_fam_uubP[M hM] := measure_uub k. +by exists M%:num => /= -[]; rewrite /case_sumr => x [|] /=; case: b => //=; + rewrite (le_lt_trans _ (hM x))// /mzero. +Qed. + +#[export] +HB.instance Definition _ b := Kernel_isFinite.Build _ _ _ _ _ + (case_sumr k b) (case_sumr_uub b). +End fkcase_sumr. + +End case_sumr. +End CASE_SUM. + +Section case_sum'. + +Section kcase_sum'. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). +Let A : measurableType _ := unit. +Let B : measurableType _ := bool. +Variables (k : (A + B)%type -> R.-sfker X ~> Y). + +Definition case_sum' : X * (A + B)%type -> {measure set Y -> \bar R} := + fun xab => match xab with + | (x, inl a) => CASE_SUM.case_suml (k xab.2) a (x, a) + | (x, inr b) => CASE_SUM.case_sumr (k xab.2) b (x, b) + end. + +Let measurable_fun_case_sum' U : measurable U -> + measurable_fun setT (case_sum' ^~ U). +Proof. +rewrite /= => mU. +apply: (measurability _ (ErealGenInftyO.measurableE R)) => //. +move=> /= _ [_ [x ->] <-]; apply: measurableI => //. +rewrite /case_sum' /CASE_SUM.case_suml /CASE_SUM.case_sumr /=. +rewrite (_ : + (fun x : X * (unit + bool) => (let (x0, s) := x in + match s with inl _ => k x.2 x0 | inr b => if b == b then k x.2 x0 else mzero + end) U) = + (fun x : X * (unit + bool) => k x.2 x.1 U)); last first. + by apply/funext => -[x1 [a|b]] //; rewrite eqxx. +rewrite (_ : _ @^-1` _ = + ([set x1 | k (inl tt) x1 U < x%:E] `*` inl @` [set tt]) `|` + ([set x1 | k (inr false) x1 U < x%:E] `*` inr @` [set false]) `|` + ([set x1 | k (inr true) x1 U < x%:E] `*` inr @` [set true])); last first. + apply/seteqP; split. + - move=> z /=; rewrite in_itv/=; move: z => [z [[]|[|]]]//= ?. + + by do 2 left; split => //; exists tt. + + by right; split => //; exists true. + + by left; right; split => //; exists false. + - move=> z /=; rewrite in_itv/=; move: z => [z [[]|[|]]]//=. + - move=> [[[]//|]|]. + + by move=> [_ []]. + + by move=> [_ []]. + - move=> [[|]|[]//]. + + by move=> [_ []]. + + by move=> [_ [] [|]]. + - move=> [[|[]//]|]. + + by move=> [_ []]. + + by move=> [_ [] [|]]. +pose h1 := [set xub : X * (unit + bool) | k (inl tt) xub.1 U < x%:E]. +have mh1 : measurable h1. + rewrite -[X in measurable X]setTI; apply: emeasurable_fun_infty_o => //=. + have H : measurable_fun [set: X] (fun x => k (inl tt) x U) by exact/measurable_kernel. + move=> _ /= C mC; rewrite setTI. + have := H measurableT _ mC; rewrite setTI => {}H. + rewrite [X in measurable X](_ : _ = ((fun x => k (inl tt) x U) @^-1` C) `*` setT)//. + exact: measurableM. + by apply/seteqP; split => [z//=| z/= []]. +set h2 := [set xub : X * (unit + bool)| k (inr false) xub.1 U < x%:E]. +have mh2 : measurable h2. + rewrite -[X in measurable X]setTI. + apply: emeasurable_fun_infty_o => //=. + have H : measurable_fun [set: X] (fun x => k (inr false) x U) by exact/measurable_kernel. + move=> _ /= C mC; rewrite setTI. + have := H measurableT _ mC; rewrite setTI => {}H. + rewrite [X in measurable X](_ : _ = ((fun x => k (inr false) x U) @^-1` C) `*` setT)//. + exact: measurableM. + by apply/seteqP; split => [z //=|z/= []]. +set h3 := [set xub : X * (unit + bool)| k (inr true) xub.1 U < x%:E]. +have mh3 : measurable h3. + rewrite -[X in measurable X]setTI. + apply: emeasurable_fun_infty_o => //=. + have H : measurable_fun [set: X] (fun x => k (inr true) x U) by exact/measurable_kernel. + move=> _ /= C mC; rewrite setTI. + have := H measurableT _ mC; rewrite setTI => {}H. + rewrite [X in measurable X](_ : _ = ((fun x => k (inr true) x U) @^-1` C) `*` setT)//. + exact: measurableM. + by apply/seteqP; split=> [z//=|z/= []]. +apply: measurableU. +- apply: measurableU. + + apply: measurableM => //. + rewrite [X in measurable X](_ : _ = ysection h1 (inl tt))//. + * by apply: measurable_ysection. + * by apply/seteqP; split => z /=; rewrite /ysection /= inE. + + apply: measurableM => //. + rewrite [X in measurable X](_ : _ = ysection h2 (inr false))//. + * by apply: measurable_ysection. + * by apply/seteqP; split => z /=; rewrite /ysection /= inE. +- apply: measurableM => //. + rewrite [X in measurable X](_ : _ = ysection h3 (inr true))//. + + by apply: measurable_ysection. + + by apply/seteqP; split => z /=; rewrite /ysection /= inE. +Qed. + +#[export] +HB.instance Definition _ := isKernel.Build _ _ _ _ _ + (case_sum') (measurable_fun_case_sum'). +End kcase_sum'. + +Section sfkcase_sum'. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). +Let A : measurableType _ := unit. +Let B : measurableType _ := bool. +Variables (k : (A + B)%type -> R.-sfker X ~> Y). + +Let sfinite_case_sum' : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> case_sum' k x U = mseries (k_ ^~ x) 0 U. +Proof. +rewrite /=. +set f : A + B -> (R.-fker _ ~> _)^nat := + fun ab : A + B => sval (cid (sfinite_kernel (k ab))). +set Hf := fun ab : A + B => svalP (cid (sfinite_kernel (k ab))). +rewrite /= in Hf. +exists (fun n => [the R.-ker _ ~> _ of case_sum' (fun ab => [the R.-fker _ ~> _ of f ab n])]). + move=> n /=. + have [rtt Hrtt] := measure_uub (f (inl tt) n). + have [rfalse Hrfalse] := measure_uub (f (inr false) n). + have [rtrue Hrtrue] := measure_uub (f (inr true) n). + exists (maxr rtt (maxr rfalse rtrue)) => //= -[x [[]|[|]]] /=. + by rewrite 2!EFin_max lt_max Hrtt. + by rewrite /CASE_SUM.case_sumr /= 2!EFin_max 2!lt_max Hrtrue 2!orbT. + by rewrite /CASE_SUM.case_sumr /= 2!EFin_max 2!lt_max Hrfalse orbT. +move=> [x [[]|[|]]] U mU/=-. +by rewrite (Hf (inl tt) x _ mU). +by rewrite (Hf (inr true) x _ mU). +by rewrite (Hf (inr false) x _ mU). +Qed. + +#[export] +HB.instance Definition _ := @Kernel_isSFinite_subdef.Build _ _ _ _ _ + (case_sum' k) (sfinite_case_sum'). +End sfkcase_sum'. + +End case_sum'. + +Section case_sum. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). +Let A : measurableType _ := unit. +Let B : measurableType _ := bool. + +Import CASE_SUM. + +(* case analysis on the datatype unit + bool *) +Definition case_sum (f : R.-sfker X ~> (A + B)%type) + (k : (A + B)%type-> R.-sfker X ~> Y) : R.-sfker X ~> Y := + [the R.-sfker X ~> (A + B)%type of f] \; + [the R.-sfker X * (A + B) ~> Y of case_sum' k]. + +End case_sum. + +(* counting measure as a kernel *) +Section kcounting. +Context d (G : measurableType d) (R : realType). + +Definition kcounting : G -> {measure set nat -> \bar R} := fun=> counting. + +Let mkcounting U : measurable U -> measurable_fun setT (kcounting ^~ U). +Proof. by []. Qed. + +HB.instance Definition _ := isKernel.Build _ _ _ _ _ kcounting mkcounting. + +Let sfkcounting : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> kcounting x U = mseries (k_ ^~ x) 0 U. +Proof. +exists (fun n => [the R.-fker _ ~> _ of + @kdirac _ _ G nat R _ (@measurable_cst _ _ _ _ setT n)]). + by move=> n /=; exact: measure_uub. +by move=> g U mU; rewrite /kcounting/= counting_dirac. +Qed. + +HB.instance Definition _ := + Kernel_isSFinite_subdef.Build _ _ _ _ R kcounting sfkcounting. + +End kcounting. + +(* formalization of the iterate construct of Staton ESOP 2017, Sect. 4.2 *) +Section iterate. +Context d {G : measurableType d} {R : realType}. +Let A : measurableType _ := unit. +Let B : measurableType _ := bool. + +(* formalization of iterate^n + Gamma |-p iterate^n t from x = u : B *) +Variables (t : R.-sfker (G * A) ~> (A + B)%type) + (u : G -> A) (mu : measurable_fun setT u). + +Fixpoint iterate_ n : R.-sfker G ~> B := + match n with + | 0%N => case_sum (letin (ret mu) t) + (fun x => match x with + | inl a => fail + | inr b => ret (measurable_cst b) + end) + | m.+1 => case_sum (letin (ret mu) t) + (fun x => match x with + | inl a => iterate_ m + | inr b => fail + end) + end. + +(* formalization of iterate (A = unit, B = bool) + Gamma, x : A |-p t : A + B Gamma |-d u : A +----------------------------------------------- + Gamma |-p iterate t from x = u : B *) +Definition iterate : R.-sfker G ~> B := case_nat (kcounting R) iterate_. + +End iterate. + +(* an s-finite kernel to test that two expressions are different *) +Section lift_neq. +Context {R : realType} d (G : measurableType d). +Variables (f : G -> bool) (g : G -> bool). + +Definition flift_neq : G -> bool := fun x' => f x' != g x'. + +Hypotheses (mf : measurable_fun setT f) (mg : measurable_fun setT g). + +(* see also emeasurable_fun_neq *) +Lemma measurable_fun_flift_neq : measurable_fun setT flift_neq. +Proof. +apply: (measurable_fun_bool true). +rewrite setTI /flift_neq /= (_ : _ @^-1` _ = ([set x | f x] `&` [set x | ~~ g x]) `|` + ([set x | ~~ f x] `&` [set x | g x])). + apply: measurableU; apply: measurableI. + - by rewrite -[X in measurable X]setTI; exact: mf. + - rewrite [X in measurable X](_ : _ = ~` [set x | g x]); last first. + by apply/seteqP; split => x /= /negP. + by apply: measurableC; rewrite -[X in measurable X]setTI; exact: mg. + - rewrite [X in measurable X](_ : _ = ~` [set x | f x]); last first. + by apply/seteqP; split => x /= /negP. + by apply: measurableC; rewrite -[X in measurable X]setTI; exact: mf. + - by rewrite -[X in measurable X]setTI; exact: mg. +by apply/seteqP; split => x /=; move: (f x) (g x) => [|] [|]//=; intuition. +Qed. + +Definition lift_neq : R.-sfker G ~> bool := ret measurable_fun_flift_neq. + +End lift_neq. + +Section von_neumann_trick. +Context d {T : measurableType d} {R : realType}. + +Definition kinrtt {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} := + @measurable_cst _ _ T1 _ setT (@inl unit T2 tt). + +Definition finlb d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) + : T1 * bool -> T2 + bool := fun t1b => inr t1b.2. + +Lemma minlb {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} : + measurable_fun setT (@finlb _ _ T1 T2). +Proof. exact: measurableT_comp. Qed. + +Variable (D : pprobability bool R (* biased coin *)). + +Definition von_neumann_trick' : R.-sfker (T * unit) ~> (unit + bool)%type := + letin (sample_cst D) + (letin (sample_cst D) + (letin (lift_neq macc1of3 macc2of3) + (ite (macc3of4) + (letin (ret macc1of4) (ret minlb)) + (ret kinrtt)))). + +Definition von_neumann_trick : R.-sfker T ~> bool := + iterate von_neumann_trick' ktt. + +End von_neumann_trick. diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 7b12eed1b6..37d56a89ad 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -1,15 +1,15 @@ From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. -From mathcomp Require Import rat. -From mathcomp.classical Require Import mathcomp_extra boolp classical_sets. -From mathcomp.classical Require Import functions cardinality fsbigop. -Require Import signed reals ereal topology normedtype sequences esum measure. -Require Import lebesgue_measure numfun lebesgue_integral exp kernel trigo. -Require Import prob_lang. +From mathcomp Require Import rat interval_inference. +From mathcomp Require Import mathcomp_extra boolp classical_sets. +From mathcomp Require Import functions cardinality fsbigop. +From mathcomp Require Import reals ereal topology normedtype sequences esum measure. +From mathcomp Require Import lebesgue_measure numfun lebesgue_integral exp kernel trigo. +From mathcomp Require Import prob_lang. (******************************************************************************) (* Semantics of a probabilistic programming language using s-finite kernels *) -(* (wip) *) +(* (wip about definition of Lebesgue and counting measures) *) (******************************************************************************) Set Implicit Arguments. @@ -41,6 +41,9 @@ Qed. Definition gauss01_density : R -> R := gauss_density 0 1. +Hypothesis integral_gauss01_density : + (\int[lebesgue_measure]_x (gauss01_density x)%:E = 1%E)%E. + Lemma gauss01_densityE x : gauss01_density x = (sqrtr (pi *+ 2))^-1 * expR (- (x ^+ 2) / 2%:R). Proof. by rewrite /gauss01_density /gauss_density mul1r subr0 divr1. Qed. @@ -68,9 +71,6 @@ Proof. by rewrite /mgauss01 integral_ge0//= => x _; rewrite lee_fin gauss_density_ge0. Qed. -Axiom integral_gauss01_density : - (\int[lebesgue_measure]_x (gauss01_density x)%:E = 1%E)%E. - Let mgauss01_sigma_additive : semi_sigma_additive mgauss01. Proof. move=> /= F mF tF mUF. @@ -81,7 +81,7 @@ rewrite /mgauss01/= integral_bigcup//=; last first. rewrite (_ : (fun x => _) = (EFin \o gauss01_density)); last first. by apply/funext => x; rewrite gee0_abs// lee_fin gauss_density_ge0. apply: le_lt_trans. - apply: (@subset_integral _ _ _ _ _ setT) => //=. + apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. apply/EFin_measurable_fun. exact: measurable_fun_gauss_density. by move=> ? _; rewrite lee_fin gauss_density_ge0. @@ -105,8 +105,14 @@ End gauss. Section gauss_lebesgue. Import Notations. Context d (T : measurableType d) (R : realType). +Hypothesis integral_gauss01_density : + (\int[@lebesgue_measure R]_x (gauss01_density x)%:E = 1%E)%E. + +Let f1 (x : R) := ((gauss01_density x) ^-1)%R. -Let f1 (x : R) := (gauss01_density x) ^-1. +Hypothesis integral_mgauss01 : forall U, measurable U -> + \int[mgauss01 (R:=R)]_(y in U) (f1 y)%:E = + \int[lebesgue_measure]_(x0 in U) (gauss01_density x0 * f1 x0)%:E. Let mf1 : measurable_fun setT f1. Proof. @@ -121,10 +127,10 @@ Qed. Variable mu : {measure set mR R -> \bar R}. Definition staton_lebesgue : R.-sfker T ~> _ := - letin (sample (@gauss01 R)) + letin (sample_cst (gauss01 integral_gauss01_density : pprobability _ _)) (letin - (score (measurableT_comp mf1 var2of2)) - (ret var2of3)). + (score (measurableT_comp mf1 macc1of2)) + (ret macc1of3)). Lemma staton_lebesgueE x U : measurable U -> staton_lebesgue x U = lebesgue_measure U. @@ -133,16 +139,113 @@ move=> mU; rewrite [in LHS]/staton_lebesgue/=. rewrite [in LHS]letinE /=. transitivity (\int[@mgauss01 R]_(y in U) (f1 y)%:E). rewrite -[in RHS](setTI U) integral_setI_indic//=. - apply: eq_integral => /= r _. + apply: eq_integral => //= r. rewrite letinE/= ge0_integral_mscale//= ger0_norm//; last first. by rewrite invr_ge0// gauss_density_ge0. - by rewrite integral_dirac// indicT mul1e diracE indicE. -transitivity (\int[lebesgue_measure]_(x in U) (gauss01_density x * f1 x)%:E). - admit. + by rewrite integral_dirac// diracT mul1e diracE indicE. +rewrite integral_mgauss01//. transitivity (\int[lebesgue_measure]_(x in U) (\1_U x)%:E). apply: eq_integral => /= y yU. by rewrite /f1 divrr ?indicE ?yU// unitfE gt_eqF// gauss_density_gt0. by rewrite integral_indic//= setIid. -Abort. +Qed. End gauss_lebesgue. + +(* TODO: move this elsewhere *) +(* assuming x > 0 *) +Definition Gamma {R : realType} (x : R) : \bar R := + \int[lebesgue_measure]_(t in `[0%R, +oo[%classic) (expR (- t) * powR t (x - 1))%:E. + +Definition Rfact {R : realType} (x : R) := Gamma (x + 1)%R. + +Section poisson. +Variable R : realType. +Local Open Scope ring_scope. +Hypothesis integral_poisson_density : forall k, + (\int[lebesgue_measure]_x (@poisson R k x)%:E = 1%E)%E. + +(* density function for poisson *) +Definition poisson1 := @poisson R 1%N. + +Lemma poisson1_ge0 (x : R) : 0 <= poisson1 x. +Proof. exact: poisson_ge0. Qed. + +Definition mpoisson1 (V : set R) : \bar R := + (\int[lebesgue_measure]_(x in V) (poisson1 x)%:E)%E. + +Lemma measurable_fun_poisson1 : measurable_fun setT poisson1. +Proof. exact: measurable_poisson. Qed. + +Let mpoisson10 : mpoisson1 set0 = 0%E. +Proof. by rewrite /mpoisson1 integral_set0. Qed. + +Lemma mpoisson1_ge0 A : (0 <= mpoisson1 A)%E. +Proof. +apply: integral_ge0 => x Ax. +by rewrite lee_fin poisson1_ge0. +Qed. + +Let mpoisson1_sigma_additive : semi_sigma_additive mpoisson1. +Proof. +move=> /= F mF tF mUF. +rewrite /mpoisson1/= integral_bigcup//=; last first. + apply/integrableP; split. + apply/EFin_measurable_fun. + exact: measurable_funS (measurable_poisson _). + rewrite (_ : (fun x => _) = (EFin \o poisson1)); last first. + by apply/funext => x; rewrite gee0_abs// lee_fin poisson1_ge0//. + apply: le_lt_trans. + apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. + by apply/EFin_measurable_fun; exact: measurable_poisson. + by move=> ? _; rewrite lee_fin poisson1_ge0//. + by rewrite /= integral_poisson_density// ltry. +apply: is_cvg_ereal_nneg_natsum_cond => n _ _. +by apply: integral_ge0 => /= x ?; rewrite lee_fin poisson1_ge0. +Qed. + +HB.instance Definition _ := isMeasure.Build _ _ _ + mpoisson1 mpoisson10 mpoisson1_ge0 mpoisson1_sigma_additive. + +Let mpoisson1_setT : mpoisson1 [set: _] = 1%E. +Proof. +rewrite /mpoisson1. +rewrite /poisson1. +by rewrite integral_poisson_density. +Qed. + +HB.instance Definition _ := @Measure_isProbability.Build _ _ R mpoisson1 mpoisson1_setT. + +Definition poisson' := [the probability _ _ of mpoisson1]. + +End poisson. + +(* Staton's definition of the counting measure + Staton ESOP 2017, Sect. 4.2, equation (13) *) +Section staton_counting. +Context d (X : measurableType d). +Variable R : realType. +Import Notations. +Hypothesis integral_poisson_density : forall k, + (\int[lebesgue_measure]_x (@poisson R k x)%:E = 1%E)%E. + +Let f1 x := ((poisson1 (x : R)) ^-1)%R. + +Let mf1 : measurable_fun setT f1. +rewrite /f1 /poisson1 /poisson. +apply: (measurable_comp (F := [set r : R | r != 0%R])) => //. +- exact: open_measurable. +- move=> /= r [t ? <-]. + by case: ifPn => // t0; rewrite gt_eqF ?mulr_gt0 ?expR_gt0//= invrK ltr0n. +- apply: open_continuous_measurable_fun => //. + by apply/in_setP => x /= x0; exact: inv_continuous. +- exact: measurable_poisson. +Qed. + +Definition staton_counting : R.-sfker X ~> _ := + letin (sample_cst (@poisson' R integral_poisson_density : pprobability _ _)) + (letin + (score (measurableT_comp mf1 macc1of2)) + (ret macc1of3)). + +End staton_counting. From 801316a3f91cc9e856621b2835f35ab6e4318b7c Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Sat, 16 Sep 2023 15:19:41 +0900 Subject: [PATCH 03/48] rm duplicate, more uniform naming --- theories/lang_syntax.v | 38 +++++----- theories/lang_syntax_examples.v | 100 ++++++++++++++++++--------- theories/prob_lang.v | 118 ++++++++++++++++---------------- theories/prob_lang_wip.v | 23 ++++--- 4 files changed, 155 insertions(+), 124 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 253cbe837b..8601c576fa 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -130,7 +130,7 @@ by rewrite /mswap/= kE. Qed. HB.instance Definition _ := - Kernel_isSFinite_subdef.Build _ _ _ Z R (mkswap k) mkswap_sfinite. + isSFiniteKernel_subdef.Build _ _ _ Z R (mkswap k) mkswap_sfinite. End mswap_sfinite_kernel. @@ -220,7 +220,7 @@ under eq_integral. under eq_integral do rewrite retE /=. over. rewrite (sfinite_Fubini (T' z) (U' z) (fun x => \d_(x.1, x.2) A ))//; last first. - apply/EFin_measurable_fun => /=; rewrite (_ : (fun x => _) = mindic R mA)//. + apply/measurable_EFinP => /=; rewrite (_ : (fun x => _) = mindic R mA)//. by apply/funext => -[]. rewrite /=. apply: eq_integral => y _. @@ -449,12 +449,6 @@ Notation "'Normalize' e" := (exp_normalize e) Notation "'if' e1 'then' e2 'else' e3" := (exp_if e1 e2 e3) (in custom expr at level 1) : lang_scope. -Local Open Scope lang_scope. -Example three_letin {R : realType} x y z : @exp R P [::] _ := - [let x := return {1}:R in - let y := return #x in - let z := return #y in return #z]. - Section free_vars. Context {R : realType}. @@ -548,7 +542,7 @@ by move=> z U mU; by rewrite /kweak/= hs. Qed. HB.instance Definition _ := - Kernel_isSFinite_subdef.Build _ _ _ _ _ (@kweak g h x t f) sf. + isSFiniteKernel_subdef.Build _ _ _ _ _ (@kweak g h x t f) sf. End sfkernel_weak. @@ -609,7 +603,7 @@ Inductive evalD : forall g t, exp D g t -> | eval_normalize g t (e : exp P g t) k : e -P> k -> - [Normalize e] -D> normalize k point ; measurable_fun_mnormalize k + [Normalize e] -D> normalize_pt k ; measurable_normalize_pt k | evalD_if g t e f mf (e1 : exp D g t) f1 mf1 e2 f2 mf2 : e -D> f ; mf -> e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> @@ -630,8 +624,8 @@ with evalP : forall g t, exp P g t -> pval R g t -> Prop := [let str := e1 in e2] -P> letin' k1 k2 | eval_sample g t (e : exp _ _ (Prob t)) - (p : mctx g -> pprobability (mtyp t) R) mp : - e -D> p ; mp -> [Sample e] -P> sample p mp + (f : mctx g -> probability (mtyp t) R) mf : + e -D> f ; mf -> [Sample e] -P> sample f mf | eval_score g (e : exp _ g _) f mf : e -D> f ; mf -> [Score e] -P> kscore mf @@ -750,13 +744,13 @@ all: (rewrite {g t e u v mu mv hu}). inj_ex H6; subst e5. inj_ex H5; subst e4. by rewrite (IH1 _ H4) (IH2 _ H8). -- move=> g t e p mp ev IH k. +- move=> g t e f mf ev IH k. inversion 1; subst g0. inj_ex H5; subst t0. inj_ex H5; subst e1. inj_ex H7; subst k. - have ? := IH _ _ H3; subst p1. - by have -> : mp = mp1 by []. + have ? := IH _ _ H3; subst f1. + by have -> : mf = mf1 by []. - move=> g e f mf ev IH k. inversion 1; subst g0. inj_ex H0; subst e0. @@ -875,12 +869,12 @@ all: rewrite {g t e u v eu}. inj_ex H5; subst e4. inj_ex H6; subst e5. by rewrite (IH1 _ H4) (IH2 _ H8). -- move=> g t e p mp ep IH v. +- move=> g t e f mf ep IH v. inversion 1; subst g0 t0. inj_ex H7; subst v. inj_ex H5; subst e1. - have ? := IH _ _ H3; subst p1. - by have -> : mp = mp1 by []. + have ? := IH _ _ H3; subst f1. + by have -> : mf = mf1 by []. - move=> g e f mf ev IH k. inversion 1; subst g0. inj_ex H0; subst e0. @@ -931,7 +925,7 @@ all: rewrite {z g t}. - move=> g h e [f [mf H]]. by exists (poisson h \o f); eexists; exact: eval_poisson. - move=> g t e [k ek]. - by exists (normalize k point); eexists; exact: eval_normalize. + by exists (normalize_pt k); eexists; exact: eval_normalize. - move=> g t1 t2 x e1 [k1 ev1] e2 [k2 ev2]. by exists (letin' k1 k2); exact: eval_letin. - move=> g t e [f [/= mf ef]]. @@ -1072,10 +1066,10 @@ Lemma execD_bernoulli g r (r1 : (r%:num <= 1)%R) : existT _ (cst [the probability _ _ of bernoulli r1]) (measurable_cst _). Proof. exact/execD_evalD/eval_bernoulli. Qed. -Lemma execD_normalize g t (e : exp P g t) : +Lemma execD_normalize_pt g t (e : exp P g t) : @execD g _ [Normalize e] = - existT _ (normalize (execP e) point : _ -> pprobability _ _) - (measurable_fun_mnormalize (execP e)). + existT _ (normalize_pt (execP e) : _ -> pprobability _ _) + (measurable_normalize_pt (execP e)). Proof. exact/execD_evalD/eval_normalize/evalP_execP. Qed. Lemma execD_poisson g n (e : exp D g Real) : diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index 87424ac899..e02fec2395 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -5,7 +5,7 @@ From mathcomp Require Import interval_inference. From mathcomp Require Import mathcomp_extra boolp classical_sets. From mathcomp Require Import functions cardinality fsbigop. From mathcomp Require Import reals ereal topology normedtype sequences esum. -From mathcomp Require Import measure lebesgue_measure numfun lebesgue_integral. +From mathcomp Require Import measure lebesgue_measure numfun lebesgue_integral. From mathcomp Require Import kernel prob_lang lang_syntax_util lang_syntax. From mathcomp Require Import ring lra. @@ -139,19 +139,19 @@ Qed. Local Close Scope lang_scope. (* simple tests to check bidirectional hints *) -Module tests_bidi. -Section tests_bidi. +Module bidi_tests. +Section bidi_tests. Local Open Scope lang_scope. Import Notations. Context (R : realType). -Definition v1 x : @exp R P [::] _ := [ +Definition bidi_test1 x : @exp R P [::] _ := [ let x := return {1}:R in return #x]. -Definition v2 (a b : string) +Definition bidi_test2 (a b : string) (a := "a") (b := "b") - (* (H : infer (b != a)) *) + (* (ba : infer (b != a)) *) : @exp R P [::] _ := [ let a := return {1}:R in let b := return {true}:B in @@ -159,17 +159,28 @@ Definition v2 (a b : string) let d := return {4}:R in *) return (#a, #b)]. -Definition v3 (a b c d : string) (H1 : infer (b != a)) (H2 : infer (c != a)) - (H3 : infer (c != b)) (H4 : infer (a != b)) (H5 : infer (a != c)) - (H6 : infer (b != c)) : @exp R P [::] _ := [ +Definition bidi_test3 (a b c d : string) + (ba : infer (b != a)) (ca : infer (c != a)) + (cb : infer (c != b)) (ab : infer (a != b)) + (ac : infer (a != c)) (bc : infer (b != c)) : @exp R P [::] _ := [ let a := return {1}:R in let b := return {2}:R in let c := return {3}:R in (* let d := return {4}:R in *) return (#b, #a)]. -End tests_bidi. -End tests_bidi. +Definition bidi_test4 (a b c d : string) + (ba : infer (b != a)) (ca : infer (c != a)) + (cb : infer (c != b)) (ab : infer (a != b)) + (ac : infer (a != c)) (bc : infer (b != c)) : @exp R P [::] _ := [ + let a := return {1}:R in + let b := return {2}:R in + let c := return {3}:R in + (* let d := return {4}:R in *) + return {exp_poisson O [#c(*{exp_var c erefl}*)]}]. + +End bidi_tests. +End bidi_tests. Section trivial_example. Local Open Scope lang_scope. @@ -181,7 +192,8 @@ Lemma exec_normalize_return g x r : @dirac _ (measurableTypeR R) r _ :> probability _ R. (* NB: \d_r notation? *) Proof. -by rewrite execD_normalize execP_return execD_real//=; exact: normalize_kdirac. +rewrite execD_normalize_pt execP_return execD_real/=. +exact: normalize_kdirac. Qed. End trivial_example. @@ -241,7 +253,7 @@ Qed. Lemma exec_sample_pair_TorT : (projT1 (execD sample_pair_syntax)) tt [set p | p.1 || p.2] = (2 / 3)%:E. Proof. -rewrite execD_normalize normalizeE/= exec_sample_pair0. +rewrite execD_normalize_pt normalizeE/= exec_sample_pair0. do 4 rewrite mem_set//=. rewrite eqe ifF; last by apply/negbTE/negP => /orP[/eqP|//]; lra. rewrite exec_sample_pair0; do 3 rewrite mem_set//; rewrite memNset//=. @@ -265,7 +277,7 @@ Lemma exec_bernoulli13_score : execD bernoulli13_score = execD (exp_bernoulli (1 / 5%:R)%:nng (p1S 4)). Proof. apply: eq_execD. -rewrite execD_bernoulli/= /bernoulli13_score execD_normalize 2!execP_letin. +rewrite execD_bernoulli/= /bernoulli13_score execD_normalize_pt 2!execP_letin. rewrite execP_sample/= execD_bernoulli/= execP_if /= exp_var'E. rewrite (execD_var_erefl "x")/= !execP_return/= 2!execP_score 2!execD_real/=. apply: funext=> g; apply: eq_probability => U. @@ -290,14 +302,12 @@ rewrite -EFinD/= eqe ifF; last first. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. -rewrite !integral_dirac//= !diracT/= !mul1e ger0_norm//. +rewrite !integral_dirac//= !diracT /= !mul1e ger0_norm// !indicT. rewrite exp_var'E (execD_var_erefl "x")/=. rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. -rewrite indicT/= !mulr1. by rewrite muleDl//; congr (_ + _)%E; - rewrite -!EFinM; - congr (_%:E); - rewrite indicE /onem; case: (_ \in _); field. + rewrite -!EFinM; congr (_%:E); + rewrite !indicE /onem /=; case: (_ \in _); field. Qed. Definition bernoulli12_score := [Normalize @@ -309,7 +319,7 @@ Lemma exec_bernoulli12_score : execD bernoulli12_score = execD (exp_bernoulli (1 / 3%:R)%:nng (p1S 2)). Proof. apply: eq_execD. -rewrite execD_bernoulli/= /bernoulli12_score execD_normalize 2!execP_letin. +rewrite execD_bernoulli/= /bernoulli12_score execD_normalize_pt 2!execP_letin. rewrite execP_sample/= execD_bernoulli/= execP_if /= exp_var'E. rewrite (execD_var_erefl "x")/= !execP_return/= 2!execP_score 2!execD_real/=. apply: funext=> g; apply: eq_probability => U. @@ -327,7 +337,7 @@ rewrite !integral_dirac// !diracT !mul1e. rewrite iteE/= !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_indic//= !iteE/= /mscale/=. -rewrite setTI diracE !in_setT !mule1. +rewrite setTI !diracT !mule1. rewrite ger0_norm//. rewrite -EFinD/= eqe ifF; last first. apply/negbTE/negP => /orP[/eqP|//]. @@ -335,8 +345,7 @@ rewrite -EFinD/= eqe ifF; last first. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. -rewrite !integral_dirac//= !diracT /= !mul1e ger0_norm//. -rewrite !indicT/= mulr1. +rewrite !integral_dirac//= !diracT /= !mul1e ger0_norm// !indicT. rewrite exp_var'E (execD_var_erefl "x")/=. rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. by rewrite muleDl//; congr (_ + _)%E; @@ -358,9 +367,9 @@ Lemma exec_bernoulli14_score : execD bernoulli14_score = execD (exp_bernoulli (5%:R / 11%:R)%:nng p511). Proof. apply: eq_execD. -rewrite execD_bernoulli/= execD_normalize 2!execP_letin. +rewrite execD_bernoulli/= execD_normalize_pt 2!execP_letin. rewrite execP_sample/= execD_bernoulli/= execP_if /= !exp_var'E. -rewrite !execP_return/= 2!execP_score 2!execD_real/=. +rewrite !execP_return/= 2!execP_score !execD_real/=. rewrite !(execD_var_erefl "x")/=. apply: funext=> g; apply: eq_probability => U. rewrite normalizeE !letin'E/=. @@ -516,12 +525,43 @@ Local Open Scope lang_scope. Import Notations. Context {R : realType}. +Section tests. + +Local Notation "$ str" := (@exp_var _ _ str%string _ erefl) + (in custom expr at level 1, format "$ str"). + +Definition staton_bus_syntax0_generic (x r u : string) + (rx : infer (r != x)) (Rx : infer (u != x)) + (ur : infer (u != r)) (xr : infer (x != r)) + (xu : infer (x != u)) (ru : infer (r != u)) : @exp R P [::] _ := + [let x := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in + let r := if #x then return {3}:R else return {10}:R in + let u := Score {exp_poisson 4 [#r]} in + return #x]. + +Fail Definition staton_bus_syntax0_generic' (x r u : string) + (rx : infer (r != x)) (Rx : infer (u != x)) + (ur : infer (u != r)) (xr : infer (x != r)) + (xu : infer (x != u)) (ru : infer (r != u)) : @exp R P [::] _ := + [let x := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in + let r := if $x then return {3}:R else return {10}:R in + let u := Score {exp_poisson 4 [$r]} in + return $x]. + +Fail Definition staton_bus_syntax0' : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in + let "r" := if ${"x"} then return {3}:R else return {10}:R in + let "_" := Score {exp_poisson 4 [${"r"}]} in + return ${"x"}]. + Definition staton_bus_syntax0 : @exp R _ [::] _ := [let "x" := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in let "r" := if #{"x"} then return {3}:R else return {10}:R in let "_" := Score {exp_poisson 4 [#{"r"}]} in return #{"x"}]. +End tests. + Definition staton_bus_syntax := [Normalize {staton_bus_syntax0}]. Let sample_bern : R.-sfker munit ~> mbool := sample_cst (bernoulli p27). @@ -568,8 +608,8 @@ by rewrite exp_var'E (execD_var_erefl "x") /=; congr ret. Qed. Lemma exec_staton_bus : execD staton_bus_syntax = - existT _ (normalize kstaton_bus' point) (measurable_fun_mnormalize _). -Proof. by rewrite execD_normalize exec_staton_bus0'. Qed. + existT _ (normalize_pt kstaton_bus') (measurable_normalize_pt _). +Proof. by rewrite execD_normalize_pt exec_staton_bus0'. Qed. Let poisson4 := @poisson R 4%N. @@ -672,8 +712,8 @@ by rewrite exp_var'E (execD_var_erefl "x") /=; congr ret. Qed. Lemma exec_statonA_bus : execD staton_busA_syntax = - existT _ (normalize kstaton_busA' point) (measurable_fun_mnormalize _). -Proof. by rewrite execD_normalize exec_staton_busA0'. Qed. + existT _ (normalize_pt kstaton_busA') (measurable_normalize_pt _). +Proof. by rewrite execD_normalize_pt exec_staton_busA0'. Qed. (* equivalence between staton_bus and staton_busA *) Lemma staton_bus_staton_busA : @@ -719,8 +759,6 @@ Section letinC. Local Open Scope lang_scope. Variable (R : realType). -Require Import Classical_Prop. - Lemma letinC g t1 t2 (e1 : @exp R P g t1) (e2 : exp P g t2) (x y : string) (xy : infer (x != y)) (yx : infer (y != x)) diff --git a/theories/prob_lang.v b/theories/prob_lang.v index c9119b43ab..bf077591d6 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -4,8 +4,10 @@ From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. From mathcomp Require Import interval_inference archimedean rat. From mathcomp Require Import mathcomp_extra boolp classical_sets. From mathcomp Require Import functions cardinality fsbigop. -From mathcomp Require Import reals ereal topology normedtype sequences esum measure. -From mathcomp Require Import lebesgue_measure numfun lebesgue_integral exp kernel. +From mathcomp Require Import reals ereal topology normedtype sequences. +From mathcomp Require Import esum measure lebesgue_measure numfun. +From mathcomp Require Import lebesgue_integral exp kernel. +From mathcomp Require Import lra. (******************************************************************************) (* Semantics of a probabilistic programming language using s-finite kernels *) @@ -13,7 +15,8 @@ From mathcomp Require Import lebesgue_measure numfun lebesgue_integral exp kerne (* bernoulli r1 == Bernoulli probability with r1 a proof that *) (* r : {nonneg R} is smaller than 1 *) (* uniform_probability a b ab0 == uniform probability over the interval [a,b] *) -(* sample_cst P == sample according to the probability P *) +(* sample mP == sample according to the probability P where mP is a *) +(* proof that P is a measurable function *) (* letin l k == execute l, augment the context, and execute k *) (* ret mf == access the context with f and return the result *) (* score mf == observe t from d, where f is the density of d and *) @@ -115,7 +118,7 @@ Lemma integral_bernoulli {R : realType} Proof. move=> f0. rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. -by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT !mul1e. +by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT 2!mul1e. Qed. Section uniform_probability. @@ -228,7 +231,7 @@ rewrite (_ : (fun x => _) = (fun x => x * (\1_(`[i%:R%:E, i.+1%:R%:E [%classic : set _) x)%:E)); last first. apply/funext => x; case: ifPn => ix; first by rewrite indicE/= mem_set ?mule1. by rewrite indicE/= memNset ?mule0// /= in_itv/=; exact/negP. -apply: emeasurable_funM => //=; apply/EFin_measurable_fun. +apply: emeasurable_funM => //=; apply/measurable_EFinP. by rewrite (_ : \1__ = mindic R (emeasurable_itv `[(i%:R)%:E, (i.+1%:R)%:E[)). Qed. @@ -291,7 +294,7 @@ rewrite big1 ?adde0//= => j jk. rewrite ifF// lte_fin lee_fin. move: jk; rewrite neq_ltn/= => /orP[|] jr. - suff : (j.+1%:R <= `|f t|)%R by rewrite leNgt => /negbTE ->; rewrite andbF. - rewrite (_ : j.+1%:R = j.+1%:~R)// floor_ge_int. + rewrite (_ : j.+1%:R = j.+1%:~R)// floor_ge_int//. move: jr; rewrite -lez_nat => /le_trans; apply. by rewrite -[leRHS](@ger0_norm _ (floor `|f t|)) ?floor_ge0. - suff : (`|f t| < j%:R)%R by rewrite ltNge => /negbTE ->. @@ -346,7 +349,7 @@ by rewrite /mseries eseries0. Qed. #[export] -HB.instance Definition _ := @Kernel_isSFinite_subdef.Build _ _ _ _ _ +HB.instance Definition _ := @isSFiniteKernel_subdef.Build _ _ _ _ _ (kiteT k) sfinite_kiteT. End sfkiteT. @@ -403,7 +406,7 @@ by rewrite /mseries eseries0. Qed. #[export] -HB.instance Definition _ := @Kernel_isSFinite_subdef.Build _ _ _ _ _ +HB.instance Definition _ := @isSFiniteKernel_subdef.Build _ _ _ _ _ (kiteF k) sfinite_kiteF. End sfkiteF. @@ -469,14 +472,28 @@ Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). Definition ret (f : X -> Y) (mf : measurable_fun setT f) : R.-pker X ~> Y := [the R.-pker _ ~> _ of kdirac mf]. +Definition sample (P : X -> pprobability Y R) (mP : measurable_fun setT P) + : R.-pker X ~> Y := + [the R.-pker _ ~> _ of kprobability mP]. + Definition sample_cst (P : pprobability Y R) : R.-pker X ~> Y := - [the R.-pker _ ~> _ of kprobability (measurable_cst P)]. + sample (measurable_cst P). -Definition sample (P : X -> pprobability Y R) (mP : measurable_fun setT P) : R.-pker X ~> Y := - [the R.-pker _ ~> _ of kprobability mP]. +Definition normalize (k : R.-ker X ~> Y) P : X -> probability Y R := + knormalize k P. + +Definition normalize_pt (k : R.-ker X ~> Y) : X -> probability Y R := + normalize k point. -Definition normalize (k : R.-sfker X ~> Y) P : X -> probability Y R := - fun x => [the probability _ _ of mnormalize (k x) P]. +Lemma measurable_normalize_pt (f : R.-ker X ~> Y) : + measurable_fun [set: X] (normalize_pt f : X -> pprobability Y R). +Proof. +apply: (@measurability _ _ _ _ _ _ + (@pset _ _ _ : set (set (pprobability Y R)))) => //. +move=> _ -[_ [r r01] [Ys mYs <-]] <-. +apply: emeasurable_fun_infty_o => //. +exact: (measurable_kernel (knormalize f point) Ys). +Qed. Definition ite (f : X -> bool) (mf : measurable_fun setT f) (k1 k2 : R.-sfker X ~> Y) : R.-sfker X ~> Y := @@ -567,10 +584,9 @@ rewrite integral_indic ?setIT// -[X in measurable X]setTI. exact: (measurableT_comp mf). Qed. -Lemma letin_retk - (f : X -> Y) (mf : measurable_fun setT f) - (k : R.-sfker [the measurableType _ of (X * Y)%type] ~> Z) - x U : measurable U -> +Lemma letin_retk (f : X -> Y) + (mf : measurable_fun [set: X] f) (k : R.-sfker X * Y ~> Z) x U : + measurable U -> letin (ret mf) k x U = k (x, f x) U. Proof. move=> mU; rewrite letinE retE integral_dirac ?diracT ?mul1e//. @@ -910,7 +926,7 @@ Lemma score_fail (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : Proof. apply/eq_sfkernel => x U. rewrite letinE/= /sample; unlock. -rewrite ge0_integral_measure_add//= ge0_integral_mscale//= ge0_integral_mscale//=. +rewrite ge0_integral_measure_add// ge0_integral_mscale//= ge0_integral_mscale//=. rewrite integral_dirac//= integral_dirac//= !diracT/= !mul1e. by rewrite /mscale/= iteE//= iteE//= failE mule0 adde0 ger0_norm. Qed. @@ -993,8 +1009,7 @@ HB.instance Definition _ z := @isMeasure.Build _ X R (T z) (T0 z) (T_ge0 z) (@T_semi_sigma_additive z). Let sfinT z : sfinite_measure (T z). Proof. exact: sfinite_kernel_measure. Qed. -HB.instance Definition _ z := @isSFinite.Build _ X R - (T z) (sfinT z). +HB.instance Definition _ z := @isSFinite.Build _ X R (T z) (sfinT z). Definition U z : set Y -> \bar R := u z. Let U0 z : (U z) set0 = 0. Proof. by []. Qed. @@ -1027,7 +1042,7 @@ rewrite (sfinite_Fubini [the {sfinite_measure set X -> \bar R} of T z] [the {sfinite_measure set Y -> \bar R} of U z] (fun x => \d_(x.1, x.2) A ))//; last first. - apply/EFin_measurable_fun => /=; rewrite (_ : (fun x => _) = mindic R mA)//. + apply/measurable_EFinP => /=; rewrite (_ : (fun x => _) = mindic R mA)//. by apply/funext => -[]. rewrite /=. apply: eq_integral => y _. @@ -1089,9 +1104,7 @@ Qed. Lemma measurable_poisson k : measurable_fun setT (poisson k). Proof. rewrite /poisson; apply: measurable_fun_if => //. - apply: (measurable_fun_bool true) => /=. - rewrite setTI (_ : _ @^-1` _ = `]0, +oo[%classic)//. - by apply/seteqP; split => x /=; rewrite in_itv/= andbT. + exact: measurable_fun_ltr. by apply: measurable_funM => /=; [exact: measurable_funM|exact: measurableT_comp]. Qed. @@ -1168,7 +1181,7 @@ Lemma sample_and_branchE t U : sample_and_branch t U = (2 / 7%:R)%:E * \d_(3%:R : R) U + (5%:R / 7%:R)%:E * \d_(10%:R : R) U. Proof. -by rewrite /sample_and_branch letin_sample_bernoulli/= !iteE !retE// onem27. +by rewrite /sample_and_branch letin_sample_bernoulli//= !iteE/= onem27. Qed. End sample_and_branch. @@ -1177,32 +1190,16 @@ Section bernoulli_and. Context d (T : measurableType d) (R : realType). Import Notations. -Definition mand (x y : T * mbool * mbool -> mbool) - (t : T * mbool * mbool) : mbool := x t && y t. - -Lemma measurable_fun_mand (x y : T * mbool * mbool -> mbool) : - measurable_fun setT x -> measurable_fun setT y -> - measurable_fun setT (mand x y). -Proof. -move=> /= mx my; apply: (measurable_fun_bool true) => /=. -rewrite setTI [X in measurable X](_ : _ = - (x @^-1` [set true]) `&` (y @^-1` [set true])); last first. - by rewrite /mand; apply/seteqP; split => z/= /andP. -apply: measurableI. -- by rewrite -[X in measurable X]setTI; exact: mx. -- by rewrite -[X in measurable X]setTI; exact: my. -Qed. - Definition bernoulli_and : R.-sfker T ~> mbool := (letin (sample_cst [the probability _ _ of bernoulli p12]) (letin (sample_cst [the probability _ _ of bernoulli p12]) - (ret (measurable_fun_mand macc1of3 macc2of3)))). + (ret (measurable_and macc1of3 macc2of3)))). Lemma bernoulli_andE t U : bernoulli_and t U = sample_cst (bernoulli p14) t U. Proof. -rewrite /bernoulli_and 3!letin_sample_bernoulli/= /mand/= muleDr//= -muleDl//. +rewrite /bernoulli_and 3!letin_sample_bernoulli/= muleDr//= -muleDl//. rewrite !muleA -addeA -muleDl// -!EFinM !onem1S/= -splitr mulr1. have -> : (1 / 2 * (1 / 2) = 1 / 4%:R :> R)%R by rewrite mulf_div mulr1// -natrM. rewrite /bernoulli/= measure_addE/= /mscale/= -!EFinM; congr( _ + (_ * _)%:E). @@ -1419,7 +1416,7 @@ by rewrite /mseries eseries0. Qed. #[export] -HB.instance Definition _ j := @Kernel_isSFinite_subdef.Build _ _ _ _ _ +HB.instance Definition _ j := @isSFiniteKernel_subdef.Build _ _ _ _ _ (case_nat_ k j) (sfcase_nat_ j). End sfcase_nat. @@ -1465,8 +1462,8 @@ Proof. exact. Qed. Definition image_classes d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) (T : Type) (f1 : T1 -> T) (f2 : T2 -> T) := - <>. + <>. Section sum_salgebra_instance. Context d1 d2 (T1 : measurableType d1) (T2 : measurableType d2). @@ -1487,7 +1484,7 @@ Proof. exact: sigma_algebra_bigcup. Qed. HB.instance Definition sum_salgebra_mixin := @isMeasurable.Build (measure_sum_display (d1, d2)) (T1 + T2)%type (image_classes f1 f2) - (sum_salgebra_set0) (sum_salgebra_setC) (sum_salgebra_bigcup). + sum_salgebra_set0 sum_salgebra_setC sum_salgebra_bigcup. End sum_salgebra_instance. Reserved Notation "p .-sum" (at level 1, format "p .-sum"). @@ -1540,7 +1537,7 @@ by rewrite /mseries hk. Qed. #[export] -HB.instance Definition _ (a : A) := @Kernel_isSFinite_subdef.Build _ _ _ _ _ +HB.instance Definition _ (a : A) := @isSFiniteKernel_subdef.Build _ _ _ _ _ (case_suml k a) (sfinite_case_suml a). End sfkcase_suml. @@ -1610,7 +1607,7 @@ by rewrite /mseries/= eseries0. Qed. #[export] -HB.instance Definition _ b := @Kernel_isSFinite_subdef.Build _ _ _ _ _ +HB.instance Definition _ b := @isSFiniteKernel_subdef.Build _ _ _ _ _ (case_sumr k b) (sfinite_case_sumr b). End sfkcase_sumr. @@ -1685,7 +1682,7 @@ have mh1 : measurable h1. move=> _ /= C mC; rewrite setTI. have := H measurableT _ mC; rewrite setTI => {}H. rewrite [X in measurable X](_ : _ = ((fun x => k (inl tt) x U) @^-1` C) `*` setT)//. - exact: measurableM. + exact: measurableX. by apply/seteqP; split => [z//=| z/= []]. set h2 := [set xub : X * (unit + bool)| k (inr false) xub.1 U < x%:E]. have mh2 : measurable h2. @@ -1695,7 +1692,7 @@ have mh2 : measurable h2. move=> _ /= C mC; rewrite setTI. have := H measurableT _ mC; rewrite setTI => {}H. rewrite [X in measurable X](_ : _ = ((fun x => k (inr false) x U) @^-1` C) `*` setT)//. - exact: measurableM. + exact: measurableX. by apply/seteqP; split => [z //=|z/= []]. set h3 := [set xub : X * (unit + bool)| k (inr true) xub.1 U < x%:E]. have mh3 : measurable h3. @@ -1705,19 +1702,19 @@ have mh3 : measurable h3. move=> _ /= C mC; rewrite setTI. have := H measurableT _ mC; rewrite setTI => {}H. rewrite [X in measurable X](_ : _ = ((fun x => k (inr true) x U) @^-1` C) `*` setT)//. - exact: measurableM. + exact: measurableX. by apply/seteqP; split=> [z//=|z/= []]. apply: measurableU. - apply: measurableU. - + apply: measurableM => //. + + apply: measurableX => //. rewrite [X in measurable X](_ : _ = ysection h1 (inl tt))//. * by apply: measurable_ysection. * by apply/seteqP; split => z /=; rewrite /ysection /= inE. - + apply: measurableM => //. + + apply: measurableX => //. rewrite [X in measurable X](_ : _ = ysection h2 (inr false))//. * by apply: measurable_ysection. * by apply/seteqP; split => z /=; rewrite /ysection /= inE. -- apply: measurableM => //. +- apply: measurableX => //. rewrite [X in measurable X](_ : _ = ysection h3 (inr true))//. + by apply: measurable_ysection. + by apply/seteqP; split => z /=; rewrite /ysection /= inE. @@ -1759,8 +1756,8 @@ by rewrite (Hf (inr false) x _ mU). Qed. #[export] -HB.instance Definition _ := @Kernel_isSFinite_subdef.Build _ _ _ _ _ - (case_sum' k) (sfinite_case_sum'). +HB.instance Definition _ := @isSFiniteKernel_subdef.Build _ _ _ _ _ + (case_sum' k) sfinite_case_sum'. End sfkcase_sum'. End case_sum'. @@ -1802,7 +1799,7 @@ by move=> g U mU; rewrite /kcounting/= counting_dirac. Qed. HB.instance Definition _ := - Kernel_isSFinite_subdef.Build _ _ _ _ R kcounting sfkcounting. + isSFiniteKernel_subdef.Build _ _ _ _ R kcounting sfkcounting. End kcounting. @@ -1851,8 +1848,9 @@ Hypotheses (mf : measurable_fun setT f) (mg : measurable_fun setT g). (* see also emeasurable_fun_neq *) Lemma measurable_fun_flift_neq : measurable_fun setT flift_neq. Proof. -apply: (measurable_fun_bool true). -rewrite setTI /flift_neq /= (_ : _ @^-1` _ = ([set x | f x] `&` [set x | ~~ g x]) `|` +apply: (@measurable_fun_bool _ _ _ _ true). +rewrite setTI. +rewrite /flift_neq /= (_ : _ @^-1` _ = ([set x | f x] `&` [set x | ~~ g x]) `|` ([set x | ~~ f x] `&` [set x | g x])). apply: measurableU; apply: measurableI. - by rewrite -[X in measurable X]setTI; exact: mf. diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 37d56a89ad..1ac606f0a1 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -3,9 +3,9 @@ From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. From mathcomp Require Import rat interval_inference. From mathcomp Require Import mathcomp_extra boolp classical_sets. From mathcomp Require Import functions cardinality fsbigop. -From mathcomp Require Import reals ereal topology normedtype sequences esum measure. -From mathcomp Require Import lebesgue_measure numfun lebesgue_integral exp kernel trigo. -From mathcomp Require Import prob_lang. +From mathcomp Require Import reals ereal topology normedtype sequences. +From mathcomp Require Import esum measure lebesgue_measure numfun. +From mathcomp Require Import lebesgue_integral exp kernel trigo prob_lang. (******************************************************************************) (* Semantics of a probabilistic programming language using s-finite kernels *) @@ -58,7 +58,7 @@ apply: measurable_funM => //=. apply: measurableT_comp => //=. apply: measurable_funM => //=. apply: measurableT_comp => //=. -apply: measurableT_comp (measurable_exprn _) _ => /=. +apply: measurableT_comp (exprn_measurable _) _ => /=. apply: measurable_funM => //=. exact: measurable_funD. Qed. @@ -76,13 +76,13 @@ Proof. move=> /= F mF tF mUF. rewrite /mgauss01/= integral_bigcup//=; last first. apply/integrableP; split. - apply/EFin_measurable_fun. + apply/measurable_EFinP. exact: measurable_funS (measurable_fun_gauss_density 0 1). rewrite (_ : (fun x => _) = (EFin \o gauss01_density)); last first. by apply/funext => x; rewrite gee0_abs// lee_fin gauss_density_ge0. apply: le_lt_trans. apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. - apply/EFin_measurable_fun. + apply/measurable_EFinP. exact: measurable_fun_gauss_density. by move=> ? _; rewrite lee_fin gauss_density_ge0. by rewrite integral_gauss01_density// ltey. @@ -138,11 +138,12 @@ Proof. move=> mU; rewrite [in LHS]/staton_lebesgue/=. rewrite [in LHS]letinE /=. transitivity (\int[@mgauss01 R]_(y in U) (f1 y)%:E). - rewrite -[in RHS](setTI U) integral_setI_indic//=. - apply: eq_integral => //= r. + rewrite -[in RHS](setTI U) integral_mkcondr/=. + apply: eq_integral => //= r _. rewrite letinE/= ge0_integral_mscale//= ger0_norm//; last first. by rewrite invr_ge0// gauss_density_ge0. - by rewrite integral_dirac// diracT mul1e diracE indicE. + rewrite integral_dirac// diracT mul1e/= diracE epatch_indic/=. + by rewrite indicE. rewrite integral_mgauss01//. transitivity (\int[lebesgue_measure]_(x in U) (\1_U x)%:E). apply: eq_integral => /= y yU. @@ -191,13 +192,13 @@ Proof. move=> /= F mF tF mUF. rewrite /mpoisson1/= integral_bigcup//=; last first. apply/integrableP; split. - apply/EFin_measurable_fun. + apply/measurable_EFinP. exact: measurable_funS (measurable_poisson _). rewrite (_ : (fun x => _) = (EFin \o poisson1)); last first. by apply/funext => x; rewrite gee0_abs// lee_fin poisson1_ge0//. apply: le_lt_trans. apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. - by apply/EFin_measurable_fun; exact: measurable_poisson. + by apply/measurable_EFinP; exact: measurable_poisson. by move=> ? _; rewrite lee_fin poisson1_ge0//. by rewrite /= integral_poisson_density// ltry. apply: is_cvg_ereal_nneg_natsum_cond => n _ _. From ab61bf9e88c898b6e5260c7b4efa28f19a4e0245 Mon Sep 17 00:00:00 2001 From: AyumuSaito Date: Sat, 23 Sep 2023 11:15:51 +0900 Subject: [PATCH 04/48] add binary operations --- theories/lang_syntax.v | 81 +++++++++++++++++++++++++++++++++ theories/lang_syntax_examples.v | 72 +++++++++++++++++++++++++++++ 2 files changed, 153 insertions(+) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 8601c576fa..66121a9659 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -361,10 +361,52 @@ Context {R : realType}. Inductive flag := D | P. +Section binop. + +Inductive binop := +| binop_and | binop_or +| binop_add | binop_minus | binop_mult. + +Definition type_of_binop (b : binop) : typ := +match b with +| binop_and => Bool +| binop_or => Bool +| binop_add => Real +| binop_minus => Real +| binop_mult => Real +end. + +(* Import Notations. *) + +Definition fun_of_binop g (b : binop) : (mctx g -> mtyp (type_of_binop b)) -> + (mctx g -> mtyp (type_of_binop b)) -> @mctx R g -> @mtyp R (type_of_binop b) := +match b with +| binop_and => (fun f1 f2 x => f1 x && f2 x : mtyp Bool) +| binop_or => (fun f1 f2 x => f1 x || f2 x : mtyp Bool) +| binop_add => (fun f1 f2 => (f1 \+ f2)%R) +| binop_minus => (fun f1 f2 => (f1 \- f2)%R) +| binop_mult => (fun f1 f2 => (f1 \* f2)%R) +end. + +Definition mfun_of_binop g b + (f1 : @mctx R g -> @mtyp R (type_of_binop b)) (mf1 : measurable_fun setT f1) + (f2 : @mctx R g -> @mtyp R (type_of_binop b)) (mf2 : measurable_fun setT f2) : + measurable_fun [set: @mctx R g] (fun_of_binop f1 f2). +destruct b. +exact: measurable_and mf1 mf2. +exact: measurable_or mf1 mf2. +exact: measurable_funD. +exact: measurable_funB. +exact: measurable_funM. +Defined. + +End binop. + Inductive exp : flag -> ctx -> typ -> Type := | exp_unit g : exp D g Unit | exp_bool g : bool -> exp D g Bool | exp_real g : R -> exp D g Real +| exp_bin g (b : binop) : exp D g (type_of_binop b) -> exp D g (type_of_binop b) -> exp D g (type_of_binop b) | exp_pair g t1 t2 : exp D g t1 -> exp D g t2 -> exp D g (Pair t1 t2) | exp_proj1 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t1 | exp_proj2 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t2 @@ -396,6 +438,7 @@ Arguments exp {R}. Arguments exp_unit {R g}. Arguments exp_bool {R g}. Arguments exp_real {R g}. +Arguments exp_bin {R g} &. Arguments exp_pair {R g} & {t1 t2}. Arguments exp_var {R g} _ {t} H. Arguments exp_bernoulli {R g}. @@ -416,6 +459,16 @@ Notation "b ':B'" := (@exp_bool _ _ b%bool) (in custom expr at level 1) : lang_scope. Notation "r ':R'" := (@exp_real _ _ r%R) (in custom expr at level 1, format "r :R") : lang_scope. +Notation "e1 && e2" := (exp_bin binop_and e1 e2) + (in custom expr at level 1) : lang_scope. +Notation "e1 || e2" := (exp_bin binop_or e1 e2) + (in custom expr at level 1) : lang_scope. +Notation "e1 + e2" := (exp_bin binop_add e1 e2) + (in custom expr at level 1) : lang_scope. +Notation "e1 - e2" := (exp_bin binop_minus e1 e2) + (in custom expr at level 1) : lang_scope. +Notation "e1 * e2" := (exp_bin binop_mult e1 e2) + (in custom expr at level 1) : lang_scope. Notation "'return' e" := (@exp_return _ _ _ e) (in custom expr at level 2) : lang_scope. (*Notation "% str" := (@exp_var _ _ str%string _ erefl) @@ -457,6 +510,7 @@ Fixpoint free_vars k g t (e : @exp R k g t) : seq string := | exp_unit _ => [::] | exp_bool _ _ => [::] | exp_real _ _ => [::] + | exp_bin _ _ e1 e2 => free_vars e1 ++ free_vars e2 | exp_pair _ _ _ e1 e2 => free_vars e1 ++ free_vars e2 | exp_proj1 _ _ _ e => free_vars e | exp_proj2 _ _ _ e => free_vars e @@ -574,6 +628,10 @@ Inductive evalD : forall g t, exp D g t -> | eval_real g r : ([r:R] : exp D g _) -D> cst r ; kr r +| eval_bin g bop (e1 : exp D g _) f1 mf1 e2 f2 mf2 : + e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> + exp_bin bop e1 e2 -D> fun_of_binop f1 f2 ; mfun_of_binop mf1 mf2 + | eval_pair g t1 (e1 : exp D g t1) f1 mf1 t2 (e2 : exp D g t2) f2 mf2 : e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> [(e1, e2)] -D> fun x => (f1 x, f2 x) ; measurable_fun_prod mf1 mf2 @@ -676,6 +734,12 @@ all: (rewrite {g t e u v mu mv hu}). - move=> g r {}v {}mv. inversion 1; subst g0 r0. by inj_ex H3. +- move=> g bop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. + inversion 1; subst g0 bop0. + inj_ex H10; subst v. + inj_ex H5; subst e1. + inj_ex H6; subst e5. + by move: H4 H11 => /IH1 <- /IH2 <-. - move=> g t1 e1 f1 mf1 t2 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. simple inversion 1 => //; subst g0. case: H3 => ? ?; subst t0 t3. @@ -798,6 +862,12 @@ all: rewrite {g t e u v eu}. - move=> g r {}v {}mv. inversion 1; subst g0 r0. by inj_ex H3. +- move=> g bop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. + inversion 1; subst g0 bop0. + inj_ex H10; subst v. + inj_ex H5; subst e1. + inj_ex H6; subst e5. + by move: H4 H11 => /IH1 <- /IH2 <-. - move=> g t1 e1 f1 mf1 t2 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. simple inversion 1 => //; subst g0. case: H3 => ? ?; subst t0 t3. @@ -914,6 +984,8 @@ all: rewrite {z g t}. - by do 2 eexists; exact: eval_unit. - by do 2 eexists; exact: eval_bool. - by do 2 eexists; exact: eval_real. +- move=> g b e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. + by exists (fun_of_binop f1 f2); eexists; exact: eval_bin. - move=> g t1 t2 e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. by exists (fun x => (f1 x, f2 x)); eexists; exact: eval_pair. - move=> g t1 t2 e [f [mf H]]. @@ -1022,6 +1094,15 @@ Proof. exact/execD_evalD/eval_bool. Qed. Lemma execD_real g r : @execD g _ [r:R] = existT _ (cst r) (kr r). Proof. exact/execD_evalD/eval_real. Qed. +Lemma execD_bin g bop (e1 : exp D g _) (e2 : exp D g _) : + let f1 := projT1 (execD e1) in let f2 := projT1 (execD e2) in + let mf1 := projT2 (execD e1) in let mf2 := projT2 (execD e2) in + execD (exp_bin bop e1 e2) = + @existT _ _ (fun_of_binop f1 f2) (mfun_of_binop mf1 mf2). +Proof. +by move=> f1 f2 mf1 mf2; apply/execD_evalD/eval_bin; exact: evalD_execD. +Qed. + Lemma execD_pair g t1 t2 (e1 : exp D g t1) (e2 : exp D g t2) : let f1 := projT1 (execD e1) in let f2 := projT1 (execD e2) in let mf1 := projT2 (execD e1) in let mf2 := projT2 (execD e2) in diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index e02fec2395..bb6eff1072 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -260,6 +260,78 @@ rewrite exec_sample_pair0; do 3 rewrite mem_set//; rewrite memNset//=. by rewrite !mule1; congr (_%:E); field. Qed. +Definition sample_and_syntax0 : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in + let "y" := Sample {exp_bernoulli (1 / 3%:R)%:nng (p1S 2)} in + return #{"x"} && #{"y"}]. + +Lemma exec_sample_and0 (A : set bool) : + @execP R [::] _ sample_and_syntax0 tt A = + ((1 / 6)%:E * (true \in A)%:R%:E + + (1 - 1 / 6)%:E * (false \in A)%:R%:E)%E. +Proof. +rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. +rewrite (@execD_bin _ _ binop_and) !exp_var'E (execD_var_erefl "x") (execD_var_erefl "y") /=. +rewrite letin'E integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. +rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !indicE !in_setT/= !mul1e !diracE. +rewrite muleDr// -addeA; congr (_ + _)%E. +by rewrite !muleA; congr (_%:E); congr (_ * _); field. +rewrite -muleDl// !muleA -muleDl//. +by congr (_%:E); congr (_ * _); field. +Qed. + +Definition sample_bernoulli_and3 : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in + let "y" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in + let "z" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in + return #{"x"} && #{"y"} && #{"z"}]. + +Lemma exec_sample_bernoulli_and3 t U : + execP sample_bernoulli_and3 t U = + ((1 / 8)%:E * (true \in U)%:R%:E + + (1 - 1 / 8)%:E * (false \in U)%:R%:E)%E. +Proof. +rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. +rewrite !(@execD_bin _ _ binop_and) !exp_var'E. +rewrite (execD_var_erefl "x") (execD_var_erefl "y") (execD_var_erefl "z") /=. +rewrite letin'E integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. +rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. +rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !indicE !in_setT/= !mul1e !diracE. +rewrite !muleDr// -!addeA. +by congr (_ + _)%E; rewrite ?addeA !muleA -?muleDl//; +congr (_ * _)%E; congr (_%:E); field. +Qed. + +Definition sample_add_syntax0 : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in + let "y" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in + let "z" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in + return #{"x"} && #{"y"} && #{"z"}]. + +Lemma exec_sample_bernoulli_and3 t U : + execP sample_bernoulli_and3 t U = + ((1 / 8)%:E * (true \in U)%:R%:E + + (1 - 1 / 8)%:E * (false \in U)%:R%:E)%E. +Proof. +rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. +rewrite !(@execD_bin _ _ binop_and) !exp_var'E. +rewrite (execD_var_erefl "x") (execD_var_erefl "y") (execD_var_erefl "z") /=. +rewrite letin'E integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. +rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. +rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !indicE !in_setT/= !mul1e !diracE. +rewrite !muleDr// -!addeA. +by congr (_ + _)%E; rewrite ?addeA !muleA -?muleDl//; +congr (_ * _)%E; congr (_%:E); field. +Qed. + End sample_pair. Section bernoulli_examples. From 660f1960dd574ed2d0407f774e3ef06276dd71bd Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 10 Oct 2023 10:05:01 +0900 Subject: [PATCH 05/48] minor fixes, updates, rebases --- coq-mathcomp-analysis.opam | 3 +- theories/lang_syntax.v | 3 +- theories/lang_syntax_examples.v | 83 ++++++++++++--------------------- theories/prob_lang.v | 8 +--- 4 files changed, 36 insertions(+), 61 deletions(-) diff --git a/coq-mathcomp-analysis.opam b/coq-mathcomp-analysis.opam index b56eb46df9..c7f50a6eed 100644 --- a/coq-mathcomp-analysis.opam +++ b/coq-mathcomp-analysis.opam @@ -19,7 +19,8 @@ depends: [ "coq-mathcomp-solvable" "coq-mathcomp-field" "coq-mathcomp-bigenough" { (>= "1.0.0") } - "coq-mathcomp-algebra-tactics" { (>= "1.1.1") } + "coq-mathcomp-algebra-tactics" { (>= "1.2.0") } + "coq-mathcomp-zify" { (>= "1.4.0") } ] tags: [ diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 66121a9659..bf0fdc60a9 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -190,8 +190,7 @@ HB.instance Definition _ z := @isMeasure.Build _ X R (T' z) (T0 z) (T_ge0 z) (@T_semi_sigma_additive z). Let sfinT z : sfinite_measure (T' z). Proof. exact: sfinite_kernel_measure. Qed. -HB.instance Definition _ z := @isSFinite.Build _ X R - (T' z) (sfinT z). +HB.instance Definition _ z := @isSFinite.Build _ X R (T' z) (sfinT z). Definition U' z : set Y -> \bar R := u z. Let U0 z : (U' z) set0 = 0. Proof. by []. Qed. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index bb6eff1072..5db0d1901c 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -215,38 +215,38 @@ Definition sample_pair_syntax : exp _ [::] _ := Lemma exec_sample_pair0 (A : set (bool * bool)) : @execP R [::] _ sample_pair_syntax0 tt A = ((1 / 2)%:E * - ((1 / 3)%:E * ((true, true) \in A)%:R%:E + - (1 - 1 / 3)%:E * ((true, false) \in A)%:R%:E) + + ((1 / 3)%:E * \d_(true, true) A + + (1 - 1 / 3)%:E * \d_(true, false) A) + (1 - 1 / 2)%:E * - ((1 / 3)%:E * ((false, true) \in A)%:R%:E + - (1 - 1 / 3)%:E * ((false, false) \in A)%:R%:E))%E. + ((1 / 3)%:E * \d_(false, true) A + + (1 - 1 / 3)%:E * \d_(false, false) A))%E. Proof. rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. rewrite execD_pair !exp_var'E (execD_var_erefl "x") (execD_var_erefl "y") /=. rewrite letin'E ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !diracE !in_setT/= !mul1e. +rewrite !integral_dirac//= !diracT !mul1e. rewrite !letin'E !ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. -by rewrite !integral_dirac//= !diracT !mul1e !diracE. +by rewrite !integral_dirac//= !diracT !mul1e. Qed. Lemma exec_sample_pair0_TandT : @execP R [::] _ sample_pair_syntax0 tt [set (true, true)] = (1 / 6)%:E. Proof. -rewrite exec_sample_pair0 mem_set//; do 3 rewrite memNset//=. +rewrite exec_sample_pair0 !diracE mem_set//; do 3 rewrite memNset//=. by rewrite /= !mule0 mule1 !add0e mule0 adde0; congr (_%:E); lra. Qed. Lemma exec_sample_pair0_TandF : @execP R [::] _ sample_pair_syntax0 tt [set (true, false)] = (1 / 3)%:E. Proof. -rewrite exec_sample_pair0 memNset// mem_set//; do 2 rewrite memNset//. +rewrite exec_sample_pair0 !diracE memNset// mem_set//; do 2 rewrite memNset//. by rewrite /= !mule0 mule1 !add0e mule0 adde0; congr (_%:E); lra. Qed. Lemma exec_sample_pair0_TandT' : @execP R [::] _ sample_pair_syntax0 tt [set p | p.1 && p.2] = (1 / 6)%:E. Proof. -rewrite exec_sample_pair0 mem_set//; do 3 rewrite memNset//=. +rewrite exec_sample_pair0 !diracE mem_set//; do 3 rewrite memNset//=. by rewrite /= !mule0 mule1 !add0e mule0 adde0; congr (_%:E); lra. Qed. @@ -254,9 +254,9 @@ Lemma exec_sample_pair_TorT : (projT1 (execD sample_pair_syntax)) tt [set p | p.1 || p.2] = (2 / 3)%:E. Proof. rewrite execD_normalize_pt normalizeE/= exec_sample_pair0. -do 4 rewrite mem_set//=. +rewrite !diracE; do 4 rewrite mem_set//=. rewrite eqe ifF; last by apply/negbTE/negP => /orP[/eqP|//]; lra. -rewrite exec_sample_pair0; do 3 rewrite mem_set//; rewrite memNset//=. +rewrite exec_sample_pair0 !diracE; do 3 rewrite mem_set//; rewrite memNset//=. by rewrite !mule1; congr (_%:E); field. Qed. @@ -266,18 +266,17 @@ Definition sample_and_syntax0 : @exp R _ [::] _ := return #{"x"} && #{"y"}]. Lemma exec_sample_and0 (A : set bool) : - @execP R [::] _ sample_and_syntax0 tt A = - ((1 / 6)%:E * (true \in A)%:R%:E + - (1 - 1 / 6)%:E * (false \in A)%:R%:E)%E. + @execP R [::] _ sample_and_syntax0 tt A = ((1 / 6)%:E * \d_true A + + (1 - 1 / 6)%:E * \d_false A)%E. Proof. rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. rewrite (@execD_bin _ _ binop_and) !exp_var'E (execD_var_erefl "x") (execD_var_erefl "y") /=. -rewrite letin'E integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. -rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !indicE !in_setT/= !mul1e !diracE. +rewrite letin'E ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !diracT !mul1e. +rewrite !letin'E !ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !diracT !mul1e. rewrite muleDr// -addeA; congr (_ + _)%E. -by rewrite !muleA; congr (_%:E); congr (_ * _); field. + by rewrite !muleA; congr (_%:E); congr (_ * _); field. rewrite -muleDl// !muleA -muleDl//. by congr (_%:E); congr (_ * _); field. Qed. @@ -289,19 +288,18 @@ Definition sample_bernoulli_and3 : @exp R _ [::] _ := return #{"x"} && #{"y"} && #{"z"}]. Lemma exec_sample_bernoulli_and3 t U : - execP sample_bernoulli_and3 t U = - ((1 / 8)%:E * (true \in U)%:R%:E + - (1 - 1 / 8)%:E * (false \in U)%:R%:E)%E. + execP sample_bernoulli_and3 t U = ((1 / 8)%:E * \d_true U + + (1 - 1 / 8)%:E * \d_false U)%E. Proof. rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. rewrite !(@execD_bin _ _ binop_and) !exp_var'E. rewrite (execD_var_erefl "x") (execD_var_erefl "y") (execD_var_erefl "z") /=. -rewrite letin'E integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. -rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. -rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !indicE !in_setT/= !mul1e !diracE. +rewrite letin'E ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !diracT !mul1e. +rewrite !letin'E !ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !diracT !mul1e. +rewrite !letin'E !ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. +rewrite !integral_dirac//= !diracT !mul1e. rewrite !muleDr// -!addeA. by congr (_ + _)%E; rewrite ?addeA !muleA -?muleDl//; congr (_ * _)%E; congr (_%:E); field. @@ -313,25 +311,6 @@ Definition sample_add_syntax0 : @exp R _ [::] _ := let "z" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in return #{"x"} && #{"y"} && #{"z"}]. -Lemma exec_sample_bernoulli_and3 t U : - execP sample_bernoulli_and3 t U = - ((1 / 8)%:E * (true \in U)%:R%:E + - (1 - 1 / 8)%:E * (false \in U)%:R%:E)%E. -Proof. -rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. -rewrite !(@execD_bin _ _ binop_and) !exp_var'E. -rewrite (execD_var_erefl "x") (execD_var_erefl "y") (execD_var_erefl "z") /=. -rewrite letin'E integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. -rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !indicE !in_setT/= !mul1e. -rewrite !letin'E !integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !indicE !in_setT/= !mul1e !diracE. -rewrite !muleDr// -!addeA. -by congr (_ + _)%E; rewrite ?addeA !muleA -?muleDl//; -congr (_ * _)%E; congr (_%:E); field. -Qed. - End sample_pair. Section bernoulli_examples. @@ -374,12 +353,12 @@ rewrite -EFinD/= eqe ifF; last first. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. -rewrite !integral_dirac//= !diracT /= !mul1e ger0_norm// !indicT. +rewrite !integral_dirac//= !diracT !mul1e ger0_norm//. rewrite exp_var'E (execD_var_erefl "x")/=. rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. by rewrite muleDl//; congr (_ + _)%E; rewrite -!EFinM; congr (_%:E); - rewrite !indicE /onem /=; case: (_ \in _); field. + rewrite !indicT !indicE /onem /=; case: (_ \in _); field. Qed. Definition bernoulli12_score := [Normalize @@ -417,13 +396,13 @@ rewrite -EFinD/= eqe ifF; last first. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. -rewrite !integral_dirac//= !diracT /= !mul1e ger0_norm// !indicT. +rewrite !integral_dirac//= !diracT !mul1e ger0_norm//. rewrite exp_var'E (execD_var_erefl "x")/=. rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. by rewrite muleDl//; congr (_ + _)%E; rewrite -!EFinM; congr (_%:E); - rewrite indicE /onem; case: (_ \in _); field. + by rewrite !indicT !indicE /onem /=; case: (_ \in _); field. Qed. (* https://dl.acm.org/doi/pdf/10.1145/2933575.2935313 (Sect. 4) *) @@ -473,7 +452,7 @@ rewrite !indicT !mulr1. by rewrite muleDl//; congr (_ + _)%E; rewrite -!EFinM; congr (_%:E); - by rewrite indicE /onem; case: (_ \in _); field. + by rewrite !indicE /onem /=; case: (_ \in _); field. Qed. End bernoulli_examples. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index bf077591d6..13e976768f 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -88,7 +88,6 @@ Qed. Section bernoulli. Variables (R : realType) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). -Local Open Scope ring_scope. Definition bernoulli : set _ -> \bar R := measure_add @@ -97,8 +96,6 @@ Definition bernoulli : set _ -> \bar R := HB.instance Definition _ := Measure.on bernoulli. -Local Close Scope ring_scope. - Let bernoulli_setT : bernoulli [set: _] = 1. Proof. rewrite /bernoulli/= /measure_add/= /msum 2!big_ord_recr/= big_ord0 add0e/=. @@ -1020,8 +1017,7 @@ HB.instance Definition _ z := @isMeasure.Build _ Y R (U z) (U0 z) (U_ge0 z) (@U_semi_sigma_additive z). Let sfinU z : sfinite_measure (U z). Proof. exact: sfinite_kernel_measure. Qed. -HB.instance Definition _ z := @isSFinite.Build _ Y R - (U z) (sfinU z). +HB.instance Definition _ z := @isSFinite.Build _ Y R (U z) (sfinU z). Lemma letinC z A : measurable A -> letin t @@ -1055,7 +1051,7 @@ End letinC. Section constants. Variable R : realType. -Local Open Scope ring_scope. +Local Close Scope ereal_scope. Lemma onem1S n : `1- (1 / n.+1%:R) = (n%:R / n.+1%:R)%:nng%:num :> R. Proof. From 52041345dd440c5438cf117f089929a3f7124ce5 Mon Sep 17 00:00:00 2001 From: AyumuSaito Date: Tue, 14 Feb 2023 22:02:38 +0900 Subject: [PATCH 06/48] casino example and extensions - binomial - uniform - generalize measure_and - change the order of binop & add pow - fix for mR - rewrite casino0 to casino1 Co-authored-by: Reynald Affeldt --- theories/lang_syntax.v | 286 +++++++++++++++++++---- theories/lang_syntax_examples.v | 100 ++++---- theories/lang_syntax_examples_wip.v | 329 ++++++++++++++++++++++++++ theories/prob_lang.v | 346 +++++++++++++++++++++++----- theories/prob_lang_wip.v | 2 +- 5 files changed, 910 insertions(+), 153 deletions(-) create mode 100644 theories/lang_syntax_examples_wip.v diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index bf0fdc60a9..d11bc21425 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -276,7 +276,9 @@ Fixpoint measurable_of_typ (t : typ) : {d & measurableType d} := match t with | Unit => existT _ _ munit | Bool => existT _ _ mbool - | Real => existT _ _ [the measurableType _ of (mR R)] + | Real => existT _ _ + [the measurableType _ of (@measurableTypeR R)] + (* (Real_sort__canonical__measure_Measurable R) *) | Pair A B => existT _ _ [the measurableType (projT1 (measurable_of_typ A), projT1 (measurable_of_typ B)).-prod%mdisp of @@ -378,7 +380,7 @@ end. (* Import Notations. *) Definition fun_of_binop g (b : binop) : (mctx g -> mtyp (type_of_binop b)) -> - (mctx g -> mtyp (type_of_binop b)) -> @mctx R g -> @mtyp R (type_of_binop b) := + (mctx g -> mtyp (type_of_binop b)) -> @mctx R g -> @mtyp R (type_of_binop b) := match b with | binop_and => (fun f1 f2 x => f1 x && f2 x : mtyp Bool) | binop_or => (fun f1 f2 x => f1 x || f2 x : mtyp Bool) @@ -387,8 +389,8 @@ match b with | binop_mult => (fun f1 f2 => (f1 \* f2)%R) end. -Definition mfun_of_binop g b - (f1 : @mctx R g -> @mtyp R (type_of_binop b)) (mf1 : measurable_fun setT f1) +Definition mfun_of_binop g b + (f1 : @mctx R g -> @mtyp R (type_of_binop b)) (mf1 : measurable_fun setT f1) (f2 : @mctx R g -> @mtyp R (type_of_binop b)) (mf2 : measurable_fun setT f2) : measurable_fun [set: @mctx R g] (fun_of_binop f1 f2). destruct b. @@ -401,17 +403,52 @@ Defined. End binop. +Section relop. +Inductive relop := +| relop_le | relop_lt | relop_eq . + +Definition fun_of_relop g (r : relop) : (@mctx R g -> @mtyp R Real) -> + (mctx g -> mtyp Real) -> @mctx R g -> @mtyp R Bool := +match r with +| relop_le => (fun f1 f2 x => (f1 x <= f2 x)%R) +| relop_lt => (fun f1 f2 x => (f1 x < f2 x)%R) +| relop_eq => (fun f1 f2 x => (f1 x == f2 x)%R) +end. + +Definition mfun_of_relop g r + (f1 : @mctx R g -> @mtyp R Real) (mf1 : measurable_fun setT f1) + (f2 : @mctx R g -> @mtyp R Real) (mf2 : measurable_fun setT f2) : + measurable_fun [set: @mctx R g] (fun_of_relop r f1 f2). +destruct r. +exact: measurable_fun_ler. +exact: measurable_fun_ltr. +exact: measurable_fun_eqr. +Defined. + +End relop. + Inductive exp : flag -> ctx -> typ -> Type := | exp_unit g : exp D g Unit | exp_bool g : bool -> exp D g Bool | exp_real g : R -> exp D g Real -| exp_bin g (b : binop) : exp D g (type_of_binop b) -> exp D g (type_of_binop b) -> exp D g (type_of_binop b) +| exp_pow g : nat -> exp D g Real -> exp D g Real +| exp_bin (b : binop) g : exp D g (type_of_binop b) -> + exp D g (type_of_binop b) -> exp D g (type_of_binop b) +| exp_rel (r : relop) g : exp D g Real -> + exp D g Real -> exp D g Bool | exp_pair g t1 t2 : exp D g t1 -> exp D g t2 -> exp D g (Pair t1 t2) | exp_proj1 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t1 | exp_proj2 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t2 | exp_var g str t : t = lookup Unit g str -> exp D g t | exp_bernoulli g (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : exp D g (Prob Bool) +| exp_bernoulli_trunc g : + exp D g Real -> exp D g (Prob Bool) +| exp_binomial g (n : nat) (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : + exp D g (Prob Real) +| exp_binomial_trunc g (n : nat) : + exp D g Real -> exp D g (Prob Real) +| exp_uniform g (a b : R) (ab0 : (0 < b - a)%R) : exp D g (Prob Real) | exp_poisson g : nat -> exp D g Real -> exp D g Real | exp_normalize g t : exp P g t -> exp D g (Prob t) | exp_letin g t1 t2 str : exp P g t1 -> exp P ((str, t1) :: g) t2 -> @@ -437,19 +474,25 @@ Arguments exp {R}. Arguments exp_unit {R g}. Arguments exp_bool {R g}. Arguments exp_real {R g}. -Arguments exp_bin {R g} &. +Arguments exp_pow {R g}. +Arguments exp_bin {R} b {g} &. +Arguments exp_rel {R} r {g} &. Arguments exp_pair {R g} & {t1 t2}. -Arguments exp_var {R g} _ {t} H. +Arguments exp_var {R g} _ {t} & H. Arguments exp_bernoulli {R g}. +Arguments exp_bernoulli_trunc {R g} &. +Arguments exp_binomial {R g}. +Arguments exp_uniform {R g} &. +Arguments exp_binomial_trunc {R g} &. Arguments exp_poisson {R g}. Arguments exp_normalize {R g _}. Arguments exp_letin {R g} & {_ _}. -Arguments exp_sample {R g t}. +Arguments exp_sample {R g} & {t}. Arguments exp_score {R g}. Arguments exp_return {R g} & {_}. -Arguments exp_if {R z g t}. +Arguments exp_if {R z g t} &. Arguments exp_weak {R} z g h {t} x. -Arguments exp_var' {R} str {t} g. +Arguments exp_var' {R} str {t} g &. Declare Custom Entry expr. Notation "[ e ]" := e (e custom expr at level 5) : lang_scope. @@ -458,18 +501,24 @@ Notation "b ':B'" := (@exp_bool _ _ b%bool) (in custom expr at level 1) : lang_scope. Notation "r ':R'" := (@exp_real _ _ r%R) (in custom expr at level 1, format "r :R") : lang_scope. -Notation "e1 && e2" := (exp_bin binop_and e1 e2) +Notation "e ^+ n" := (exp_pow n e) (in custom expr at level 1) : lang_scope. +Notation "e1 && e2" := (exp_bin binop_and e1 e2) + (in custom expr at level 2) : lang_scope. Notation "e1 || e2" := (exp_bin binop_or e1 e2) - (in custom expr at level 1) : lang_scope. + (in custom expr at level 2) : lang_scope. Notation "e1 + e2" := (exp_bin binop_add e1 e2) - (in custom expr at level 1) : lang_scope. + (in custom expr at level 3) : lang_scope. Notation "e1 - e2" := (exp_bin binop_minus e1 e2) - (in custom expr at level 1) : lang_scope. + (in custom expr at level 3) : lang_scope. Notation "e1 * e2" := (exp_bin binop_mult e1 e2) - (in custom expr at level 1) : lang_scope. -Notation "'return' e" := (@exp_return _ _ _ e) (in custom expr at level 2) : lang_scope. +Notation "e1 <= e2" := (exp_rel relop_le e1 e2) + (in custom expr at level 2) : lang_scope. +Notation "e1 == e2" := (exp_rel relop_eq e1 e2) + (in custom expr at level 4) : lang_scope. +Notation "'return' e" := (@exp_return _ _ _ e) + (in custom expr at level 6) : lang_scope. (*Notation "% str" := (@exp_var _ _ str%string _ erefl) (in custom expr at level 1, format "% str") : lang_scope.*) (* Notation "% str H" := (@exp_var _ _ str%string _ H) @@ -485,21 +534,21 @@ Notation "\pi_1 e" := (exp_proj1 e) Notation "\pi_2 e" := (exp_proj2 e) (in custom expr at level 1) : lang_scope. Notation "'let' x ':=' e 'in' f" := (exp_letin x e f) - (in custom expr at level 3, + (in custom expr at level 5, x constr, - f custom expr at level 3, + f custom expr at level 5, left associativity) : lang_scope. Notation "{ c }" := c (in custom expr, c constr) : lang_scope. Notation "x" := x (in custom expr at level 0, x ident) : lang_scope. Notation "'Sample' e" := (exp_sample e) - (in custom expr at level 2) : lang_scope. + (in custom expr at level 5) : lang_scope. Notation "'Score' e" := (exp_score e) - (in custom expr at level 2) : lang_scope. + (in custom expr at level 5) : lang_scope. Notation "'Normalize' e" := (exp_normalize e) (in custom expr at level 0) : lang_scope. Notation "'if' e1 'then' e2 'else' e3" := (exp_if e1 e2 e3) - (in custom expr at level 1) : lang_scope. + (in custom expr at level 6) : lang_scope. Section free_vars. Context {R : realType}. @@ -509,12 +558,18 @@ Fixpoint free_vars k g t (e : @exp R k g t) : seq string := | exp_unit _ => [::] | exp_bool _ _ => [::] | exp_real _ _ => [::] + | exp_pow _ _ e => free_vars e | exp_bin _ _ e1 e2 => free_vars e1 ++ free_vars e2 + | exp_rel _ _ e1 e2 => free_vars e1 ++ free_vars e2 | exp_pair _ _ _ e1 e2 => free_vars e1 ++ free_vars e2 | exp_proj1 _ _ _ e => free_vars e | exp_proj2 _ _ _ e => free_vars e | exp_var _ x _ _ => [:: x] | exp_bernoulli _ _ _ => [::] + | exp_bernoulli_trunc _ e => free_vars e + | exp_binomial _ _ _ _ => [::] + | exp_uniform _ _ _ _ => [::] + | exp_binomial_trunc _ _ e => free_vars e | exp_poisson _ _ e => free_vars e | exp_normalize _ _ e => free_vars e | exp_letin _ _ _ x e1 e2 => free_vars e1 ++ rem x (free_vars e2) @@ -619,6 +674,20 @@ Context {R : realType}. Implicit Type (g : ctx) (str : string). Local Open Scope lang_scope. +Local Open Scope ring_scope. +Definition bernoulli0 := @bernoulli R 0%R%:nng ler01. + +HB.instance Definition _ := Probability.on bernoulli0. + +Lemma __ : Measurable.sort + (pprobability + [the measurableType (R.-ocitv.-measurable).-sigma of + g_sigma_algebraType (R.-ocitv.-measurable)] R) = +Measurable.sort (@mtyp R (Prob Real)). +rewrite /=. +(* done. *) +Abort. + Inductive evalD : forall g t, exp D g t -> forall f : dval R g t, measurable_fun setT f -> Prop := | eval_unit g : ([TT] : exp D g _) -D> cst tt ; ktt @@ -627,10 +696,17 @@ Inductive evalD : forall g t, exp D g t -> | eval_real g r : ([r:R] : exp D g _) -D> cst r ; kr r +| eval_pow g n (e : exp D g _) f mf : e -D> f ; mf -> + [e ^+ {n}] -D> (fun x => f x ^+ n) ; (measurable_funX n mf) + | eval_bin g bop (e1 : exp D g _) f1 mf1 e2 f2 mf2 : e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> exp_bin bop e1 e2 -D> fun_of_binop f1 f2 ; mfun_of_binop mf1 mf2 +| eval_rel g rop (e1 : exp D g _) f1 mf1 e2 f2 mf2 : + e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> + exp_rel rop e1 e2 -D> fun_of_relop rop f1 f2 ; mfun_of_relop rop mf1 mf2 + | eval_pair g t1 (e1 : exp D g t1) f1 mf1 t2 (e2 : exp D g t2) f2 mf2 : e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> [(e1, e2)] -D> fun x => (f1 x, f2 x) ; measurable_fun_prod mf1 mf2 @@ -649,9 +725,27 @@ Inductive evalD : forall g t, exp D g t -> | eval_var g x H : let i := index x (dom g) in exp_var x H -D> acc_typ (map snd g) i ; measurable_acc_typ (map snd g) i -| eval_bernoulli g (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : +| eval_bernoulli g r r1 : (exp_bernoulli r r1 : exp D g _) -D> cst (bernoulli r1) ; - measurable_cst _ + measurable_cst _ + +| eval_bernoulli_trunc g e r mr : + e -D> r ; mr -> + (exp_bernoulli_trunc e : exp D g _) -D> bernoulli_trunc \o r ; + measurableT_comp measurable_bernoulli_trunc mr + +| eval_binomial g n (p : {nonneg R}) (p1 : (p%:num <= 1)%R) : + (exp_binomial n p p1 : exp D g _) -D> cst (binomial_probability n p1) ; + measurable_cst _ + +| eval_binomial_trunc g n e r mr : + e -D> r ; mr -> + (exp_binomial_trunc n e : exp D g _) -D> binomial_probability_trunc n \o r ; + measurableT_comp (measurable_binomial_probability_trunc n) mr + +| eval_uniform g (a b : R) (ab0 : (0 < b - a)%R) : + (exp_uniform a b ab0 : exp D g _) -D> cst (uniform_probability ab0) ; + measurable_cst _ | eval_poisson g n (e : exp D g _) f mf : e -D> f ; mf -> @@ -681,8 +775,8 @@ with evalP : forall g t, exp P g t -> pval R g t -> Prop := [let str := e1 in e2] -P> letin' k1 k2 | eval_sample g t (e : exp _ _ (Prob t)) - (f : mctx g -> probability (mtyp t) R) mf : - e -D> f ; mf -> [Sample e] -P> sample f mf + (p : mctx g -> pprobability (mtyp t) R) mp : + e -D> p ; mp -> [Sample e] -P> sample p mp | eval_score g (e : exp _ g _) f mf : e -D> f ; mf -> [Score e] -P> kscore mf @@ -733,12 +827,23 @@ all: (rewrite {g t e u v mu mv hu}). - move=> g r {}v {}mv. inversion 1; subst g0 r0. by inj_ex H3. +- move=> g n e f mf ev IH {}v {}mv. + inversion 1; subst g0 n0. + inj_ex H4; subst v. + inj_ex H2; subst e0. + by move: H3 => /IH <-. - move=> g bop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. inversion 1; subst g0 bop0. inj_ex H10; subst v. inj_ex H5; subst e1. inj_ex H6; subst e5. by move: H4 H11 => /IH1 <- /IH2 <-. +- move=> g rop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. + inversion 1; subst g0 rop0. + inj_ex H5; subst v. + inj_ex H1; subst e1. + inj_ex H3; subst e3. + by move: H6 H7 => /IH1 <- /IH2 <-. - move=> g t1 e1 f1 mf1 t2 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. simple inversion 1 => //; subst g0. case: H3 => ? ?; subst t0 t3. @@ -759,22 +864,33 @@ all: (rewrite {g t e u v mu mv hu}). clear H9. inj_ex H7; subst e1. by rewrite (ih _ _ H4). -(* - move=> g str n {}v {}mv. - inversion 1; subst g0. - inj_ex H6; rewrite -H6. - by inj_ex H7. - inj_ex H8; rewrite -H8. - by inj_ex H9. *) - move=> g str H n {}v {}mv. inversion 1; subst g0. - (* inj_ex H7; rewrite -H7. - by inj_ex H8. *) inj_ex H9; rewrite -H9. by inj_ex H10. - move=> g r r1 {}v {}mv. inversion 1; subst g0 r0. inj_ex H3; subst v. by have -> : r1 = r3 by []. +- move=> g e r mr ev IH {}v {}mv. + inversion 1; subst g0. + inj_ex H0; subst e0. + inj_ex H4; subst v. + by rewrite (IH _ _ H2). +- move=> g n p p1 {}v {}mv. + inversion 1; subst g0 n0 p0. + inj_ex H4; subst v. + by have -> : p1 = p3 by []. +- move=> g n e f mf ev IH {}v {}mv. + inversion 1; subst g0 n0. + inj_ex H2; subst e0. + inj_ex H4; subst v. + inj_ex H5; subst mv. + by rewrite (IH _ _ H3). +- move=> g a b ab0 {}v {}mv. + inversion 1; subst g0 a0 b0. + inj_ex H4; subst v. + by have -> : ab0 = ab2. - move=> g n e0 f mf ev IH {}v {}mv. inversion 1; subst g0 n0. inj_ex H2; subst e0. @@ -807,13 +923,13 @@ all: (rewrite {g t e u v mu mv hu}). inj_ex H6; subst e5. inj_ex H5; subst e4. by rewrite (IH1 _ H4) (IH2 _ H8). -- move=> g t e f mf ev IH k. +- move=> g t e p mp ev IH k. inversion 1; subst g0. inj_ex H5; subst t0. inj_ex H5; subst e1. inj_ex H7; subst k. - have ? := IH _ _ H3; subst f1. - by have -> : mf = mf1 by []. + have ? := IH _ _ H3; subst p1. + by have -> : mp = mp1 by []. - move=> g e f mf ev IH k. inversion 1; subst g0. inj_ex H0; subst e0. @@ -861,12 +977,23 @@ all: rewrite {g t e u v eu}. - move=> g r {}v {}mv. inversion 1; subst g0 r0. by inj_ex H3. +- move=> g n e f mf ev IH {}v {}mv. + inversion 1; subst g0 n0. + inj_ex H4; subst v. + inj_ex H2; subst e0. + by move: H3 => /IH <-. - move=> g bop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. inversion 1; subst g0 bop0. inj_ex H10; subst v. inj_ex H5; subst e1. inj_ex H6; subst e5. by move: H4 H11 => /IH1 <- /IH2 <-. +- move=> g rop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. + inversion 1; subst g0 rop0. + inj_ex H5; subst v. + inj_ex H1; subst e1. + inj_ex H3; subst e3. + by move: H6 H7 => /IH1 <- /IH2 <-. - move=> g t1 e1 f1 mf1 t2 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. simple inversion 1 => //; subst g0. case: H3 => ? ?; subst t0 t3. @@ -888,22 +1015,33 @@ all: rewrite {g t e u v eu}. clear H9. inj_ex H7; subst e1. by rewrite (ih _ _ H4). -(* - move=> g str n {}v {}mv. - inversion 1; subst g0. - inj_ex H6; rewrite -H6. - by inj_ex H7. - inj_ex H8; rewrite -H8. - by inj_ex H9. *) - move=> g str H n {}v {}mv. inversion 1; subst g0. - (* inj_ex H7; rewrite -H7. - by inj_ex H8. *) inj_ex H9; rewrite -H9. by inj_ex H10. - move=> g r r1 {}v {}mv. inversion 1; subst g0 r0. inj_ex H3; subst v. by have -> : r1 = r3 by []. +- move=> g e r mr ev IH {}v {}mv. + inversion 1; subst g0. + inj_ex H0; subst e0. + inj_ex H4; subst v. + by rewrite (IH _ _ H2). +- move=> g n p p1 {}v {}mv. + inversion 1; subst g0 n0 p0. + inj_ex H4; subst v. + by have -> : p1 = p3 by []. +- move=> g n e f mf ev IH {}v {}mv. + inversion 1; subst g0 n0. + inj_ex H2; subst e0. + inj_ex H4; subst v. + inj_ex H5; subst mv. + by rewrite (IH _ _ H3). +- move=> g a b ab0 {}v {}mv. + inversion 1; subst g0 a0 b0. + inj_ex H4; subst v. + by have -> : ab0 = ab2. - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. inj_ex H2; subst e0. @@ -938,12 +1076,12 @@ all: rewrite {g t e u v eu}. inj_ex H5; subst e4. inj_ex H6; subst e5. by rewrite (IH1 _ H4) (IH2 _ H8). -- move=> g t e f mf ep IH v. +- move=> g t e p mp ep IH v. inversion 1; subst g0 t0. inj_ex H7; subst v. inj_ex H5; subst e1. - have ? := IH _ _ H3; subst f1. - by have -> : mf = mf1 by []. + have ? := IH _ _ H3; subst p1. + by have -> : mp = mp1 by []. - move=> g e f mf ev IH k. inversion 1; subst g0. inj_ex H0; subst e0. @@ -983,8 +1121,12 @@ all: rewrite {z g t}. - by do 2 eexists; exact: eval_unit. - by do 2 eexists; exact: eval_bool. - by do 2 eexists; exact: eval_real. -- move=> g b e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. +- move=> g n e [f [mf H]]. + by exists (fun x => (f x ^+ n)%R); eexists; exact: eval_pow. +- move=> b g e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. by exists (fun_of_binop f1 f2); eexists; exact: eval_bin. +- move=> r g e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. + by exists (fun_of_relop r f1 f2); eexists; exact: eval_rel. - move=> g t1 t2 e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. by exists (fun x => (f1 x, f2 x)); eexists; exact: eval_pair. - move=> g t1 t2 e [f [mf H]]. @@ -993,6 +1135,13 @@ all: rewrite {z g t}. by exists (snd \o f); eexists; exact: eval_proj2. - by move=> g x t tE; subst t; eexists; eexists; exact: eval_var. - by move=> r r1; eexists; eexists; exact: eval_bernoulli. +- move=> g e [p [mp H]]. + by exists (bernoulli_trunc \o p); eexists; exact: eval_bernoulli_trunc. +- by move=> p p1; eexists; eexists; exact: eval_binomial. +- move=> g n e [p [mp H]]. + exists (binomial_probability_trunc n \o p). + eexists; exact: (eval_binomial_trunc n). +- by eexists; eexists; exact: eval_uniform. - move=> g h e [f [mf H]]. by exists (poisson h \o f); eexists; exact: eval_poisson. - move=> g t e [k ek]. @@ -1093,13 +1242,31 @@ Proof. exact/execD_evalD/eval_bool. Qed. Lemma execD_real g r : @execD g _ [r:R] = existT _ (cst r) (kr r). Proof. exact/execD_evalD/eval_real. Qed. +Local Open Scope ring_scope. +Lemma execD_pow g (e : exp D g _) n : + let f := projT1 (execD e) in let mf := projT2 (execD e) in + execD (exp_pow n e) = + @existT _ _ (fun x => f x ^+ n) (measurable_funX n mf). +Proof. +by move=> f mf; apply/execD_evalD/eval_pow/evalD_execD. +Qed. + Lemma execD_bin g bop (e1 : exp D g _) (e2 : exp D g _) : let f1 := projT1 (execD e1) in let f2 := projT1 (execD e2) in let mf1 := projT2 (execD e1) in let mf2 := projT2 (execD e2) in execD (exp_bin bop e1 e2) = @existT _ _ (fun_of_binop f1 f2) (mfun_of_binop mf1 mf2). Proof. -by move=> f1 f2 mf1 mf2; apply/execD_evalD/eval_bin; exact: evalD_execD. +by move=> f1 f2 mf1 mf2; apply/execD_evalD/eval_bin; exact/evalD_execD. +Qed. + +Lemma execD_rel g rop (e1 : exp D g _) (e2 : exp D g _) : + let f1 := projT1 (execD e1) in let f2 := projT1 (execD e2) in + let mf1 := projT2 (execD e1) in let mf2 := projT2 (execD e2) in + execD (exp_rel rop e1 e2) = + @existT _ _ (fun_of_relop rop f1 f2) (mfun_of_relop rop mf1 mf2). +Proof. +by move=> f1 f2 mf1 mf2; apply/execD_evalD/eval_rel; exact: evalD_execD. Qed. Lemma execD_pair g t1 t2 (e1 : exp D g t1) (e2 : exp D g t2) : @@ -1146,6 +1313,27 @@ Lemma execD_bernoulli g r (r1 : (r%:num <= 1)%R) : existT _ (cst [the probability _ _ of bernoulli r1]) (measurable_cst _). Proof. exact/execD_evalD/eval_bernoulli. Qed. +Lemma execD_bernoulli_trunc g e : + @execD g _ (exp_bernoulli_trunc e) = + existT _ (bernoulli_trunc \o projT1 (execD e)) (measurableT_comp measurable_bernoulli_trunc (projT2 (execD e))). +Proof. exact/execD_evalD/eval_bernoulli_trunc/evalD_execD. Qed. + +Lemma execD_binomial g n p (p1 : (p%:num <= 1)%R) : + @execD g _ (exp_binomial n p p1) = + existT _ (cst [the probability _ _ of binomial_probability n p1]) (measurable_cst _). +Proof. exact/execD_evalD/eval_binomial. Qed. + +Lemma execD_binomial_trunc g n e : + @execD g _ (exp_binomial_trunc n e) = + existT _ (binomial_probability_trunc n \o projT1 (execD e)) + (measurableT_comp (measurable_binomial_probability_trunc n) (projT2 (execD e))). +Proof. exact/execD_evalD/eval_binomial_trunc/evalD_execD. Qed. + +Lemma execD_uniform g a b ab0 : + @execD g _ (exp_uniform a b ab0) = + existT _ (cst [the probability _ _ of uniform_probability ab0]) (measurable_cst _). +Proof. exact/execD_evalD/eval_uniform. Qed. + Lemma execD_normalize_pt g t (e : exp P g t) : @execD g _ [Normalize e] = existT _ (normalize_pt (execP e) : _ -> pprobability _ _) diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index 5db0d1901c..ec7e244c79 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -190,10 +190,9 @@ Context {R : realType}. Lemma exec_normalize_return g x r : projT1 (@execD _ g _ [Normalize return r:R]) x = @dirac _ (measurableTypeR R) r _ :> probability _ R. - (* NB: \d_r notation? *) + (* TODO: try to use the notation \d_r *) Proof. -rewrite execD_normalize_pt execP_return execD_real/=. -exact: normalize_kdirac. +by rewrite execD_normalize_pt execP_return execD_real//=; exact: normalize_kdirac. Qed. End trivial_example. @@ -305,12 +304,6 @@ by congr (_ + _)%E; rewrite ?addeA !muleA -?muleDl//; congr (_ * _)%E; congr (_%:E); field. Qed. -Definition sample_add_syntax0 : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in - let "y" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in - let "z" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in - return #{"x"} && #{"y"} && #{"z"}]. - End sample_pair. Section bernoulli_examples. @@ -349,7 +342,7 @@ rewrite !integral_indic//= !iteE/= /mscale/=. rewrite setTI !diracT !mule1. rewrite ger0_norm//. rewrite -EFinD/= eqe ifF; last first. - by apply/negbTE/negP => /orP[/eqP|//]; rewrite /onem; lra. + apply/negbTE/negP => /orP[/eqP|//]; rewrite /onem. lra. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. @@ -457,6 +450,51 @@ Qed. End bernoulli_examples. +Section binomial_examples. +Context {R : realType}. +Open Scope lang_scope. +Open Scope ring_scope. + +Definition sample_binomial3 : @exp R _ [::] _ := + [let "x" := Sample {exp_binomial 3 (1 / 2)%:nng (p1S 1)} in + return #{"x"}]. + +Lemma exec_sample_binomial3 t U : measurable U -> + execP sample_binomial3 t U = ((1 / 8)%:E * \d_(0%:R : R) U + + (3 / 8)%:E * \d_(1%:R : R) U + + (3 / 8)%:E * \d_(2%:R : R) U + + (1 / 8)%:E * \d_(3%:R : R) U)%E. +Proof. +move=> mU; rewrite /sample_binomial3 execP_letin execP_sample execP_return. +rewrite exp_var'E (execD_var_erefl "x") !execD_binomial/=. +rewrite letin'E ge0_integral_measure_sum//=; last first. + exact: measurable_fun_dirac. +rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//=; [|exact: measurable_fun_dirac..]. +rewrite !integral_dirac// /bump; [|exact: measurable_fun_dirac..]. +rewrite !binS/= !bin0 bin1 bin2 bin_small// addn0. +rewrite expr0 mulr1 mul1r subn0. +rewrite -2!addeA !mul1r. +congr _%:E. +rewrite indicT !mul1r. +congr (_ + _). + congr (_ * _). + by field. +congr (_ + _). + congr (_ * _). + rewrite expr1 /onem. + by field. +congr (_ + _). + congr (_ * _). + rewrite /onem/=. + by field. +rewrite addr0. +congr (_ * _). +rewrite /onem/=. +by field. +Qed. + +End binomial_examples. + Section hard_constraint'. Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). @@ -480,7 +518,7 @@ Proof. apply/eq_sfkernel => x U. rewrite letin'E/= /sample; unlock. rewrite ge0_integral_measure_add//= ge0_integral_mscale//= ge0_integral_mscale//=. -rewrite integral_dirac//= integral_dirac//= !diracT/= !mul1e. +rewrite !integral_dirac//= !diracT/= !mul1e. by rewrite /mscale/= iteE//= iteE//= fail'E mule0 adde0 ger0_norm. Qed. @@ -576,43 +614,12 @@ Local Open Scope lang_scope. Import Notations. Context {R : realType}. -Section tests. - -Local Notation "$ str" := (@exp_var _ _ str%string _ erefl) - (in custom expr at level 1, format "$ str"). - -Definition staton_bus_syntax0_generic (x r u : string) - (rx : infer (r != x)) (Rx : infer (u != x)) - (ur : infer (u != r)) (xr : infer (x != r)) - (xu : infer (x != u)) (ru : infer (r != u)) : @exp R P [::] _ := - [let x := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in - let r := if #x then return {3}:R else return {10}:R in - let u := Score {exp_poisson 4 [#r]} in - return #x]. - -Fail Definition staton_bus_syntax0_generic' (x r u : string) - (rx : infer (r != x)) (Rx : infer (u != x)) - (ur : infer (u != r)) (xr : infer (x != r)) - (xu : infer (x != u)) (ru : infer (r != u)) : @exp R P [::] _ := - [let x := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in - let r := if $x then return {3}:R else return {10}:R in - let u := Score {exp_poisson 4 [$r]} in - return $x]. - -Fail Definition staton_bus_syntax0' : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in - let "r" := if ${"x"} then return {3}:R else return {10}:R in - let "_" := Score {exp_poisson 4 [${"r"}]} in - return ${"x"}]. - Definition staton_bus_syntax0 : @exp R _ [::] _ := [let "x" := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in let "r" := if #{"x"} then return {3}:R else return {10}:R in let "_" := Score {exp_poisson 4 [#{"r"}]} in return #{"x"}]. -End tests. - Definition staton_bus_syntax := [Normalize {staton_bus_syntax0}]. Let sample_bern : R.-sfker munit ~> mbool := sample_cst (bernoulli p27). @@ -810,6 +817,11 @@ Section letinC. Local Open Scope lang_scope. Variable (R : realType). +Require Import Classical_Prop. (* TODO: mv *) + +Let weak_head g {t1 t2} x (e : @exp R P g t2) (xg : x \notin dom g) := + exp_weak P [::] _ (x, t1) e xg. + Lemma letinC g t1 t2 (e1 : @exp R P g t1) (e2 : exp P g t2) (x y : string) (xy : infer (x != y)) (yx : infer (y != x)) @@ -817,11 +829,11 @@ Lemma letinC g t1 t2 (e1 : @exp R P g t1) (e2 : exp P g t2) forall U, measurable U -> execP [ let x := e1 in - let y := {exp_weak _ [::] _ (x, t1) e2 xg} in + let y := {weak_head e2 xg} in return (#x, #y)] ^~ U = execP [ let y := e2 in - let x := {exp_weak _ [::] _ (y, t2) e1 yg} in + let x := {weak_head e1 yg} in return (#x, #y)] ^~ U. Proof. move=> U mU; apply/funext => z. diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v new file mode 100644 index 0000000000..c2fc64319a --- /dev/null +++ b/theories/lang_syntax_examples_wip.v @@ -0,0 +1,329 @@ +Require Import String. +From HB Require Import structures. +From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. +From mathcomp.classical Require Import mathcomp_extra boolp classical_sets. +From mathcomp.classical Require Import functions cardinality fsbigop. +Require Import signed reals ereal topology normedtype sequences esum measure. +Require Import lebesgue_measure numfun lebesgue_integral kernel prob_lang. +Require Import lang_syntax_util lang_syntax. +From mathcomp Require Import ring lra. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import Order.TTheory GRing.Theory Num.Def Num.Theory. +Import numFieldTopology.Exports. + +Local Open Scope classical_set_scope. +Local Open Scope ring_scope. +(* Local Open Scope ereal_scope. *) + +Section casino_example. +Open Scope ring_scope. +Open Scope lang_scope. +Context (R : realType). +Lemma a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. + +Definition casino0 : exp _ [::] Bool := + [let "p" := Sample {exp_uniform 0 1 a01} in + let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:R then return TT else Score {0}:R in + let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in + return {1}:R <= #{"a2"}]. + +Example e1 : @exp R _ [::] _ := [{1}:R + {2}:R * {2}:R ^+ {3%nat}]. + +Lemma exe1 : projT1 (execD e1) = (fun x => 17). +Proof. +rewrite /e1 (@execD_bin _ _ binop_add) (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow/= !execD_real /=. +apply: funext => x /=. +lra. +Qed. + +(* Arguments exp_bin {R g b} &. *) +Definition casino1 : @exp R _ [::] _ := + [let "p" := Sample {exp_uniform 0 1 a01} in + let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:R then return TT else Score {0}:R in + let "a2" := Sample + {exp_bernoulli_trunc [{1}:R - #{"p"} ^+ {3%nat}]} in + return #{"a2"}]. + +(* exec [let x := sample (binomial n e) in *) +(* return (x >= 1)] = *) +(* exec [let y := sample (bernoulli (1 - e^n)) in *) +(* return y] *) +Lemma binomial_le1 x y n g e t U : + (0 <= (projT1 (execD e) t : mtyp Real) <= 1) -> + execP ([let x := Sample {exp_binomial_trunc n e} in + return {1}:R <= #x] : @exp R _ g _) t U = + execP [let y := Sample {exp_bernoulli_trunc [{1}:R - e ^+ n]} in + return #y] t U. +Proof. +rewrite !execP_letin !execP_sample execD_binomial_trunc execD_bernoulli_trunc/=. +rewrite !exp_var'E/=. + exact/ctx_prf_head. + exact/ctx_prf_head. +move=> H0 H1. +rewrite !execP_return !execD_rel. + have /= := @execD_var R ((y, Bool) :: g) y. + rewrite eqxx => /(_ H0) ->. + have /= := @execD_var R ((x, Real) :: g) x. + rewrite eqxx => /(_ H1) -> /=. +rewrite (@execD_bin _ _ binop_minus) execD_pow !execD_real/=. +rewrite 2!letin'E/=. +set p := projT1 (execD e) t. +move => /andP[p0 p1]. +rewrite (@integral_bernoulli_trunc _ _ (fun x => \d_x U))//; last first. + apply/andP; split. + exact: (onemX_ge0 _ p0 p1). + apply/onem_le1/exprn_ge0/p0. +rewrite (@integral_binomial_probabilty_trunc _ n p _ _ (fun y => \d_(1 <= y)%R U))//; last first. + (* move=>/= _ y0 my0. + rewrite setTI. *) + apply: measurable_fun_dirac. + have := @subsetT _ U; rewrite setT_bool => UT. + have [->|->|->|->] /= := subset_set2 UT. + exact: measurable0. + rewrite [X in measurable X](_ : _ = `[1, +oo[%classic) //. + apply/seteqP. + admit. + rewrite [X in measurable X](_ : _ = `]-oo, 1[%classic) //. + apply/seteqP. + admit. + admit. + (* rewrite diracE. *) +rewrite !big_ord_recl/=. +have -> : (1 <= 0 :> R) = false by lra. +rewrite /bump. +under eq_bigr => i _. + rewrite /= natrD. + have -> : 1 <= 1 + i%:R :> R. + by rewrite lerDl. + over. +rewrite addeC -ge0_sume_distrl. + congr (_ + _)%E; congr (_ * _)%E. + have -> : (\sum_(i < n) (p ^+ (n - (1 + i)) * `1-p ^+ (1 + i) *+ 'C(n, 1 + i))%:E)%E = + (\sum_(i < n.+1) (p ^+ (n - i) * `1-p ^+ i *+ 'C(n, i))%:E - (p ^+ n)%:E)%E. + rewrite big_ord_recl/= subn0 addeC addeA. + rewrite bin0 mulr1 mulr1n. + have <- : 0%E = ((- p ^+ n)%:E + (p ^+ n)%:E)%E. + rewrite EFinN. + congr _%:E. + lra. + by rewrite add0e. + congr _%E. + rewrite sumEFin. + rewrite !EFinB EFin_expe. + congr (_ - _)%E. + rewrite -(@exprDn_comm _ p `1-p n); last first. + by rewrite /GRing.comm/onem; lra. + rewrite /onem addrC. + have -> : 1 - p + p = 1 by lra. + by rewrite expr1n. + rewrite subn0 expr0 bin0 mulr1 mulr1n. + rewrite /onem. + congr _%:E. + set pn := p ^+ n. + lra. +move=> i _. +apply/mulrn_wge0/mulr_ge0; apply/exprn_ge0. +exact: p0. +apply/onem_ge0/p1. +Admitted. + +Lemma __ : uniform_probability a01 `[0, (1 / 2)] = (1 / 2)%:E. +Proof. +rewrite /uniform_probability /mscale/= /mrestr. +Abort. + +Lemma letin'_sample_uniform d d' (T : measurableType d) + (T' : measurableType d') (a b : R) (ab0 : (0 < b - a)%R) + (u : R.-sfker [the measurableType _ of (_ * T)%type] ~> T') x y : + letin' (sample_cst (uniform_probability ab0)) u x y = + (((b - a)^-1)%:E * \int[lebesgue_measure]_(x0 in `[a, b]) u (x0, x) y)%E. +Proof. +rewrite letin'E/=. +rewrite ge0_integral_mscale//=; last admit. +transitivity ( +(((b - a)^-1)%:E * \int[lebesgue_measure]_(x0 in `[a, b]) u (x0, x) y)%E +). + admit. +by []. +Admitted. + +Lemma execP_letin_uniform g u a b p (Hap : infer (a != p)) +(Hbp : infer (b != p)) (s0 s1 : @exp R D ((a, Unit) :: (b, Real) :: (p, Real) :: g) Real -> exp P ((p, Real) :: g) u) : + (forall (t : R) x U, t \in `[0, 1] -> execP (s0 [#p]) (t, x) U = execP (s1 [#p]) (t, x) U) -> + execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s0 [#p]}] = + execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s1 [#p]}]. +Proof. +move=> s01. +rewrite !execP_letin execP_sample execD_uniform/=. +apply: eq_sfkernel => x U. +rewrite 2!letin'_sample_uniform. +congr (_ * _)%E. +apply: eq_integral => t t01. +apply: s01. +by rewrite inE in t01. +Qed. + +Lemma casino01 : execP casino0 = execP casino1. +Proof. +rewrite /casino0 /casino1. +pose s0 := fun x => [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} + in let "_" := if #{"a1"} == {5}:R then return TT else Score {0}:R + in let "a2" := Sample {exp_binomial_trunc 3 [x]} in {exp_return [{1}:R <= #{"a2"}]}] : @exp R _ _ _. +pose s1 := fun x => [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} + in let "_" := if #{"a1"} == {5}:R then return TT else Score {0}:R + in let "a2" + := Sample {exp_bernoulli_trunc [{1}:R - x ^+ {3%N}]} + in {exp_return [#{"a2"}]}] : @exp R _ _ _. +apply: (@execP_letin_uniform [::] _ _ _ _ _ _ (s0 _) (s1 _)). +move=> e k k01 ek. +rewrite /s0/s1. +rewrite 2!execP_letin. +rewrite 2![in RHS]execP_letin. +congr (letin' _ _ _ _). +congr (letin' _ _ _). +(* apply: eq_sfkernel => x U. *) +rewrite !execP_letin !execP_sample execD_bernoulli_trunc execD_binomial_trunc/=. +rewrite (@execD_bin _ _ binop_minus) !execP_return execD_rel execD_pow/=. +rewrite !execD_real !exp_var'E !(execD_var_erefl "a2") (execD_var_erefl "p")/=. + +have H := (@binomial_le1 "a2" "a2" 3 _ [#{"p"}]). +(* f_equal. *) +(* congr ((letin' _ _ _) _). +apply: eq_sfkernel => ?. +have : 0 <= e <= 1. + admit. +move=> H. +rewrite (binomial_le1 _ _ _ _ H). +under eq_integral. +rewrite execP_letin. +congr (letin'). + +(* apply: eq_sfkernel => x U. *) +have : 0 <= (projT1 (execD e) x : mtyp Real) <= 1. + by rewrite ek. +move=> H. +apply: (binomial_le1 _ _ _ _ H). *) +Admitted. + +(* guard test *) +Definition test_guard : @exp R _ [::] _ := [ + let "p" := Sample {exp_bernoulli (1 / 3)%:nng (p1S 2)} in + let "_" := if #{"p"} then return TT else Score {0}:R in + return #{"p"} +]. + +Lemma exec_guard t U : execP test_guard t U = ((1 / 3)%:E * @dirac _ _ true R U)%E. +Proof. +rewrite /test_guard 2!execP_letin execP_sample execD_bernoulli execP_if/=. +rewrite !execP_return !exp_var'E !(execD_var_erefl "p") execD_unit execP_score execD_real/=. +rewrite letin'E ge0_integral_measure_sum//. +rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//= !integral_dirac//. +rewrite !letin'E !iteE/= integral_dirac// ge0_integral_mscale//=. +by rewrite normr0 mul0e !mule0 !adde0 !diracT !mul1e. +Qed. + +Lemma exec_casino t U : + execP casino0 t U = ((10 / 99)%:E * @dirac _ _ true R U + + (1 / 99)%:E * @dirac _ _ false R U)%E . +Proof. +rewrite /casino0 !execP_letin !execP_sample execD_uniform/=. +rewrite !execD_binomial_trunc execP_if !execP_return !execP_score/=. +rewrite !execD_rel !execD_real execD_unit/=. +rewrite !exp_var'E (execD_var_erefl "p") (execD_var_erefl "a1"). +rewrite (execD_var_erefl "p") (execD_var_erefl "a2")/=. +rewrite letin'E/= /uniform_probability ge0_integral_mscale//=. +rewrite subr0 invr1 mul1e. +under eq_integral. + +Admitted. + +Definition uniform_syntax : @exp R _ [::] _ := + [let "p" := Sample {exp_uniform 0 1 a01} in + return #{"p"}]. + +About integralT_nnsfun. + +Lemma exec_uniform_syntax t U : measurable U -> + execP uniform_syntax t U = uniform_probability a01 U. +Proof. +move=> mU. +rewrite /uniform_syntax execP_letin execP_sample execP_return !execD_uniform. +rewrite exp_var'E (execD_var_erefl "p")/=. +rewrite letin'E /=. +rewrite {1}/uniform_probability. +set x := (X in mscale _ X). +set k := (X in mscale X _). +transitivity ((k%:num)%:E * \int[x]_y \d_y U)%E. + rewrite -ge0_integral_mscale //. + exact: measurable_fun_dirac. +rewrite /uniform_probability /mscale /=. +congr (_ * _)%E. +under eq_integral do rewrite diracE. +rewrite /=. +rewrite /lebesgue_measure/=. +rewrite -integral_cst. +set f := (mrestr lebesgue_measure (measurable_itv `[0%R, 1%R])). +admit. +rewrite -(@ge0_integral_mscale _ _ _ x setT measurableT k (fun y => \d_y U)) //. +rewrite /uniform_probability/mscale/=/integral//=. +Admitted. + +Definition binomial_le : @exp R _ [::] Bool := + [let "a2" := Sample {exp_binomial 3 (1 / 2)%:nng (p1S 1)} in + return {1}:R <= #{"a2"}]. + +Lemma exec_binomial_le t U : + execP binomial_le t U = ((7 / 8)%:E * @dirac _ _ true R U + + (1 / 8)%:E * @dirac _ _ false R U)%E. +Proof. +rewrite /binomial_le execP_letin execP_sample execP_return execD_rel execD_real. +rewrite exp_var'E (execD_var_erefl "a2") execD_binomial. +rewrite letin'E//= /binomial_probability ge0_integral_measure_sum//=. +rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//=. +rewrite !integral_dirac// /bump. +rewrite !binS/= !bin0 bin1 bin2 bin_small// addn0. +rewrite addeC adde0. +congr (_ + _)%:E. +have -> : (1 <= 1)%R. admit. +have -> : (1 <= 2)%R. admit. +have -> : (1 <= 3)%R. admit. +rewrite -!mulrDl indicT !mul1r. +congr (_ * _). +rewrite /onem addn0 add0n. +by field. +congr (_ * _). +by field. +(* by rewrite ler10. *) +Admitted. + +Definition binomial_guard : @exp R _ [::] Real := + [let "a1" := Sample {exp_binomial 3 (1 / 2)%:nng (p1S 1)} in + let "_" := if #{"a1"} == {1}:R then return TT else Score {0}:R in + return #{"a1"}]. + +Lemma exec_binomial_guard t U : + execP binomial_guard t U = ((7 / 8)%:E * @dirac _ R 1%R R U + + (1 / 8)%:E * @dirac _ R 0%R R U)%E. +Proof. +rewrite /binomial_guard !execP_letin execP_sample execP_return execP_if. +rewrite !exp_var'E execD_rel !(execD_var_erefl "a1") execP_return execD_unit execD_binomial execD_real execP_score execD_real. +rewrite !letin'E//= /binomial_probability ge0_integral_measure_sum//=. +rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//=. +rewrite !integral_dirac// /bump. +rewrite (* indicT *) !binS/= !bin0 bin1 bin2 bin_small// addn0 (* !mul1e *). +rewrite !letin'E//= !iteE/= !diracE/=. +have -> : (0 == 1)%R = false; first by admit. +have -> : (1 == 1)%R; first by admit. +have -> : (2 == 1)%R = false; first by admit. +have -> : (3 == 1)%R = false; first by admit. +rewrite addeC adde0. +Admitted. + +End casino_example. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 13e976768f..f5f111292e 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -7,7 +7,7 @@ From mathcomp Require Import functions cardinality fsbigop. From mathcomp Require Import reals ereal topology normedtype sequences. From mathcomp Require Import esum measure lebesgue_measure numfun. From mathcomp Require Import lebesgue_integral exp kernel. -From mathcomp Require Import lra. +From mathcomp Require Import ring lra. (******************************************************************************) (* Semantics of a probabilistic programming language using s-finite kernels *) @@ -43,6 +43,26 @@ Local Open Scope classical_set_scope. Local Open Scope ring_scope. Local Open Scope ereal_scope. +(* Definition mR (R : realType) : Type := R. +HB.instance Definition _ (R : realType) := Measurable.on (mR R). +(* [the measurableType (R.-ocitv.-measurable).-sigma of + salgebraType (R.-ocitv.-measurable)]. *) *) + +Module Notations. +(*Notation var1of2 := (@measurable_fst _ _ _ _). +Notation var2of2 := (@measurable_snd _ _ _ _). +Notation var1of3 := (measurableT_comp (@measurable_fst _ _ _ _) + (@measurable_fst _ _ _ _)). +Notation var2of3 := (measurableT_comp (@measurable_snd _ _ _ _) + (@measurable_fst _ _ _ _)). +Notation var3of3 := (@measurable_snd _ _ _ _).*) + +(* Definition mR R := [the measurableType (R.-ocitv.-measurable).-sigma of + salgebraType (R.-ocitv.-measurable)]. *) +Notation munit := (unit : measurableType _). +Notation mbool := (bool : measurableType _). +End Notations. + (* TODO: PR *) Lemma onem_nonneg_proof (R : numDomainType) (p : {nonneg R}) : (p%:num <= 1 -> 0 <= `1-(p%:num))%R. @@ -86,6 +106,34 @@ subst p2. by f_equal. Qed. +Section constants. +Variable R : realType. +Local Open Scope ring_scope. + +Lemma onem1S n : `1- (1 / n.+1%:R) = (n%:R / n.+1%:R)%:nng%:num :> R. +Proof. +by rewrite /onem/= -{1}(@divrr _ n.+1%:R) ?unitfE// -mulrBl -natr1 addrK. +Qed. + +Lemma p1S n : (1 / n.+1%:R)%:nng%:num <= 1 :> R. +Proof. by rewrite ler_pdivrMr//= mul1r ler1n. Qed. + +Lemma p12 : (1 / 2%:R)%:nng%:num <= 1 :> R. Proof. by rewrite p1S. Qed. + +Lemma p14 : (1 / 4%:R)%:nng%:num <= 1 :> R. Proof. by rewrite p1S. Qed. + +Lemma onem27 : `1- (2 / 7%:R) = (5%:R / 7%:R)%:nng%:num :> R. +Proof. by apply/eqP; rewrite subr_eq/= -mulrDl -natrD divrr// unitfE. Qed. + +Lemma p27 : (2 / 7%:R)%:nng%:num <= 1 :> R. +Proof. by rewrite /= lter_pdivrMr// mul1r ler_nat. Qed. + +End constants. +Arguments p12 {R}. +Arguments p14 {R}. +Arguments p27 {R}. +Arguments p1S {R}. + Section bernoulli. Variables (R : realType) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). @@ -107,27 +155,230 @@ HB.instance Definition _ := End bernoulli. -Lemma integral_bernoulli {R : realType} - (p : {nonneg R}) (p1 : (p%:num <= 1)%R) (f : bool -> set bool -> _) U : - (forall x y, 0 <= f x y) -> - \int[bernoulli p1]_y (f y ) U = - p%:num%:E * f true U + (`1-(p%:num))%:E * f false U. +Lemma integral_bernoulli {R : realType} (p : {nonneg R}) (p1 : (p%:num <= 1)%R) + (f : bool -> \bar R) : (forall x, 0 <= f x) -> + \int[bernoulli p1]_y (f y) = + p%:num%:E * f true + (`1-(p%:num))%:E * f false. Proof. move=> f0. rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT 2!mul1e. Qed. -Section uniform_probability. +Section bernoulli_trunc. +Variables (R : realType). +Local Open Scope ring_scope. + +Lemma sumbool_ler (x y : R) : {(x <= y)%R} + {(x > y)%R}. +Proof. +have [_|_] := leP x y. +by apply left. +by apply right. +Qed. + +Definition bernoulli0 := @bernoulli R 0%R%:nng ler01. + +HB.instance Definition _ := Probability.on bernoulli0. + +Definition bernoulli_trunc (p : R) := match (sumbool_ler 0%R p) with +| left l0p => match (sumbool_ler (NngNum l0p)%:num 1%R) with + | left lp1 => [the probability _ _ of @bernoulli R (NngNum l0p) lp1] + | right _ => [the probability _ _ of bernoulli0] + end +| right _ => [the probability _ _ of bernoulli0] +end. + +HB.instance Definition _ (p : R) := Probability.on (bernoulli_trunc p). + +Lemma measurable_bernoulli_trunc : measurable_fun setT (bernoulli_trunc : _ -> pprobability _ _). +Proof. +(* move=> _ Y mY. *) +(* rewrite setTI. *) +apply: (@measurability _ _ _ _ _ _ + (@pset _ _ _ : set (set (pprobability _ R)))) => //. +move=> _ -[_ [r r01] [Ys mYs <-]] <-. +apply: emeasurable_fun_infty_o => //. +rewrite /bernoulli_trunc. +(* apply: measurable_fun_case_suml. +case: sumbool_ler. *) +(* case: (sumbool_ler 0 p) => [p0'/=|]. + case: (sumbool_ler p 1) => [p1'/=|]. *) +(* exact: (measurable_kernel (knormalize f point) Ys). *) +Admitted. + +End bernoulli_trunc. + +Arguments bernoulli_trunc {R}. +Arguments measurable_bernoulli_trunc {R}. + +Lemma integral_bernoulli_trunc {R : realType} (p : R) (f : bool -> \bar R) : + (0 <= p <= 1)%R -> + (forall x, 0 <= f x) -> + \int[bernoulli_trunc p]_y (f y) = + p%:E * f true + (`1-p)%:E * f false. +Proof. +move=> /andP[p0 p1] f0. +rewrite /bernoulli_trunc. +case: (sumbool_ler 0 p) => [p0'/=|]. + case: (sumbool_ler p 1) => [p1'/=|]. + by rewrite integral_bernoulli. + by rewrite ltNge p1. +by rewrite ltNge p0. +Qed. + +Section binomial_probability. +Context {R : realType} (n : nat) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). +Local Open Scope ring_scope. + +(* C(n, k) * p^(n-k) * (1-p)^k *) +Definition bino_term (k : nat) :{nonneg R} := + (p%:num^+(n-k)%N * (NngNum (onem_ge0 p1))%:num^+k *+ 'C(n, k))%:nng. + +Lemma bino_term0 : + bino_term 0 = (p%:num^+n)%:nng. +Proof. +rewrite /bino_term bin0 subn0/=. +apply/val_inj => /=. +by field. +Qed. + +Lemma bino_term1 : + bino_term 1 = (p%:num^+(n-1)%N * (NngNum (onem_ge0 p1))%:num *+ n)%:nng. +Proof. +rewrite /bino_term bin1/=. +apply/val_inj => /=. +by rewrite expr1. +Qed. + +(* \sum_(k < n.+1) (bino_coef p n k) * \d_k. *) +Definition binomial_probability : set (measurableTypeR R) -> \bar R := + @msum _ _ R + (fun k => [the measure _ _ of mscale (bino_term k) + [the measure _ _ of @dirac _ (measurableTypeR R) k%:R R]]) n.+1. + +HB.instance Definition _ := Measure.on binomial_probability. + +Let binomial_setT : binomial_probability [set: _] = 1%:E. +Proof. +rewrite /binomial_probability/msum/mscale/bino_term/=/mscale/=. +under eq_bigr do rewrite diracT mule1. +rewrite sumEFin. +rewrite -exprDn_comm; last by rewrite /GRing.comm mulrC. +by rewrite add_onemK; congr _%:E; rewrite expr1n. +Qed. + +HB.instance Definition _ := + @Measure_isProbability.Build _ _ R binomial_probability binomial_setT. + +End binomial_probability. + +Section integral_binomial. +Variables (R : realType) (d : measure_display) (T : measurableType d). + +Lemma integral_binomial (n : nat) (p : {nonneg R}) + (p1 : (p%:num <= 1)%R) (f : R -> \bar R) + (mf : measurable_fun setT f) : + (forall x, 0 <= f x) -> \int[binomial_probability n p1]_y (f y) = + \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k%:R. +Proof. +move=> f0. +rewrite ge0_integral_measure_sum//=; last first. +apply: eq_bigr => i _. +rewrite ge0_integral_mscale//=; last first. +rewrite integral_dirac//=; last first. +by rewrite diracT mul1e. +Qed. + +End integral_binomial. + +Section binomial_trunc. +Variables (R : realType). +Local Open Scope ring_scope. + +Definition binomial_probability0 := @binomial_probability R 0 0%:nng%R ler01. + +Definition binomial_probability_trunc n (p : R) := + (* (p1 : (p%:num <= 1)%R) := *) + match (sumbool_ler 0%R p) with + | left l0p => match (sumbool_ler (NngNum l0p)%:num 1%R) with + | left lp1 => [the probability _ _ of @binomial_probability R n (NngNum l0p) lp1] + | right _ => [the probability _ _ of binomial_probability0] + end + | right _ => [the probability _ _ of binomial_probability0] + end. + +Lemma measurable_binomial_probability_trunc (n : nat) + : measurable_fun setT (binomial_probability_trunc n : R -> pprobability _ _). +Admitted. + +End binomial_trunc. + +Arguments binomial_probability_trunc {R}. +Arguments measurable_binomial_probability_trunc {R}. + +Section integral_binomial_trunc. +Variables (R : realType) (d : measure_display) (T : measurableType d). + +Lemma integral_binomial_probabilty_trunc (n : nat) (p : R) + (p0 : (0 <= p)%R) (p1 : ((NngNum p0)%:num <= 1)%R) (f : R -> \bar R) (mf : measurable_fun setT f) : + (forall x, 0 <= f x) -> \int[binomial_probability_trunc n p]_y (f y) = \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k%:R. +Proof. +move=> f0. +rewrite /binomial_probability_trunc/=. +case: (sumbool_ler 0 p) => [p0'/=|]. + case: (sumbool_ler p 1) => [p1'/=|]. + by rewrite integral_binomial. + by rewrite ltNge p1. +by rewrite ltNge p0. +Qed. + +End integral_binomial_trunc. + +Section binomial_example. Context {R : realType}. -Variables (a b : R) (ab0 : (0 < b - a)%R). +Open Scope ring_scope. + +Lemma binomial3_2 : @binomial_probability R 3 _ (p1S 1) [set 2%:R] = (3 / 8)%:E. +Proof. +rewrite /binomial_probability/msum !big_ord_recl/= big_ord0 adde0 bino_term0. +rewrite /mscale/= !diracE /bump/=. +repeat rewrite ?binS ?bin0 ?bin1 ?bin_small//. +rewrite memNset//=; last by move/eqP; rewrite eqr_nat. +rewrite memNset//=; last by move/eqP; rewrite eqr_nat. +rewrite mem_set//=. +rewrite memNset//=; last by move/eqP; rewrite eqr_nat. +congr _%:E. +rewrite expr0 !mulr1 !mulr0 !add0r !addn0 !add0n /onem. +by field. +Qed. -Definition uniform_probability := mscale (invr_nonneg (NngNum (ltW ab0))) - (mrestr lebesgue_measure (measurable_itv `[a, b])). +End binomial_example. + +Section uniform_probability. +Context (R : realType) (a b : R) (ab0 : (0 < b - a)%R). +Definition uniform_probability +(* : set _ -> \bar R *) + := @mscale _ _ R (invr_nonneg (NngNum (ltW ab0))) + (mrestr lebesgue_measure (measurable_itv `[a, b])). + +(** TODO: set R -> \bar R を書くとMeasure.onが通らない **) +(** **) HB.instance Definition _ := Measure.on uniform_probability. -Let uniform_probability_setT : uniform_probability [set: _] = 1. +(* Let uniform0 : uniform_probability set0 = 0. +Proof. exact: measure0. Qed. + +Let uniform_ge0 U : 0 <= uniform_probability U. +Proof. exact: measure_ge0. Qed. + +Let uniform_sigma_additive : semi_sigma_additive uniform_probability. +Proof. move=> /= F mF tF mUF; exact: measure_semi_sigma_additive. Qed. + +HB.instance Definition _ := isMeasure.Build _ _ _ uniform_probability + uniform0 uniform_ge0 uniform_sigma_additive. *) + +Let uniform_probability_setT : uniform_probability [set: _] = 1%:E. Proof. rewrite /uniform_probability /mscale/= /mrestr/=. rewrite setTI lebesgue_measure_itv/= lte_fin. @@ -614,22 +865,6 @@ Proof. by rewrite /fail letinE ge0_integral_mscale//= normr0 mul0e. Qed. End hard_constraint. Arguments fail {d d' X Y R}. -Module Notations. - -(*Notation var1of2 := (@measurable_fst _ _ _ _). -Notation var2of2 := (@measurable_snd _ _ _ _). -Notation var1of3 := (measurableT_comp (@measurable_fst _ _ _ _) - (@measurable_fst _ _ _ _)). -Notation var2of3 := (measurableT_comp (@measurable_snd _ _ _ _) - (@measurable_fst _ _ _ _)). -Notation var3of3 := (@measurable_snd _ _ _ _).*) - -Notation mR := measurableTypeR. -Notation munit := (unit : measurableType _). -Notation mbool := (bool : measurableType _). - -End Notations. - Section cst_fun. Context d (T : measurableType d) (R : realType). @@ -1048,35 +1283,6 @@ Qed. End letinC. (* sample programs *) - -Section constants. -Variable R : realType. -Local Close Scope ereal_scope. - -Lemma onem1S n : `1- (1 / n.+1%:R) = (n%:R / n.+1%:R)%:nng%:num :> R. -Proof. -by rewrite /onem/= -{1}(@divrr _ n.+1%:R) ?unitfE// -mulrBl -natr1 addrK. -Qed. - -Lemma p1S n : (1 / n.+1%:R)%:nng%:num <= 1 :> R. -Proof. by rewrite ler_pdivrMr//= mul1r ler1n. Qed. - -Lemma p12 : (1 / 2%:R)%:nng%:num <= 1 :> R. Proof. by rewrite p1S. Qed. - -Lemma p14 : (1 / 4%:R)%:nng%:num <= 1 :> R. Proof. by rewrite p1S. Qed. - -Lemma onem27 : `1- (2 / 7%:R) = (5%:R / 7%:R)%:nng%:num :> R. -Proof. by apply/eqP; rewrite subr_eq/= -mulrDl -natrD divrr// unitfE. Qed. - -Lemma p27 : (2 / 7%:R)%:nng%:num <= 1 :> R. -Proof. by rewrite /= lter_pdivrMr// mul1r ler_nat. Qed. - -End constants. -Arguments p12 {R}. -Arguments p14 {R}. -Arguments p27 {R}. -Arguments p1S {R}. - Section poisson. Variable R : realType. Local Open Scope ring_scope. @@ -1171,7 +1377,7 @@ Context d (T : measurableType d) (R : realType). Definition sample_and_branch : R.-sfker T ~> R := letin (sample_cst [the probability _ _ of bernoulli p27]) (* T -> B *) - (ite macc1of2 (ret k3) (ret k10)). + (ite macc1of2 (ret (@k3 _ _ R)) (ret k10)). Lemma sample_and_branchE t U : sample_and_branch t U = (2 / 7%:R)%:E * \d_(3%:R : R) U + @@ -1190,6 +1396,28 @@ Definition bernoulli_and : R.-sfker T ~> mbool := (letin (sample_cst [the probability _ _ of bernoulli p12]) (letin (sample_cst [the probability _ _ of bernoulli p12]) (ret (measurable_and macc1of3 macc2of3)))). +(* +Definition mand (x y : T * mbool * mbool -> mbool) + (t : T * mbool * mbool) : mbool := x t && y t. + +Lemma measurable_fun_mand (x y : T * mbool * mbool -> mbool) : + measurable_fun setT x -> measurable_fun setT y -> + measurable_fun setT (mand x y). +Proof. +move=> /= mx my; apply: (measurable_fun_bool true). +rewrite [X in measurable X](_ : _ = + (x @^-1` [set true]) `&` (y @^-1` [set true])); last first. + by rewrite /mand; apply/seteqP; split => z/= /andP. +apply: measurableI. +- by rewrite -[X in measurable X]setTI; exact: mx. +- by rewrite -[X in measurable X]setTI; exact: my. +Qed. + +Definition bernoulli_and : R.-sfker T ~> mbool := + (letin (sample_cst [the probability _ _ of bernoulli p12]) + (letin (sample_cst [the probability _ _ of bernoulli p12]) + (ret (measurable_fun_mand macc1of3 macc2of3)))). +*) Lemma bernoulli_andE t U : bernoulli_and t U = @@ -1262,7 +1490,7 @@ Lemma staton_busE P (t : R) U : ((2 / 7%:R)%:E * (poisson4 3%:R)%:E * \d_true U + (5%:R / 7%:R)%:E * (poisson4 10%:R)%:E * \d_false U) * N^-1%:E. Proof. -rewrite /staton_bus normalizeE /= !kstaton_bus_poissonE !diracT !mule1 ifF //. +rewrite /staton_bus normalizeE !kstaton_bus_poissonE !diracT !mule1 ifF //. apply/negbTE; rewrite gt_eqF// lte_fin. by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// poisson_gt0// ltr0n. Qed. @@ -1883,7 +2111,7 @@ Definition von_neumann_trick' : R.-sfker (T * unit) ~> (unit + bool)%type := letin (sample_cst D) (letin (sample_cst D) (letin (lift_neq macc1of3 macc2of3) - (ite (macc3of4) + (ite macc3of4 (letin (ret macc1of4) (ret minlb)) (ret kinrtt)))). diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 1ac606f0a1..79d2c24888 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -124,7 +124,7 @@ apply: (measurable_comp (F := [set r : R | r != 0%R])) => //. - exact: measurable_fun_gauss_density. Qed. -Variable mu : {measure set mR R -> \bar R}. +Variable mu : {measure set R -> \bar R}. Definition staton_lebesgue : R.-sfker T ~> _ := letin (sample_cst (gauss01 integral_gauss01_density : pprobability _ _)) From fcea1d0a65e40b508cc164298a6e8670f715372a Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 9 Jan 2024 19:29:56 +0900 Subject: [PATCH 07/48] bernoulli_trunc measurable --- theories/prob_lang.v | 84 ++++++++++++++++++++++++++++++++++++-------- 1 file changed, 70 insertions(+), 14 deletions(-) diff --git a/theories/prob_lang.v b/theories/prob_lang.v index f5f111292e..59e2d0b1ea 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -176,35 +176,91 @@ by apply left. by apply right. Qed. +(* TODO: move? *) +Definition dep_uncurry := +(fun (A : Type) (B : A -> Type) (C : Type) (f : forall a : A, B a -> C) (p : {a : A & B a}) => let (a, Ba) := p in f a Ba) : +forall [A : Type] [B : A -> Type] [C : Type], +(forall a : A, B a -> C) -> {a : A & B a} -> C. + Definition bernoulli0 := @bernoulli R 0%R%:nng ler01. HB.instance Definition _ := Probability.on bernoulli0. -Definition bernoulli_trunc (p : R) := match (sumbool_ler 0%R p) with -| left l0p => match (sumbool_ler (NngNum l0p)%:num 1%R) with +Definition bernoulli_trunc (p : R) := match sumbool_ler 0%R p with +| left l0p => match sumbool_ler (NngNum l0p)%:num 1%R with | left lp1 => [the probability _ _ of @bernoulli R (NngNum l0p) lp1] | right _ => [the probability _ _ of bernoulli0] end | right _ => [the probability _ _ of bernoulli0] end. -HB.instance Definition _ (p : R) := Probability.on (bernoulli_trunc p). +(*HB.instance Definition _ (p : R) := Probability.on (bernoulli_trunc p).*) + +Let simpe := (@mule0 R, @adde0 R, @mule1 R, @add0e R). -Lemma measurable_bernoulli_trunc : measurable_fun setT (bernoulli_trunc : _ -> pprobability _ _). +Lemma measurable_bernoulli_trunc : + measurable_fun setT (bernoulli_trunc : _ -> pprobability _ _). Proof. -(* move=> _ Y mY. *) -(* rewrite setTI. *) apply: (@measurability _ _ _ _ _ _ (@pset _ _ _ : set (set (pprobability _ R)))) => //. move=> _ -[_ [r r01] [Ys mYs <-]] <-. -apply: emeasurable_fun_infty_o => //. -rewrite /bernoulli_trunc. -(* apply: measurable_fun_case_suml. -case: sumbool_ler. *) -(* case: (sumbool_ler 0 p) => [p0'/=|]. - case: (sumbool_ler p 1) => [p1'/=|]. *) -(* exact: (measurable_kernel (knormalize f point) Ys). *) -Admitted. +apply: emeasurable_fun_infty_o => //=. +rewrite /bernoulli_trunc/=. +have := @subsetT _ Ys; rewrite setT_bool => UT. +have [->|->|->|->] /= := subset_set2 UT. +- rewrite [X in measurable_fun _ X](_ : _ = (fun=> 0%E))//. + apply/funext => x/=. + by case: sumbool_ler. +- rewrite [X in measurable_fun _ X](_ : _ = (fun x => if 0 <= x <= 1 then x%:E else 0%E))//. + apply: measurable_fun_ifT => //=; apply: measurable_and => //; apply: (measurable_fun_bool true) => //=. + rewrite (_ : _ @^-1` _ = `[0, +oo[%classic)//. + by apply/seteqP; split => [x|x] /=; rewrite in_itv/= andbT. + by rewrite (_ : _ @^-1` _ = `]-oo, 1]%classic). + apply/funext => x/=; case: sumbool_ler => /= x0. + case: sumbool_ler => /= x1. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. + by rewrite mem_set//= memNset//= ?simpe x0 x1. + rewrite /bernoulli0. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. + by rewrite mem_set//= memNset//= ?simpe x0/= leNgt x1. + rewrite /bernoulli0. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. + by rewrite mem_set//= memNset//= ?simpe leNgt x0. +- rewrite [X in measurable_fun _ X](_ : _ = (fun x => if 0 <= x <= 1 then (`1-x)%:E else 1%E))//. + apply: measurable_fun_ifT => //=. + apply: measurable_and => //; apply: (measurable_fun_bool true) => //=. + rewrite (_ : _ @^-1` _ = `[0, +oo[%classic)//. + by apply/seteqP; split => [x|x] /=; rewrite in_itv/= andbT. + by rewrite (_ : _ @^-1` _ = `]-oo, 1]%classic). + by apply/EFin_measurable_fun; apply/measurable_funB. + apply/funext => x/=; case: sumbool_ler => /= x0. + case: sumbool_ler => /= x1. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. + by rewrite memNset//= mem_set//= ?simpe x0 x1/=. + rewrite /bernoulli0. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. + by rewrite memNset//= mem_set//= ?simpe x0/= leNgt x1/= onem0. + rewrite /bernoulli0. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. + by rewrite memNset//= mem_set//= leNgt x0/= ?simpe onem0. +- rewrite [X in measurable_fun _ X](_ : _ = (fun=> 1%E))//; apply/funext => x/=. + case: sumbool_ler => /= x0. + case: sumbool_ler => /= x1. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. + rewrite mem_set//=; last by left. + rewrite mem_set//=; last by right. + by rewrite ?simpe -EFinD add_onemK. + rewrite /bernoulli0. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. + rewrite mem_set//=; last by left. + rewrite mem_set//=; last by right. + by rewrite ?simpe onem0. + rewrite /bernoulli0. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. + rewrite mem_set//=; last by left. + rewrite mem_set//=; last by right. + by rewrite ?simpe onem0. +Qed. End bernoulli_trunc. From 6d0e67e8fdbdc13441c0e782af25357d6cf95938 Mon Sep 17 00:00:00 2001 From: AyumuSaito Date: Fri, 12 Jan 2024 17:43:09 +0900 Subject: [PATCH 08/48] add Nat and 0<=p<=1 problem - also fix casino1 --- theories/lang_syntax.v | 48 +++-- theories/lang_syntax_examples.v | 29 +-- theories/lang_syntax_examples_wip.v | 302 ++++++++++++++++++++++------ theories/prob_lang.v | 68 ++++--- 4 files changed, 320 insertions(+), 127 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index d11bc21425..0d1c28eb14 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -266,7 +266,7 @@ Import Notations. Context {R : realType}. Inductive typ := -| Unit | Bool | Real +| Unit | Bool | Nat | Real | Pair : typ -> typ -> typ | Prob : typ -> typ. @@ -276,6 +276,7 @@ Fixpoint measurable_of_typ (t : typ) : {d & measurableType d} := match t with | Unit => existT _ _ munit | Bool => existT _ _ mbool + | Nat => existT _ _ (nat : measurableType _) | Real => existT _ _ [the measurableType _ of (@measurableTypeR R)] (* (Real_sort__canonical__measure_Measurable R) *) @@ -407,22 +408,22 @@ Section relop. Inductive relop := | relop_le | relop_lt | relop_eq . -Definition fun_of_relop g (r : relop) : (@mctx R g -> @mtyp R Real) -> - (mctx g -> mtyp Real) -> @mctx R g -> @mtyp R Bool := +Definition fun_of_relop g (r : relop) : (@mctx R g -> @mtyp R Nat) -> + (mctx g -> mtyp Nat) -> @mctx R g -> @mtyp R Bool := match r with -| relop_le => (fun f1 f2 x => (f1 x <= f2 x)%R) -| relop_lt => (fun f1 f2 x => (f1 x < f2 x)%R) -| relop_eq => (fun f1 f2 x => (f1 x == f2 x)%R) +| relop_le => (fun f1 f2 x => (f1 x <= f2 x)%N) +| relop_lt => (fun f1 f2 x => (f1 x < f2 x)%N) +| relop_eq => (fun f1 f2 x => (f1 x == f2 x)%N) end. Definition mfun_of_relop g r - (f1 : @mctx R g -> @mtyp R Real) (mf1 : measurable_fun setT f1) - (f2 : @mctx R g -> @mtyp R Real) (mf2 : measurable_fun setT f2) : + (f1 : @mctx R g -> @mtyp R Nat) (mf1 : measurable_fun setT f1) + (f2 : @mctx R g -> @mtyp R Nat) (mf2 : measurable_fun setT f2) : measurable_fun [set: @mctx R g] (fun_of_relop r f1 f2). destruct r. -exact: measurable_fun_ler. -exact: measurable_fun_ltr. -exact: measurable_fun_eqr. +exact: measurable_fun_leq. +exact: measurable_fun_ltn. +exact: measurable_fun_eqn. Defined. End relop. @@ -430,12 +431,13 @@ End relop. Inductive exp : flag -> ctx -> typ -> Type := | exp_unit g : exp D g Unit | exp_bool g : bool -> exp D g Bool +| exp_nat g : nat -> exp D g Nat | exp_real g : R -> exp D g Real | exp_pow g : nat -> exp D g Real -> exp D g Real | exp_bin (b : binop) g : exp D g (type_of_binop b) -> exp D g (type_of_binop b) -> exp D g (type_of_binop b) -| exp_rel (r : relop) g : exp D g Real -> - exp D g Real -> exp D g Bool +| exp_rel (r : relop) g : exp D g Nat -> + exp D g Nat -> exp D g Bool | exp_pair g t1 t2 : exp D g t1 -> exp D g t2 -> exp D g (Pair t1 t2) | exp_proj1 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t1 | exp_proj2 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t2 @@ -445,9 +447,9 @@ Inductive exp : flag -> ctx -> typ -> Type := | exp_bernoulli_trunc g : exp D g Real -> exp D g (Prob Bool) | exp_binomial g (n : nat) (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : - exp D g (Prob Real) + exp D g (Prob Nat) | exp_binomial_trunc g (n : nat) : - exp D g Real -> exp D g (Prob Real) + exp D g Real -> exp D g (Prob Nat) | exp_uniform g (a b : R) (ab0 : (0 < b - a)%R) : exp D g (Prob Real) | exp_poisson g : nat -> exp D g Real -> exp D g Real | exp_normalize g t : exp P g t -> exp D g (Prob t) @@ -473,6 +475,7 @@ End syntax_of_expressions. Arguments exp {R}. Arguments exp_unit {R g}. Arguments exp_bool {R g}. +Arguments exp_nat {R g}. Arguments exp_real {R g}. Arguments exp_pow {R g}. Arguments exp_bin {R} b {g} &. @@ -499,6 +502,8 @@ Notation "[ e ]" := e (e custom expr at level 5) : lang_scope. Notation "'TT'" := (exp_unit) (in custom expr at level 1) : lang_scope. Notation "b ':B'" := (@exp_bool _ _ b%bool) (in custom expr at level 1) : lang_scope. +Notation "n ':N'" := (@exp_nat _ _ n%N) + (in custom expr at level 1) : lang_scope. Notation "r ':R'" := (@exp_real _ _ r%R) (in custom expr at level 1, format "r :R") : lang_scope. Notation "e ^+ n" := (exp_pow n e) @@ -557,6 +562,7 @@ Fixpoint free_vars k g t (e : @exp R k g t) : seq string := match e with | exp_unit _ => [::] | exp_bool _ _ => [::] + | exp_nat _ _ => [::] | exp_real _ _ => [::] | exp_pow _ _ e => free_vars e | exp_bin _ _ e1 e2 => free_vars e1 ++ free_vars e2 @@ -694,6 +700,8 @@ Inductive evalD : forall g t, exp D g t -> | eval_bool g b : ([b:B] : exp D g _) -D> cst b ; kb b +| eval_nat g n : ([n:N] : exp D g _) -D> cst n; kn n + | eval_real g r : ([r:R] : exp D g _) -D> cst r ; kr r | eval_pow g n (e : exp D g _) f mf : e -D> f ; mf -> @@ -824,6 +832,9 @@ all: (rewrite {g t e u v mu mv hu}). - move=> g b {}v {}mv. inversion 1; subst g0 b0. by inj_ex H3. +- move=> g n {}v {}mv. + inversion 1; subst g0 n0. + by inj_ex H3. - move=> g r {}v {}mv. inversion 1; subst g0 r0. by inj_ex H3. @@ -974,6 +985,9 @@ all: rewrite {g t e u v eu}. - move=> g b {}v {}mv. inversion 1; subst g0 b0. by inj_ex H3. +- move=> g n {}v {}mv. + inversion 1; subst g0 n0. + by inj_ex H3. - move=> g r {}v {}mv. inversion 1; subst g0 r0. by inj_ex H3. @@ -1120,6 +1134,7 @@ elim: e. all: rewrite {z g t}. - by do 2 eexists; exact: eval_unit. - by do 2 eexists; exact: eval_bool. +- by do 2 eexists; exact: eval_nat. - by do 2 eexists; exact: eval_real. - move=> g n e [f [mf H]]. by exists (fun x => (f x ^+ n)%R); eexists; exact: eval_pow. @@ -1239,6 +1254,9 @@ Proof. exact/execD_evalD/eval_unit. Qed. Lemma execD_bool g b : @execD g _ [b:B] = existT _ (cst b) (kb b). Proof. exact/execD_evalD/eval_bool. Qed. +Lemma execD_nat g n : @execD g _ [n:N] = existT _ (cst n) (kn n). +Proof. exact/execD_evalD/eval_nat. Qed. + Lemma execD_real g r : @execD g _ [r:R] = existT _ (cst r) (kr r). Proof. exact/execD_evalD/eval_real. Qed. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index ec7e244c79..77e7a78dda 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -460,37 +460,22 @@ Definition sample_binomial3 : @exp R _ [::] _ := return #{"x"}]. Lemma exec_sample_binomial3 t U : measurable U -> - execP sample_binomial3 t U = ((1 / 8)%:E * \d_(0%:R : R) U + - (3 / 8)%:E * \d_(1%:R : R) U + - (3 / 8)%:E * \d_(2%:R : R) U + - (1 / 8)%:E * \d_(3%:R : R) U)%E. + execP sample_binomial3 t U = ((1 / 8)%:E * \d_0%N U + + (3 / 8)%:E * \d_1%N U + + (3 / 8)%:E * \d_2%N U + + (1 / 8)%:E * \d_3%N U)%E. Proof. move=> mU; rewrite /sample_binomial3 execP_letin execP_sample execP_return. rewrite exp_var'E (execD_var_erefl "x") !execD_binomial/=. -rewrite letin'E ge0_integral_measure_sum//=; last first. - exact: measurable_fun_dirac. +rewrite letin'E ge0_integral_measure_sum//=. rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//=; [|exact: measurable_fun_dirac..]. rewrite !integral_dirac// /bump; [|exact: measurable_fun_dirac..]. rewrite !binS/= !bin0 bin1 bin2 bin_small// addn0. rewrite expr0 mulr1 mul1r subn0. rewrite -2!addeA !mul1r. congr _%:E. -rewrite indicT !mul1r. -congr (_ + _). - congr (_ * _). - by field. -congr (_ + _). - congr (_ * _). - rewrite expr1 /onem. - by field. -congr (_ + _). - congr (_ * _). - rewrite /onem/=. - by field. -rewrite addr0. -congr (_ * _). -rewrite /onem/=. -by field. +rewrite indicT !mul1r /onem !addrA addr0 expr1/=. +by congr (_ + _ + _ + _); congr (_ * _); field. Qed. End binomial_examples. diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index c2fc64319a..9e0c7b1198 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -25,12 +25,15 @@ Open Scope lang_scope. Context (R : realType). Lemma a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. +(* Definition ex : exp _ [::] _ := @exp_bernoulli R [::] (1 / 2)%:nng (p1S 1). +Example ex1 : projT1 (execD ex) tt = 1%:E. *) + Definition casino0 : exp _ [::] Bool := [let "p" := Sample {exp_uniform 0 1 a01} in let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:R then return TT else Score {0}:R in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in - return {1}:R <= #{"a2"}]. + return {1}:N <= #{"a2"}]. Example e1 : @exp R _ [::] _ := [{1}:R + {2}:R * {2}:R ^+ {3%nat}]. @@ -46,20 +49,136 @@ Qed. Definition casino1 : @exp R _ [::] _ := [let "p" := Sample {exp_uniform 0 1 a01} in let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:R then return TT else Score {0}:R in - let "a2" := Sample - {exp_bernoulli_trunc [{1}:R - #{"p"} ^+ {3%nat}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + let "a2" := + Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]} in return #{"a2"}]. +Definition casino1' : @exp R _ [::] _ := + [let "p" := Sample {exp_uniform 0 1 a01} in + let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. + +Lemma bernoulli_truncE (p : R) U : + (0 <= p <= 1)%R -> + (bernoulli_trunc p U = + p%:E * \d_true U + (`1-p)%:E * \d_false U)%E. +Proof. +move=> /andP[p0 p1]. +rewrite /bernoulli_trunc. +case: (sumbool_ler 0 p) => [{}p0/=|]. + case: (sumbool_ler p 1) => [{}p1/=|]. + by rewrite /bernoulli/= measure_addE. + by rewrite ltNge p1. +by rewrite ltNge p0. +Qed. + +Lemma binomial_le1' n p U : + 0 <= p <= 1 -> + (\int[binomial_probability_trunc n p]_y0 \d_(0 < y0)%N U = + bernoulli_trunc (1 - `1-p ^+ n) U :> \bar R)%E. +Proof. +move=> /andP[p0 p1]. +rewrite bernoulli_truncE; last first. + apply/andP; split. + apply/onemX_ge0; rewrite /onem; lra. + apply/onem_le1/exprn_ge0; rewrite /onem; lra. +rewrite (@integral_binomial_probabilty_trunc _ n p _ _ (fun y => \d_(1 <= y)%N U))//; last first. +rewrite !big_ord_recl/=. +rewrite /bump. +under eq_bigr => i _. + rewrite /=. + have -> : (0 < 1 + i)%N => //. + over. +rewrite addeC -ge0_sume_distrl. + congr (_ + _)%E; congr (_ * _)%E. + have -> : (\sum_(i < n) (p ^+ (1 + i) * `1-p ^+ (n - (1 + i)) *+ 'C(n, 1 + i))%:E)%E = + (\sum_(i < n.+1) (p ^+ i * `1-p ^+ (n - i) *+ 'C(n, i))%:E - (`1-p ^+ n)%:E)%E. + rewrite big_ord_recl/= expr0 subn0 mul1r bin0 mulr1n addeC addeA. + have <- : 0%E = ((- `1-p ^+ n)%:E + (`1-p ^+ n)%:E)%E. + rewrite EFinN. + congr _%:E. + lra. + by rewrite add0e. + congr _%E. + rewrite sumEFin. + rewrite !EFinB EFin_expe. + congr (_ - _)%E. + under eq_bigr do rewrite mulrC. + rewrite -(@exprDn_comm _ `1-p p n); last first. + by rewrite /GRing.comm/onem; lra. + rewrite /onem addrC. + have -> : p + (1 - p) = 1 by lra. + by rewrite expr1n. + rewrite subn0 expr0 bin0 mulr1n. + rewrite /onem. + congr _%:E. + set pn := (1-p) ^+ n. + lra. +move=> i _. +apply/mulrn_wge0/mulr_ge0; apply/exprn_ge0. +exact: p0. +apply/onem_ge0/p1. +Qed. + +Lemma binomial_le1 n p U : + 0 <= p <= 1 -> + (\int[binomial_probability_trunc n p]_y0 \d_(0 < y0)%N U = + \int[bernoulli_trunc (1 - `1-p ^+ n)]_y0 \d_y0 U :> \bar R)%E. +Proof. +move=> /andP[p0 p1]. +rewrite (@integral_bernoulli_trunc _ _ (fun x => \d_x U))//; last first. + apply/andP; split. + apply: onemX_ge0; rewrite /onem; lra. + apply/onem_le1/exprn_ge0; rewrite /onem; lra. +rewrite (@integral_binomial_probabilty_trunc _ n p _ _ (fun y => \d_(1 <= y)%N U))//; last first. +rewrite !big_ord_recl/=. +rewrite /bump. +under eq_bigr => i _. + rewrite /=. + have -> : (0 < 1 + i)%N => //. + over. +rewrite addeC -ge0_sume_distrl. + congr (_ + _)%E; congr (_ * _)%E. + have -> : (\sum_(i < n) (p ^+ (1 + i) * `1-p ^+ (n - (1 + i)) *+ 'C(n, 1 + i))%:E)%E = + (\sum_(i < n.+1) (p ^+ i * `1-p ^+ (n - i) *+ 'C(n, i))%:E - (`1-p ^+ n)%:E)%E. + rewrite big_ord_recl/= expr0 subn0 mul1r bin0 mulr1n addeC addeA. + have <- : 0%E = ((- `1-p ^+ n)%:E + (`1-p ^+ n)%:E)%E. + rewrite EFinN. + congr _%:E. + lra. + by rewrite add0e. + congr _%E. + rewrite sumEFin. + rewrite !EFinB EFin_expe. + congr (_ - _)%E. + under eq_bigr do rewrite mulrC. + rewrite -(@exprDn_comm _ `1-p p n); last first. + by rewrite /GRing.comm/onem; lra. + rewrite /onem addrC. + have -> : p + (1 - p) = 1 by lra. + by rewrite expr1n. + rewrite subn0 expr0 bin0 mulr1n. + rewrite /onem. + congr _%:E. + set pn := (1-p) ^+ n. + lra. +move=> i _. +apply/mulrn_wge0/mulr_ge0; apply/exprn_ge0. +exact: p0. +apply/onem_ge0/p1. +Qed. + (* exec [let x := sample (binomial n e) in *) (* return (x >= 1)] = *) (* exec [let y := sample (bernoulli (1 - e^n)) in *) (* return y] *) -Lemma binomial_le1 x y n g e t U : +(* Lemma binomial_le1' x y n g e t U : (0 <= (projT1 (execD e) t : mtyp Real) <= 1) -> execP ([let x := Sample {exp_binomial_trunc n e} in - return {1}:R <= #x] : @exp R _ g _) t U = - execP [let y := Sample {exp_bernoulli_trunc [{1}:R - e ^+ n]} in + return {1}:N <= #x] : @exp R _ g _) t U = + execP [let y := Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - e]} ^+ n]} in return #y] t U. Proof. rewrite !execP_letin !execP_sample execD_binomial_trunc execD_bernoulli_trunc/=. @@ -70,20 +189,19 @@ move=> H0 H1. rewrite !execP_return !execD_rel. have /= := @execD_var R ((y, Bool) :: g) y. rewrite eqxx => /(_ H0) ->. - have /= := @execD_var R ((x, Real) :: g) x. + have /= := @execD_var R ((x, Nat) :: g) x. rewrite eqxx => /(_ H1) -> /=. -rewrite (@execD_bin _ _ binop_minus) execD_pow !execD_real/=. +rewrite (@execD_bin _ _ binop_minus) execD_pow (@execD_bin _ _ binop_minus). +rewrite !execD_real execD_nat/=. rewrite 2!letin'E/=. set p := projT1 (execD e) t. move => /andP[p0 p1]. rewrite (@integral_bernoulli_trunc _ _ (fun x => \d_x U))//; last first. apply/andP; split. - exact: (onemX_ge0 _ p0 p1). - apply/onem_le1/exprn_ge0/p0. -rewrite (@integral_binomial_probabilty_trunc _ n p _ _ (fun y => \d_(1 <= y)%R U))//; last first. - (* move=>/= _ y0 my0. - rewrite setTI. *) - apply: measurable_fun_dirac. + apply/onemX_ge0; lra. + apply/onem_le1/exprn_ge0; lra. +rewrite (@integral_binomial_probabilty_trunc _ n p _ _ (fun y => \d_(1 <= y)%N U))//; last first. + (* apply: measurable_fun_dirac. have := @subsetT _ U; rewrite setT_bool => UT. have [->|->|->|->] /= := subset_set2 UT. exact: measurable0. @@ -94,9 +212,15 @@ rewrite (@integral_binomial_probabilty_trunc _ n p _ _ (fun y => \d_(1 <= y)%R U apply/seteqP. admit. admit. - (* rewrite diracE. *) + rewrite diracE. *) rewrite !big_ord_recl/=. -have -> : (1 <= 0 :> R) = false by lra. +rewrite addeC /bump/=. +congr (_ + _)%E. + under eq_bigr => i _. + have -> : (0 < 1 + i)%N => //. + over. +congr (_ * _)%E. +have -> : (0 < 0)%N = false. rewrite /bump. under eq_bigr => i _. rewrite /= natrD. @@ -132,7 +256,7 @@ move=> i _. apply/mulrn_wge0/mulr_ge0; apply/exprn_ge0. exact: p0. apply/onem_ge0/p1. -Admitted. +Admitted. *) Lemma __ : uniform_probability a01 `[0, (1 / 2)] = (1 / 2)%:E. Proof. @@ -154,11 +278,17 @@ transitivity ( by []. Admitted. -Lemma execP_letin_uniform g u a b p (Hap : infer (a != p)) -(Hbp : infer (b != p)) (s0 s1 : @exp R D ((a, Unit) :: (b, Real) :: (p, Real) :: g) Real -> exp P ((p, Real) :: g) u) : - (forall (t : R) x U, t \in `[0, 1] -> execP (s0 [#p]) (t, x) U = execP (s1 [#p]) (t, x) U) -> - execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s0 [#p]}] = - execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s1 [#p]}]. +Let weak_head fl g {t1 t2} x (e : @exp R fl g t2) (xg : x \notin dom g) := + exp_weak fl [::] _ (x, t1) e xg. + +Lemma execP_letin_uniform g u p + (* (Hap : infer (a != p)) (Hbp : infer (b != p)) *) + (* (bg : b \notin dom ((p, Real) :: g)) (ag : b \notin dom ((b, Nat) :: (p, Real) :: g)) *) + (s0 s1 : exp P ((p, Real) :: g) u) : + (forall (t : R) x U, 0 <= t <= 1 -> + execP s0 (t, x) U = execP s1 (t, x) U) -> + execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s0}] = + execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s1}]. Proof. move=> s01. rewrite !execP_letin execP_sample execD_uniform/=. @@ -170,47 +300,93 @@ apply: s01. by rewrite inE in t01. Qed. -Lemma casino01 : execP casino0 = execP casino1. +(* Lemma casino01 : execP casino0 = execP casino1. Proof. rewrite /casino0 /casino1. -pose s0 := fun x => [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} - in let "_" := if #{"a1"} == {5}:R then return TT else Score {0}:R - in let "a2" := Sample {exp_binomial_trunc 3 [x]} in {exp_return [{1}:R <= #{"a2"}]}] : @exp R _ _ _. -pose s1 := fun x => [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} - in let "_" := if #{"a1"} == {5}:R then return TT else Score {0}:R - in let "a2" - := Sample {exp_bernoulli_trunc [{1}:R - x ^+ {3%N}]} - in {exp_return [#{"a2"}]}] : @exp R _ _ _. -apply: (@execP_letin_uniform [::] _ _ _ _ _ _ (s0 _) (s1 _)). -move=> e k k01 ek. -rewrite /s0/s1. -rewrite 2!execP_letin. -rewrite 2![in RHS]execP_letin. -congr (letin' _ _ _ _). -congr (letin' _ _ _). -(* apply: eq_sfkernel => x U. *) -rewrite !execP_letin !execP_sample execD_bernoulli_trunc execD_binomial_trunc/=. -rewrite (@execD_bin _ _ binop_minus) !execP_return execD_rel execD_pow/=. -rewrite !execD_real !exp_var'E !(execD_var_erefl "a2") (execD_var_erefl "p")/=. - -have H := (@binomial_le1 "a2" "a2" 3 _ [#{"p"}]). -(* f_equal. *) -(* congr ((letin' _ _ _) _). -apply: eq_sfkernel => ?. -have : 0 <= e <= 1. - admit. -move=> H. -rewrite (binomial_le1 _ _ _ _ H). -under eq_integral. -rewrite execP_letin. -congr (letin'). - -(* apply: eq_sfkernel => x U. *) -have : 0 <= (projT1 (execD e) x : mtyp Real) <= 1. - by rewrite ek. -move=> H. -apply: (binomial_le1 _ _ _ _ H). *) -Admitted. +rewrite !execP_letin !execP_sample execD_uniform !execD_binomial_trunc /=. +rewrite execP_if execP_score !execP_return !execD_bernoulli_trunc/=. +rewrite !execD_rel (@execD_bin _ _ binop_minus) !execD_real execD_pow/=. +rewrite !execD_nat execD_unit/=. +rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a1")/=. +rewrite !(execD_var_erefl "a2")/=. +do 3 congr letin'. +apply: eq_sfkernel => x U. +rewrite !letin'E/=. +apply: binomial_le1. + +rewrite /=. +Abort. *) + +(* Lemma casino01 : execP casino0 = execP casino1. +Proof. +rewrite /casino0 /casino1. +pose s0 := + [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in + {exp_return [{1}:N <= #{"a2"}]}]. +pose s1 := + [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + let "a2" := Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%N}]} in + {exp_return [#{"a2"}]}]. +have := (@execP_letin_uniform [::] Bool "p" (s0 R (found "p" Real [::]) _ _) (s1 R (found "p" Real [::]) _ _)). +apply. +move=> p x U r01. +rewrite /s0/s1/=. +rewrite !execP_letin !execP_sample !execD_binomial_trunc /=. +rewrite execP_if execP_score !execP_return !execD_bernoulli_trunc/=. +rewrite !execD_rel (@execD_bin _ _ binop_minus) execD_pow. +rewrite (@execD_bin _ _ binop_minus) !execD_real/=. +rewrite !execD_nat execD_unit/=. +rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a1")/=. +rewrite !(execD_var_erefl "a2")/=. +rewrite !letin'E/=. +move: r01 => /andP[r0 r1]. +rewrite !integral_binomial_probabilty_trunc//=. +apply: eq_bigr => i _. +congr (_ * _)%E. +rewrite !letin'E iteE/=. +congr (\int[_]_y _)%E. +apply: funext => x0. +rewrite !letin'E/=. +by apply/binomial_le1/andP. +Qed. *) + +Lemma casino01' : execP casino0 = execP casino1'. +Proof. +rewrite /casino0 /casino1. +pose s0 := + [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in + return {1}:N <= #{"a2"}]. +(* pose s1 := + [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + return {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%N}]}]. *) +have := (@execP_letin_uniform [::] Bool "p" (s0 R (found "p" Real [::]) _ _) _). +apply. +move=> p x U r01. +rewrite /s0/=. +rewrite !execP_letin !execP_sample !execD_binomial_trunc /=. +rewrite execP_if execP_score !execP_return !execD_bernoulli_trunc/=. +rewrite !execD_rel (@execD_bin _ _ binop_minus) execD_pow. +rewrite (@execD_bin _ _ binop_minus) !execD_real/=. +rewrite !execD_nat execD_unit/=. +rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a1")/=. +rewrite !(execD_var_erefl "a2")/=. +rewrite !letin'E/=. +move: r01 => /andP[r0 r1]. +rewrite !integral_binomial_probabilty_trunc//=. +apply: eq_bigr => i _. +congr (_ * _)%E. +rewrite !letin'E iteE/=. +congr (\int[_]_y _)%E. +apply: funext => x0. +rewrite !letin'E/=. +by apply/binomial_le1'/andP. +Qed. (* guard test *) Definition test_guard : @exp R _ [::] _ := [ diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 59e2d0b1ea..f8255e2fa6 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -45,7 +45,7 @@ Local Open Scope ereal_scope. (* Definition mR (R : realType) : Type := R. HB.instance Definition _ (R : realType) := Measurable.on (mR R). -(* [the measurableType (R.-ocitv.-measurable).-sigma of +(* [the measurableType (R.-ocitv.-measurable).-sigma of salgebraType (R.-ocitv.-measurable)]. *) *) Module Notations. @@ -57,10 +57,11 @@ Notation var2of3 := (measurableT_comp (@measurable_snd _ _ _ _) (@measurable_fst _ _ _ _)). Notation var3of3 := (@measurable_snd _ _ _ _).*) -(* Definition mR R := [the measurableType (R.-ocitv.-measurable).-sigma of +(* Definition mR R := [the measurableType (R.-ocitv.-measurable).-sigma of salgebraType (R.-ocitv.-measurable)]. *) Notation munit := (unit : measurableType _). Notation mbool := (bool : measurableType _). +Notation mnat := (nat : measurableType _). End Notations. (* TODO: PR *) @@ -137,7 +138,7 @@ Arguments p1S {R}. Section bernoulli. Variables (R : realType) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). -Definition bernoulli : set _ -> \bar R := +Definition bernoulli : set bool -> \bar R := measure_add [the measure _ _ of mscale p [the measure _ _ of dirac true]] [the measure _ _ of mscale (onem_nonneg p1) [the measure _ _ of dirac false]]. @@ -212,10 +213,7 @@ have [->|->|->|->] /= := subset_set2 UT. apply/funext => x/=. by case: sumbool_ler. - rewrite [X in measurable_fun _ X](_ : _ = (fun x => if 0 <= x <= 1 then x%:E else 0%E))//. - apply: measurable_fun_ifT => //=; apply: measurable_and => //; apply: (measurable_fun_bool true) => //=. - rewrite (_ : _ @^-1` _ = `[0, +oo[%classic)//. - by apply/seteqP; split => [x|x] /=; rewrite in_itv/= andbT. - by rewrite (_ : _ @^-1` _ = `]-oo, 1]%classic). + by apply: measurable_fun_ifT => //=; apply: measurable_and => //; exact: measurable_fun_ler. apply/funext => x/=; case: sumbool_ler => /= x0. case: sumbool_ler => /= x1. rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. @@ -228,11 +226,8 @@ have [->|->|->|->] /= := subset_set2 UT. by rewrite mem_set//= memNset//= ?simpe leNgt x0. - rewrite [X in measurable_fun _ X](_ : _ = (fun x => if 0 <= x <= 1 then (`1-x)%:E else 1%E))//. apply: measurable_fun_ifT => //=. - apply: measurable_and => //; apply: (measurable_fun_bool true) => //=. - rewrite (_ : _ @^-1` _ = `[0, +oo[%classic)//. - by apply/seteqP; split => [x|x] /=; rewrite in_itv/= andbT. - by rewrite (_ : _ @^-1` _ = `]-oo, 1]%classic). - by apply/EFin_measurable_fun; apply/measurable_funB. + by apply: measurable_and => //; exact: measurable_fun_ler. + by apply/EFin_measurable_fun; apply/measurable_funB. apply/funext => x/=; case: sumbool_ler => /= x0. case: sumbool_ler => /= x1. rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. @@ -287,11 +282,11 @@ Context {R : realType} (n : nat) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). Local Open Scope ring_scope. (* C(n, k) * p^(n-k) * (1-p)^k *) -Definition bino_term (k : nat) :{nonneg R} := - (p%:num^+(n-k)%N * (NngNum (onem_ge0 p1))%:num^+k *+ 'C(n, k))%:nng. +Definition bino_term (k : nat) : {nonneg R} := + (p%:num^+k * (NngNum (onem_ge0 p1))%:num^+(n-k)%N *+ 'C(n, k))%:nng. Lemma bino_term0 : - bino_term 0 = (p%:num^+n)%:nng. + bino_term 0 = ((NngNum (onem_ge0 p1))%:num^+n)%:nng. Proof. rewrite /bino_term bin0 subn0/=. apply/val_inj => /=. @@ -299,18 +294,22 @@ by field. Qed. Lemma bino_term1 : - bino_term 1 = (p%:num^+(n-1)%N * (NngNum (onem_ge0 p1))%:num *+ n)%:nng. + bino_term 1 = (p%:num * (NngNum (onem_ge0 p1))%:num^+(n-1)%N *+ n)%:nng. Proof. rewrite /bino_term bin1/=. apply/val_inj => /=. by rewrite expr1. Qed. +Import Notations. + +(* Check \sum_(k < n.+1) (fun k => [the measure _ _ of mscale (bino_term k) + [the measure _ _ of \d_k]]). *) (* \sum_(k < n.+1) (bino_coef p n k) * \d_k. *) -Definition binomial_probability : set (measurableTypeR R) -> \bar R := +Definition binomial_probability : set nat -> \bar R := @msum _ _ R (fun k => [the measure _ _ of mscale (bino_term k) - [the measure _ _ of @dirac _ (measurableTypeR R) k%:R R]]) n.+1. + [the measure _ _ of \d_k]]) n.+1. HB.instance Definition _ := Measure.on binomial_probability. @@ -319,8 +318,11 @@ Proof. rewrite /binomial_probability/msum/mscale/bino_term/=/mscale/=. under eq_bigr do rewrite diracT mule1. rewrite sumEFin. +under eq_bigr=> i _. + rewrite mulrC. + over. rewrite -exprDn_comm; last by rewrite /GRing.comm mulrC. -by rewrite add_onemK; congr _%:E; rewrite expr1n. +by rewrite addrC add_onemK; congr _%:E; rewrite expr1n. Qed. HB.instance Definition _ := @@ -332,10 +334,10 @@ Section integral_binomial. Variables (R : realType) (d : measure_display) (T : measurableType d). Lemma integral_binomial (n : nat) (p : {nonneg R}) - (p1 : (p%:num <= 1)%R) (f : R -> \bar R) + (p1 : (p%:num <= 1)%R) (f : nat -> \bar R) (mf : measurable_fun setT f) : (forall x, 0 <= f x) -> \int[binomial_probability n p1]_y (f y) = - \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k%:R. + \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k. Proof. move=> f0. rewrite ge0_integral_measure_sum//=; last first. @@ -365,6 +367,14 @@ Definition binomial_probability_trunc n (p : R) := Lemma measurable_binomial_probability_trunc (n : nat) : measurable_fun setT (binomial_probability_trunc n : R -> pprobability _ _). +Proof. +apply: (@measurability _ _ _ _ _ _ + (@pset _ _ _ : set (set (pprobability _ R)))) => //. +move=> _ -[_ [r r01] [Ys mYs <-]] <-. +apply: emeasurable_fun_infty_o => //=. +rewrite /binomial_probability_trunc/=. +(* have := @subsetT _ Ys; rewrite setT_bool => UT. +have [->|->|->|->] /= := subset_set2 UT. *) Admitted. End binomial_trunc. @@ -375,9 +385,10 @@ Arguments measurable_binomial_probability_trunc {R}. Section integral_binomial_trunc. Variables (R : realType) (d : measure_display) (T : measurableType d). +Import Notations. Lemma integral_binomial_probabilty_trunc (n : nat) (p : R) - (p0 : (0 <= p)%R) (p1 : ((NngNum p0)%:num <= 1)%R) (f : R -> \bar R) (mf : measurable_fun setT f) : - (forall x, 0 <= f x) -> \int[binomial_probability_trunc n p]_y (f y) = \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k%:R. + (p0 : (0 <= p)%R) (p1 : ((NngNum p0)%:num <= 1)%R) (f : mnat -> \bar R) (mf : measurable_fun setT f) : + (forall x, 0 <= f x) -> \int[binomial_probability_trunc n p]_y (f y) = \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k. Proof. move=> f0. rewrite /binomial_probability_trunc/=. @@ -394,15 +405,15 @@ Section binomial_example. Context {R : realType}. Open Scope ring_scope. -Lemma binomial3_2 : @binomial_probability R 3 _ (p1S 1) [set 2%:R] = (3 / 8)%:E. +Lemma binomial3_2 : @binomial_probability R 3 _ (p1S 1) [set 2%N] = (3 / 8)%:E. Proof. rewrite /binomial_probability/msum !big_ord_recl/= big_ord0 adde0 bino_term0. rewrite /mscale/= !diracE /bump/=. repeat rewrite ?binS ?bin0 ?bin1 ?bin_small//. -rewrite memNset//=; last by move/eqP; rewrite eqr_nat. -rewrite memNset//=; last by move/eqP; rewrite eqr_nat. +rewrite memNset//=. +rewrite memNset//=. rewrite mem_set//=. -rewrite memNset//=; last by move/eqP; rewrite eqr_nat. +rewrite memNset//=. congr _%:E. rewrite expr0 !mulr1 !mulr0 !add0r !addn0 !add0n /onem. by field. @@ -930,12 +941,15 @@ Definition k10 : measurable_fun _ _ := kr 10%:R. Definition ktt := @measurable_cst _ _ T _ setT tt. Definition kb (b : bool) := @measurable_cst _ _ T _ setT b. +Definition kn (n : nat) := @measurable_cst _ _ T _ setT n. + End cst_fun. Arguments kr {d T R}. Arguments k3 {d T R}. Arguments k10 {d T R}. Arguments ktt {d T}. Arguments kb {d T}. +Arguments kn {d T}. Section iter_mprod. Import Notations. From b9680629a1a029510010c9a3aa10659abe42ed9a Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 16 Jan 2024 09:44:11 +0900 Subject: [PATCH 09/48] progress - several admits proved - patch - admits in examples wip --- theories/lang_syntax_examples_wip.v | 121 +++++----- theories/prob_lang.v | 356 ++++++++++++++++++---------- 2 files changed, 281 insertions(+), 196 deletions(-) diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index 9e0c7b1198..cae3d82597 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -266,17 +266,20 @@ Abort. Lemma letin'_sample_uniform d d' (T : measurableType d) (T' : measurableType d') (a b : R) (ab0 : (0 < b - a)%R) (u : R.-sfker [the measurableType _ of (_ * T)%type] ~> T') x y : + measurable y -> letin' (sample_cst (uniform_probability ab0)) u x y = - (((b - a)^-1)%:E * \int[lebesgue_measure]_(x0 in `[a, b]) u (x0, x) y)%E. + ((b - a)^-1%:E * \int[lebesgue_measure]_(x0 in `[a, b]) u (x0, x) y)%E. Proof. -rewrite letin'E/=. -rewrite ge0_integral_mscale//=; last admit. -transitivity ( -(((b - a)^-1)%:E * \int[lebesgue_measure]_(x0 in `[a, b]) u (x0, x) y)%E -). - admit. -by []. -Admitted. +move=> my; rewrite letin'E/=. +rewrite integral_uniform//= => _ /= Y mY /=. +have /= := measurable_kernel u _ my measurableT _ mY. +move/measurable_ysection => /(_ R x) /=. +set A := (X in measurable X). +set B := (X in _ -> measurable X). +suff : A = B by move=> ->. +rewrite {}/A {}/B !setTI /ysection/= (*TODO: lemma?*) /preimage/=. +by apply/seteqP; split => [z|z] /=; rewrite inE/=. +Qed. Let weak_head fl g {t1 t2} x (e : @exp R fl g t2) (xg : x \notin dom g) := exp_weak fl [::] _ (x, t1) e xg. @@ -287,13 +290,13 @@ Lemma execP_letin_uniform g u p (s0 s1 : exp P ((p, Real) :: g) u) : (forall (t : R) x U, 0 <= t <= 1 -> execP s0 (t, x) U = execP s1 (t, x) U) -> - execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s0}] = - execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s1}]. + forall x U, measurable U -> + execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s0}] x U = + execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s1}] x U. Proof. -move=> s01. +move=> s01 x U mU. rewrite !execP_letin execP_sample execD_uniform/=. -apply: eq_sfkernel => x U. -rewrite 2!letin'_sample_uniform. +rewrite !letin'_sample_uniform//. congr (_ * _)%E. apply: eq_integral => t t01. apply: s01. @@ -353,8 +356,9 @@ rewrite !letin'E/=. by apply/binomial_le1/andP. Qed. *) -Lemma casino01' : execP casino0 = execP casino1'. +Lemma casino01' y V : measurable V -> execP casino0 y V = execP casino1' y V. Proof. +move=> mV. rewrite /casino0 /casino1. pose s0 := [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in @@ -366,7 +370,7 @@ pose s0 := let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in return {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%N}]}]. *) have := (@execP_letin_uniform [::] Bool "p" (s0 R (found "p" Real [::]) _ _) _). -apply. +apply => //. move=> p x U r01. rewrite /s0/=. rewrite !execP_letin !execP_sample !execD_binomial_trunc /=. @@ -395,7 +399,7 @@ Definition test_guard : @exp R _ [::] _ := [ return #{"p"} ]. -Lemma exec_guard t U : execP test_guard t U = ((1 / 3)%:E * @dirac _ _ true R U)%E. +Lemma exec_guard t U : execP test_guard t U = ((1 / 3)%:E * \d_true U)%E. Proof. rewrite /test_guard 2!execP_letin execP_sample execD_bernoulli execP_if/=. rewrite !execP_return !exp_var'E !(execD_var_erefl "p") execD_unit execP_score execD_real/=. @@ -406,8 +410,7 @@ by rewrite normr0 mul0e !mule0 !adde0 !diracT !mul1e. Qed. Lemma exec_casino t U : - execP casino0 t U = ((10 / 99)%:E * @dirac _ _ true R U + - (1 / 99)%:E * @dirac _ _ false R U)%E . + execP casino0 t U = ((10 / 99)%:E * \d_true U + (1 / 99)%:E * \d_false U)%E . Proof. rewrite /casino0 !execP_letin !execP_sample execD_uniform/=. rewrite !execD_binomial_trunc execP_if !execP_return !execP_score/=. @@ -417,15 +420,12 @@ rewrite (execD_var_erefl "p") (execD_var_erefl "a2")/=. rewrite letin'E/= /uniform_probability ge0_integral_mscale//=. rewrite subr0 invr1 mul1e. under eq_integral. - Admitted. Definition uniform_syntax : @exp R _ [::] _ := [let "p" := Sample {exp_uniform 0 1 a01} in return #{"p"}]. -About integralT_nnsfun. - Lemma exec_uniform_syntax t U : measurable U -> execP uniform_syntax t U = uniform_probability a01 U. Proof. @@ -433,33 +433,22 @@ move=> mU. rewrite /uniform_syntax execP_letin execP_sample execP_return !execD_uniform. rewrite exp_var'E (execD_var_erefl "p")/=. rewrite letin'E /=. +rewrite integral_uniform//=; last exact: measurable_fun_dirac. +rewrite subr0 invr1 mul1e. rewrite {1}/uniform_probability. -set x := (X in mscale _ X). -set k := (X in mscale X _). -transitivity ((k%:num)%:E * \int[x]_y \d_y U)%E. - rewrite -ge0_integral_mscale //. - exact: measurable_fun_dirac. -rewrite /uniform_probability /mscale /=. -congr (_ * _)%E. -under eq_integral do rewrite diracE. -rewrite /=. -rewrite /lebesgue_measure/=. -rewrite -integral_cst. -set f := (mrestr lebesgue_measure (measurable_itv `[0%R, 1%R])). -admit. -rewrite -(@ge0_integral_mscale _ _ _ x setT measurableT k (fun y => \d_y U)) //. -rewrite /uniform_probability/mscale/=/integral//=. -Admitted. +rewrite /mscale/= subr0 invr1 mul1e. +by rewrite integral_indic. +Qed. Definition binomial_le : @exp R _ [::] Bool := [let "a2" := Sample {exp_binomial 3 (1 / 2)%:nng (p1S 1)} in - return {1}:R <= #{"a2"}]. + return {1}:N <= #{"a2"}]. Lemma exec_binomial_le t U : - execP binomial_le t U = ((7 / 8)%:E * @dirac _ _ true R U + - (1 / 8)%:E * @dirac _ _ false R U)%E. + execP binomial_le t U = ((7 / 8)%:E * \d_true U + + (1 / 8)%:E * \d_false U)%E. Proof. -rewrite /binomial_le execP_letin execP_sample execP_return execD_rel execD_real. +rewrite /binomial_le execP_letin execP_sample execP_return execD_rel execD_nat. rewrite exp_var'E (execD_var_erefl "a2") execD_binomial. rewrite letin'E//= /binomial_probability ge0_integral_measure_sum//=. rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//=. @@ -467,39 +456,39 @@ rewrite !integral_dirac// /bump. rewrite !binS/= !bin0 bin1 bin2 bin_small// addn0. rewrite addeC adde0. congr (_ + _)%:E. -have -> : (1 <= 1)%R. admit. -have -> : (1 <= 2)%R. admit. -have -> : (1 <= 3)%R. admit. -rewrite -!mulrDl indicT !mul1r. -congr (_ * _). -rewrite /onem addn0 add0n. -by field. -congr (_ * _). -by field. -(* by rewrite ler10. *) -Admitted. + rewrite !indicT !(mul0n,add0n,lt0n,mul1r)/=. + rewrite -!mulrDl; congr (_ * _). + rewrite /onem. + lra. +rewrite !expr0 ltnn indicT/= !(mul1r,mul1e) /onem. +lra. +Qed. -Definition binomial_guard : @exp R _ [::] Real := +Definition binomial_guard : @exp R _ [::] Nat := [let "a1" := Sample {exp_binomial 3 (1 / 2)%:nng (p1S 1)} in - let "_" := if #{"a1"} == {1}:R then return TT else Score {0}:R in + let "_" := if #{"a1"} == {1}:N then return TT else Score {0}:R in return #{"a1"}]. Lemma exec_binomial_guard t U : - execP binomial_guard t U = ((7 / 8)%:E * @dirac _ R 1%R R U + - (1 / 8)%:E * @dirac _ R 0%R R U)%E. + execP binomial_guard t U = ((3 / 8)%:E * \d_1%N U(* + + (1 / 8)%:E * \d_0%N U*))%E. Proof. rewrite /binomial_guard !execP_letin execP_sample execP_return execP_if. -rewrite !exp_var'E execD_rel !(execD_var_erefl "a1") execP_return execD_unit execD_binomial execD_real execP_score execD_real. +rewrite !exp_var'E execD_rel !(execD_var_erefl "a1") execP_return. +rewrite execD_unit execD_binomial execD_nat execP_score execD_real. rewrite !letin'E//= /binomial_probability ge0_integral_measure_sum//=. rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//=. -rewrite !integral_dirac// /bump. -rewrite (* indicT *) !binS/= !bin0 bin1 bin2 bin_small// addn0 (* !mul1e *). +rewrite !integral_dirac//. +rewrite /bump/=. +rewrite !binS/= !bin0 bin1 bin2 bin_small//. +rewrite !diracT !addn0 !expr0 !subn0 !mulr1n !mul1r !expr1 !mul1e. rewrite !letin'E//= !iteE/= !diracE/=. -have -> : (0 == 1)%R = false; first by admit. -have -> : (1 == 1)%R; first by admit. -have -> : (2 == 1)%R = false; first by admit. -have -> : (3 == 1)%R = false; first by admit. -rewrite addeC adde0. -Admitted. +rewrite !ge0_integral_mscale//=. +rewrite !integral_dirac// !diracT//. +rewrite !(normr0,mul0e,mule0,add0e,add0n,mul1e,adde0). +rewrite /onem. +congr (_%:E * _)%E. +lra. +Qed. End casino_example. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index f8255e2fa6..e9d9abc03c 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -2,9 +2,9 @@ From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. From mathcomp Require Import interval_inference archimedean rat. -From mathcomp Require Import mathcomp_extra boolp classical_sets. -From mathcomp Require Import functions cardinality fsbigop. -From mathcomp Require Import reals ereal topology normedtype sequences. +From mathcomp Require Import mathcomp_extra boolp classical_sets functions. +From mathcomp Require Import cardinality fsbigop reals ereal. +From mathcomp Require Import topology normedtype sequences. From mathcomp Require Import esum measure lebesgue_measure numfun. From mathcomp Require Import lebesgue_integral exp kernel. From mathcomp Require Import ring lra. @@ -170,11 +170,11 @@ Section bernoulli_trunc. Variables (R : realType). Local Open Scope ring_scope. -Lemma sumbool_ler (x y : R) : {(x <= y)%R} + {(x > y)%R}. +Lemma sumbool_ler (x y : R) : (x <= y)%R + (x > y)%R. Proof. have [_|_] := leP x y. -by apply left. -by apply right. +by apply (*left*) inl. +by apply (*right*) inr. Qed. (* TODO: move? *) @@ -188,11 +188,11 @@ Definition bernoulli0 := @bernoulli R 0%R%:nng ler01. HB.instance Definition _ := Probability.on bernoulli0. Definition bernoulli_trunc (p : R) := match sumbool_ler 0%R p with -| left l0p => match sumbool_ler (NngNum l0p)%:num 1%R with - | left lp1 => [the probability _ _ of @bernoulli R (NngNum l0p) lp1] - | right _ => [the probability _ _ of bernoulli0] +| inl l0p => match sumbool_ler (NngNum l0p)%:num 1%R with + | inl lp1 => [the probability _ _ of @bernoulli R (NngNum l0p) lp1] + | inr _ => [the probability _ _ of bernoulli0] end -| right _ => [the probability _ _ of bernoulli0] +| inr _ => [the probability _ _ of bernoulli0] end. (*HB.instance Definition _ (p : R) := Probability.on (bernoulli_trunc p).*) @@ -204,55 +204,52 @@ Lemma measurable_bernoulli_trunc : Proof. apply: (@measurability _ _ _ _ _ _ (@pset _ _ _ : set (set (pprobability _ R)))) => //. -move=> _ -[_ [r r01] [Ys mYs <-]] <-. -apply: emeasurable_fun_infty_o => //=. +move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. rewrite /bernoulli_trunc/=. have := @subsetT _ Ys; rewrite setT_bool => UT. have [->|->|->|->] /= := subset_set2 UT. -- rewrite [X in measurable_fun _ X](_ : _ = (fun=> 0%E))//. +- rewrite [X in measurable_fun _ X](_ : _ = cst 0%E)//. apply/funext => x/=. by case: sumbool_ler. - rewrite [X in measurable_fun _ X](_ : _ = (fun x => if 0 <= x <= 1 then x%:E else 0%E))//. by apply: measurable_fun_ifT => //=; apply: measurable_and => //; exact: measurable_fun_ler. apply/funext => x/=; case: sumbool_ler => /= x0. case: sumbool_ler => /= x1. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. - by rewrite mem_set//= memNset//= ?simpe x0 x1. - rewrite /bernoulli0. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. - by rewrite mem_set//= memNset//= ?simpe x0/= leNgt x1. - rewrite /bernoulli0. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. - by rewrite mem_set//= memNset//= ?simpe leNgt x0. -- rewrite [X in measurable_fun _ X](_ : _ = (fun x => if 0 <= x <= 1 then (`1-x)%:E else 1%E))//. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. + by rewrite /mscale/= !diracE mem_set//= memNset//= ?simpe x0 x1. + rewrite /bernoulli0 /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. + by rewrite /mscale/= !diracE mem_set//= memNset//= ?simpe x0/= leNgt x1. + rewrite /bernoulli0 /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. + by rewrite /mscale/= !diracE mem_set//= memNset//= ?simpe leNgt x0. +- rewrite [X in measurable_fun _ X](_ : _ = + (fun x => if 0 <= x <= 1 then (`1-x)%:E else 1%E))//. apply: measurable_fun_ifT => //=. by apply: measurable_and => //; exact: measurable_fun_ler. by apply/EFin_measurable_fun; apply/measurable_funB. apply/funext => x/=; case: sumbool_ler => /= x0. case: sumbool_ler => /= x1. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. - by rewrite memNset//= mem_set//= ?simpe x0 x1/=. - rewrite /bernoulli0. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. - by rewrite memNset//= mem_set//= ?simpe x0/= leNgt x1/= onem0. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. + by rewrite /mscale/= !diracE memNset//= mem_set//= ?simpe x0 x1/=. + rewrite /bernoulli0 /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. + by rewrite /mscale/= !diracE memNset//= mem_set//= ?simpe x0/= leNgt x1/= onem0. rewrite /bernoulli0. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. - by rewrite memNset//= mem_set//= leNgt x0/= ?simpe onem0. -- rewrite [X in measurable_fun _ X](_ : _ = (fun=> 1%E))//; apply/funext => x/=. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. + by rewrite /mscale/= !diracE memNset//= mem_set//= leNgt x0/= ?simpe onem0. +- rewrite [X in measurable_fun _ X](_ : _ = cst 1%E)//; apply/funext => x/=. case: sumbool_ler => /= x0. case: sumbool_ler => /= x1. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. - rewrite mem_set//=; last by left. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. + rewrite /mscale/= !diracE mem_set//=; last by left. rewrite mem_set//=; last by right. by rewrite ?simpe -EFinD add_onemK. rewrite /bernoulli0. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. - rewrite mem_set//=; last by left. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. + rewrite /mscale/= !diracE mem_set//=; last by left. rewrite mem_set//=; last by right. by rewrite ?simpe onem0. rewrite /bernoulli0. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//= /mscale/= !diracE. - rewrite mem_set//=; last by left. + rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. + rewrite /mscale/= !diracE mem_set//=; last by left. rewrite mem_set//=; last by right. by rewrite ?simpe onem0. Qed. @@ -263,15 +260,12 @@ Arguments bernoulli_trunc {R}. Arguments measurable_bernoulli_trunc {R}. Lemma integral_bernoulli_trunc {R : realType} (p : R) (f : bool -> \bar R) : - (0 <= p <= 1)%R -> - (forall x, 0 <= f x) -> - \int[bernoulli_trunc p]_y (f y) = - p%:E * f true + (`1-p)%:E * f false. -Proof. -move=> /andP[p0 p1] f0. -rewrite /bernoulli_trunc. -case: (sumbool_ler 0 p) => [p0'/=|]. - case: (sumbool_ler p 1) => [p1'/=|]. + (0 <= p <= 1)%R -> (forall x, 0 <= f x) -> + \int[bernoulli_trunc p]_y (f y) = p%:E * f true + (`1-p)%:E * f false. +Proof. +move=> /andP[p0 p1] f0; rewrite /bernoulli_trunc. +case: sumbool_ler => [? /=|]. + case: (sumbool_ler p 1) => [? /=|]. by rewrite integral_bernoulli. by rewrite ltNge p1. by rewrite ltNge p0. @@ -285,7 +279,7 @@ Local Open Scope ring_scope. Definition bino_term (k : nat) : {nonneg R} := (p%:num^+k * (NngNum (onem_ge0 p1))%:num^+(n-k)%N *+ 'C(n, k))%:nng. -Lemma bino_term0 : +Lemma bino_term0 : bino_term 0 = ((NngNum (onem_ge0 p1))%:num^+n)%:nng. Proof. rewrite /bino_term bin0 subn0/=. @@ -293,7 +287,7 @@ apply/val_inj => /=. by field. Qed. -Lemma bino_term1 : +Lemma bino_term1 : bino_term 1 = (p%:num * (NngNum (onem_ge0 p1))%:num^+(n-1)%N *+ n)%:nng. Proof. rewrite /bino_term bin1/=. @@ -333,49 +327,123 @@ End binomial_probability. Section integral_binomial. Variables (R : realType) (d : measure_display) (T : measurableType d). -Lemma integral_binomial (n : nat) (p : {nonneg R}) +Lemma integral_binomial (n : nat) (p : {nonneg R}) (p1 : (p%:num <= 1)%R) (f : nat -> \bar R) (mf : measurable_fun setT f) : (forall x, 0 <= f x) -> \int[binomial_probability n p1]_y (f y) = \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k. Proof. -move=> f0. -rewrite ge0_integral_measure_sum//=; last first. -apply: eq_bigr => i _. -rewrite ge0_integral_mscale//=; last first. -rewrite integral_dirac//=; last first. -by rewrite diracT mul1e. +move=> f0; rewrite ge0_integral_measure_sum//=; apply: eq_bigr => i _. +by rewrite ge0_integral_mscale//= integral_dirac//= diracT mul1e. Qed. End integral_binomial. +(* X + Y is a measurableType if X and Y are *) +HB.instance Definition _ (X Y : pointedType) := + isPointed.Build (X + Y)%type (@inl X Y point). + +Section measurable_sum. +Context d d' (X : measurableType d) (Y : measurableType d'). + +Definition measurable_sum : set (set (X + Y)) := setT. + +Let sum0 : measurable_sum set0. Proof. by []. Qed. + +Let sumC A : measurable_sum A -> measurable_sum (~` A). Proof. by []. Qed. + +Let sumU (F : (set (X + Y))^nat) : (forall i, measurable_sum (F i)) -> + measurable_sum (\bigcup_i F i). +Proof. by []. Qed. + +HB.instance Definition _ := @isMeasurable.Build default_measure_display (X + Y)%type + measurable_sum sum0 sumC sumU. + +End measurable_sum. + +Lemma measurable_fun_sum dA dB d' (A : measurableType dA) (B : measurableType dB) + (Y : measurableType d') (f : A -> Y) (g : B -> Y) : + measurable_fun setT f -> measurable_fun setT g -> + measurable_fun setT (fun tb : A + B => + match tb with inl a => f a | inr b => g b end). +Proof. +move=> mx my/= _ Z mZ /=; rewrite setTI /=. +rewrite (_ : _ @^-1` Z = inl @` (f @^-1` Z) `|` inr @` (g @^-1` Z)). + exact: measurableU. +apply/seteqP; split. + by move=> [a Zxa|b Zxb]/=; [left; exists a|right; exists b]. +by move=> z [/= [a Zxa <-//=]|]/= [b Zyb <-//=]. +Qed. + +(* TODO: measurable_fun_if_pair -> measurable_fun_if_pair_bool? *) +Lemma measurable_fun_if_pair_nat d d' (X : measurableType d) + (Y : measurableType d') (f g : X -> Y) (n : nat) : + measurable_fun setT f -> measurable_fun setT g -> + measurable_fun setT (fun xn => if xn.2 == n then f xn.1 else g xn.1). +Proof. +move=> mx my; apply: measurable_fun_ifT => //=. +- have h : measurable_fun [set: nat] (fun t => t == n) by []. + exact: (@measurableT_comp _ _ _ _ _ _ _ _ _ h). +- exact: measurableT_comp. +- exact: measurableT_comp. +Qed. + Section binomial_trunc. Variables (R : realType). Local Open Scope ring_scope. Definition binomial_probability0 := @binomial_probability R 0 0%:nng%R ler01. -Definition binomial_probability_trunc n (p : R) := - (* (p1 : (p%:num <= 1)%R) := *) +Definition binomial_probability_trunc n (p : R) := match (sumbool_ler 0%R p) with - | left l0p => match (sumbool_ler (NngNum l0p)%:num 1%R) with - | left lp1 => [the probability _ _ of @binomial_probability R n (NngNum l0p) lp1] - | right _ => [the probability _ _ of binomial_probability0] + | inl l0p => match (sumbool_ler (NngNum l0p)%:num 1%R) with + | inl lp1 => [the probability _ _ of @binomial_probability R n (NngNum l0p) lp1] + | inr _ => [the probability _ _ of binomial_probability0] end - | right _ => [the probability _ _ of binomial_probability0] + | inr _ => [the probability _ _ of binomial_probability0] end. -Lemma measurable_binomial_probability_trunc (n : nat) +Lemma measurable_binomial_probability_trunc (n : nat) : measurable_fun setT (binomial_probability_trunc n : R -> pprobability _ _). Proof. apply: (@measurability _ _ _ _ _ _ (@pset _ _ _ : set (set (pprobability _ R)))) => //. -move=> _ -[_ [r r01] [Ys mYs <-]] <-. -apply: emeasurable_fun_infty_o => //=. +move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. rewrite /binomial_probability_trunc/=. -(* have := @subsetT _ Ys; rewrite setT_bool => UT. -have [->|->|->|->] /= := subset_set2 UT. *) -Admitted. +set f := (X in measurable_fun _ X). +rewrite (_ : f = fun x => if 0 <= x <= 1 then (\sum_(m < n.+1) + match sumbool_ler 0 x with + | inl l0p => + match sumbool_ler x 1 with + | inl lp1 => mscale (@bino_term _ n (NngNum l0p) lp1 m) (\d_(nat_of_ord m)) Ys + | inr _ => (x ^+ m * `1-x ^+ (n - m) *+ 'C(n, m))%:E * \d_(nat_of_ord m) Ys + end + | inr _ => (x ^+ m * `1-x ^+ (n - m) *+ 'C(n, m))%:E * \d_(nat_of_ord m) Ys + end)%E + else \d_0%N Ys)//. + move=> ?; apply: measurable_fun_ifT => //=. + by apply: measurable_and => //; exact: measurable_fun_ler. + apply: emeasurable_fun_sum => m /=. + rewrite /mscale/= [X in measurable_fun _ X](_ : _ = (fun x => + ((x ^+ m * `1-x ^+ (n - m) *+ 'C(n, m))%:E * \d_(nat_of_ord m) Ys)%E)); last first. + by apply:funext => x; case: sumbool_ler => // x0; case: sumbool_ler. + apply: emeasurable_funM => //; apply/EFin_measurable_fun => //. + under eq_fun do rewrite -mulrnAr. + apply: measurable_funM => //. + under eq_fun do rewrite -mulr_natr. + apply: measurable_funM => //=. + apply: measurable_fun_pow. + exact: measurable_funB. +rewrite {}/f. +apply/funext => x. +case: sumbool_ler => /= x0. + case: sumbool_ler => /= x1. + by rewrite /binomial_probability/= /msum /= /bino_term/= x0 x1. + rewrite /binomial_probability /= /msum big_ord_recl/= big_ord0 /mscale. + by rewrite /= expr0 mul1r subnn expr0 bin0 mul1e adde0 !leNgt x1 andbF. +rewrite /binomial_probability /= /msum big_ord_recl/= big_ord0 /mscale. +by rewrite /= expr0 mul1r subnn expr0 bin0 mul1e adde0 !leNgt x0/=. +Qed. End binomial_trunc. @@ -386,15 +454,16 @@ Section integral_binomial_trunc. Variables (R : realType) (d : measure_display) (T : measurableType d). Import Notations. -Lemma integral_binomial_probabilty_trunc (n : nat) (p : R) - (p0 : (0 <= p)%R) (p1 : ((NngNum p0)%:num <= 1)%R) (f : mnat -> \bar R) (mf : measurable_fun setT f) : - (forall x, 0 <= f x) -> \int[binomial_probability_trunc n p]_y (f y) = \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k. +Lemma integral_binomial_probabilty_trunc (n : nat) (p : R) + (p0 : (0 <= p)%R) (p1 : ((NngNum p0)%:num <= 1)%R) (f : mnat -> \bar R) + (mf : measurable_fun setT f) : (forall x, 0 <= f x) -> + \int[binomial_probability_trunc n p]_y (f y) = + \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k. Proof. -move=> f0. -rewrite /binomial_probability_trunc/=. -case: (sumbool_ler 0 p) => [p0'/=|]. - case: (sumbool_ler p 1) => [p1'/=|]. - by rewrite integral_binomial. +move=> f0; rewrite /binomial_probability_trunc/=. +case: sumbool_ler => [? /=|]. + case: sumbool_ler => [?/=|]. + by rewrite integral_binomial. by rewrite ltNge p1. by rewrite ltNge p0. Qed. @@ -406,7 +475,7 @@ Context {R : realType}. Open Scope ring_scope. Lemma binomial3_2 : @binomial_probability R 3 _ (p1S 1) [set 2%N] = (3 / 8)%:E. -Proof. +Proof. rewrite /binomial_probability/msum !big_ord_recl/= big_ord0 adde0 bino_term0. rewrite /mscale/= !diracE /bump/=. repeat rewrite ?binS ?bin0 ?bin1 ?bin_small//. @@ -424,13 +493,11 @@ End binomial_example. Section uniform_probability. Context (R : realType) (a b : R) (ab0 : (0 < b - a)%R). -Definition uniform_probability -(* : set _ -> \bar R *) - := @mscale _ _ R (invr_nonneg (NngNum (ltW ab0))) +Definition uniform_probability (* : set _ -> \bar R *) := + @mscale _ _ R (invr_nonneg (NngNum (ltW ab0))) (mrestr lebesgue_measure (measurable_itv `[a, b])). -(** TODO: set R -> \bar R を書くとMeasure.onが通らない **) -(** **) +(* NB: set R -> \bar R を書くとMeasure.onが通らない *) HB.instance Definition _ := Measure.on uniform_probability. (* Let uniform0 : uniform_probability set0 = 0. @@ -457,6 +524,84 @@ HB.instance Definition _ := @Measure_isProbability.Build _ _ R End uniform_probability. +Section integral_uniform. +Context {R : realType}. + +Let integral_uniform_indic (a b : R) (ab0 : (0 < b - a)%R) E : + measurable E -> let m := uniform_probability ab0 in + \int[m]_x (\1_E x)%:E = + (b - a)^-1%:E * \int[lebesgue_measure]_(x in `[a, b]) (\1_E x)%:E. +Proof. +move=> mE m. +by rewrite !integral_indic//= /uniform_probability/= /mscale/= /mrestr setIT. +Qed. + +Import HBNNSimple. + +Let integral_uniform_nnsfun (f : {nnsfun measurableTypeR R >-> R}) + (a b : R) (ab0 : (0 < b - a)%R) : let m := uniform_probability ab0 in + \int[m]_x (f x)%:E = + (b - a)^-1%:E * \int[lebesgue_measure]_(x in `[a, b]) (f x)%:E. +Proof. +move=> m. +under [LHS]eq_integral do rewrite fimfunE -fsumEFin//. +rewrite [LHS]ge0_integral_fsum//; last 2 first. + - by move=> r; exact/EFin_measurable_fun/measurableT_comp. + - by move=> n x _; rewrite EFinM nnfun_muleindic_ge0. +rewrite -[RHS]ge0_integralZl//; last 3 first. + - exact/EFin_measurable_fun/measurable_funTS. + - by move=> x _; rewrite lee_fin. + - by rewrite lee_fin invr_ge0// ltW. +under [RHS]eq_integral. + move=> x xD; rewrite fimfunE -fsumEFin// ge0_mule_fsumr; last first. + by move=> r; rewrite EFinM nnfun_muleindic_ge0. + over. +rewrite [RHS]ge0_integral_fsum//; last 2 first. + - by move=> r; apply/EFin_measurable_fun; do 2 apply/measurableT_comp => //. + - move=> n x _; rewrite EFinM mule_ge0//; last by rewrite nnfun_muleindic_ge0. + by rewrite lee_fin invr_ge0// ltW. +apply: eq_fsbigr => r _; rewrite ge0_integralZl//. +- by rewrite !integralZl_indic_nnsfun//= integral_uniform_indic// muleCA. +- exact/EFin_measurable_fun/measurableT_comp. +- by move=> t _; rewrite nnfun_muleindic_ge0. +- by rewrite lee_fin invr_ge0// ltW. +Qed. + +Lemma integral_uniform (f : measurableTypeR R -> \bar R) (a b : R) + (ab0 : (0 < b - a)%R) : measurable_fun setT f -> (forall x, 0 <= f x) -> + let m := uniform_probability ab0 in + \int[m]_x f x = + (b - a)^-1%:E * \int[lebesgue_measure]_(x in `[a, b]) f x. +Proof. +move=> mf f0 m. +have [f_ [ndf_ f_f]] := approximation measurableT mf (fun x0 _ => f0 x0). +transitivity (lim (\int[m]_(x in setT) (f_ n x)%:E @[n --> \oo])). + rewrite -monotone_convergence//=. + - apply: eq_integral => ? /[!inE] xD; apply/esym/cvg_lim => //=. + exact: f_f. + - by move=> n; exact/EFin_measurable_fun/measurable_funTS. + - by move=> n ? _; rewrite lee_fin. + - by move=> ? _ ? ? mn; rewrite lee_fin; exact/lefP/ndf_. +rewrite [X in _ = _ * X](_ : _ = lim + (\int[lebesgue_measure]_(x in `[a, b]) (f_ n x)%:E @[n --> \oo])); last first. + rewrite -monotone_convergence//=. + - by apply: eq_integral => ? /[!inE] xD; apply/esym/cvg_lim => //; exact: f_f. + - by move=> n; exact/EFin_measurable_fun/measurable_funTS. + - by move=> n ? _; rewrite lee_fin. + - by move=> ? _ ? ? /ndf_ /lefP; rewrite lee_fin. +rewrite -limeMl//. + apply: congr_lim. + by apply/funext => n /=; exact: integral_uniform_nnsfun. +apply/ereal_nondecreasing_is_cvgn => ? ? ab; apply: ge0_le_integral => //=. +- by move=> ? _; rewrite lee_fin. +- exact/EFin_measurable_fun/measurable_funTS. +- by move=> ? _; rewrite lee_fin. +- exact/EFin_measurable_fun/measurable_funTS. +- by move=> ? _; rewrite lee_fin; move/ndf_ : ab => /lefP. +Qed. + +End integral_uniform. + Section mscore. Context d (T : measurableType d) (R : realType). Variable f : T -> R. @@ -1486,7 +1631,7 @@ Qed. Definition bernoulli_and : R.-sfker T ~> mbool := (letin (sample_cst [the probability _ _ of bernoulli p12]) (letin (sample_cst [the probability _ _ of bernoulli p12]) - (ret (measurable_fun_mand macc1of3 macc2of3)))). + (ret (measurable_fun_mand macc1of3 macc2of3)))). *) Lemma bernoulli_andE t U : @@ -1618,55 +1763,6 @@ Qed. End staton_bus_exponential. -(* X + Y is a measurableType if X and Y are *) -HB.instance Definition _ (X Y : pointedType) := - isPointed.Build (X + Y)%type (@inl X Y point). - -Section measurable_sum. -Context d d' (X : measurableType d) (Y : measurableType d'). - -Definition measurable_sum : set (set (X + Y)) := setT. - -Let sum0 : measurable_sum set0. Proof. by []. Qed. - -Let sumC A : measurable_sum A -> measurable_sum (~` A). Proof. by []. Qed. - -Let sumU (F : (set (X + Y))^nat) : (forall i, measurable_sum (F i)) -> - measurable_sum (\bigcup_i F i). -Proof. by []. Qed. - -HB.instance Definition _ := @isMeasurable.Build default_measure_display (X + Y)%type - measurable_sum sum0 sumC sumU. - -End measurable_sum. - -Lemma measurable_fun_sum dA dB d' (A : measurableType dA) (B : measurableType dB) - (Y : measurableType d') (f : A -> Y) (g : B -> Y) : - measurable_fun setT f -> measurable_fun setT g -> - measurable_fun setT (fun tb : A + B => - match tb with inl a => f a | inr b => g b end). -Proof. -move=> mx my/= _ Z mZ /=; rewrite setTI /=. -rewrite (_ : _ @^-1` Z = inl @` (f @^-1` Z) `|` inr @` (g @^-1` Z)). - exact: measurableU. -apply/seteqP; split. - by move=> [a Zxa|b Zxb]/=; [left; exists a|right; exists b]. -by move=> z [/= [a Zxa <-//=]|]/= [b Zyb <-//=]. -Qed. - -(* TODO: measurable_fun_if_pair -> measurable_fun_if_pair_bool? *) -Lemma measurable_fun_if_pair_nat d d' (X : measurableType d) - (Y : measurableType d') (f g : X -> Y) (n : nat) : - measurable_fun setT f -> measurable_fun setT g -> - measurable_fun setT (fun xn => if xn.2 == n then f xn.1 else g xn.1). -Proof. -move=> mx my; apply: measurable_fun_ifT => //=. -- have h : measurable_fun [set: nat] (fun t => t == n) by []. - exact: (@measurableT_comp _ _ _ _ _ _ _ _ _ h). -- exact: measurableT_comp. -- exact: measurableT_comp. -Qed. - Module CASE_NAT. Section case_nat. Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). From 3731371db0e7f77badb5f0b7d337f3537fd71ba1 Mon Sep 17 00:00:00 2001 From: AyumuSaito Date: Tue, 14 Feb 2023 22:02:38 +0900 Subject: [PATCH 10/48] minor fixes Co-authored-by: Reynald Affeldt --- theories/prob_lang.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/prob_lang.v b/theories/prob_lang.v index e9d9abc03c..4f409d920d 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -1631,7 +1631,7 @@ Qed. Definition bernoulli_and : R.-sfker T ~> mbool := (letin (sample_cst [the probability _ _ of bernoulli p12]) (letin (sample_cst [the probability _ _ of bernoulli p12]) - (ret (measurable_fun_mand macc1of3 macc2of3)))). + (ret (measurable_fun_mand macc1of3 macc2of3)))). *) Lemma bernoulli_andE t U : From c21b83e9ee808f5f48867d61f7834ac7c8fea075 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 16 Jan 2024 09:44:11 +0900 Subject: [PATCH 11/48] several admits proved --- theories/lang_syntax_examples_wip.v | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index cae3d82597..43c69487aa 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -301,7 +301,9 @@ congr (_ * _)%E. apply: eq_integral => t t01. apply: s01. by rewrite inE in t01. -Qed. +admit. +admit. +Admitted. (* Lemma casino01 : execP casino0 = execP casino1. Proof. From accbf5d46746e9be20346af2137c17de80a5969c Mon Sep 17 00:00:00 2001 From: AyumuSaito Date: Fri, 19 Jan 2024 08:45:23 +0900 Subject: [PATCH 12/48] definition beta - define beta_probability - proof beta_ge0 --- theories/lang_syntax.v | 16 ++- theories/lang_syntax_examples_wip.v | 205 +++++++++------------------- theories/prob_lang.v | 99 +++++++++++++- 3 files changed, 168 insertions(+), 152 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 0d1c28eb14..4d838f2c28 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -451,6 +451,7 @@ Inductive exp : flag -> ctx -> typ -> Type := | exp_binomial_trunc g (n : nat) : exp D g Real -> exp D g (Prob Nat) | exp_uniform g (a b : R) (ab0 : (0 < b - a)%R) : exp D g (Prob Real) +| exp_beta g (a b : nat) : exp D g (Prob Real) | exp_poisson g : nat -> exp D g Real -> exp D g Real | exp_normalize g t : exp P g t -> exp D g (Prob t) | exp_letin g t1 t2 str : exp P g t1 -> exp P ((str, t1) :: g) t2 -> @@ -486,6 +487,7 @@ Arguments exp_bernoulli {R g}. Arguments exp_bernoulli_trunc {R g} &. Arguments exp_binomial {R g}. Arguments exp_uniform {R g} &. +Arguments exp_beta {R g} &. Arguments exp_binomial_trunc {R g} &. Arguments exp_poisson {R g}. Arguments exp_normalize {R g _}. @@ -575,6 +577,7 @@ Fixpoint free_vars k g t (e : @exp R k g t) : seq string := | exp_bernoulli_trunc _ e => free_vars e | exp_binomial _ _ _ _ => [::] | exp_uniform _ _ _ _ => [::] + | exp_beta _ _ _ => [::] | exp_binomial_trunc _ _ e => free_vars e | exp_poisson _ _ e => free_vars e | exp_normalize _ _ e => free_vars e @@ -755,6 +758,9 @@ Inductive evalD : forall g t, exp D g t -> (exp_uniform a b ab0 : exp D g _) -D> cst (uniform_probability ab0) ; measurable_cst _ +| eval_beta g (a b : nat) (p : {nonneg R}) (p1 : (p%:num <= 1)%R) : + (exp_beta a b : exp D g _) -D> cst (beta a b p1) ; measurable_cst _ + | eval_poisson g n (e : exp D g _) f mf : e -D> f ; mf -> exp_poisson n e -D> poisson n \o f ; @@ -902,11 +908,11 @@ all: (rewrite {g t e u v mu mv hu}). inversion 1; subst g0 a0 b0. inj_ex H4; subst v. by have -> : ab0 = ab2. -- move=> g n e0 f mf ev IH {}v {}mv. - inversion 1; subst g0 n0. - inj_ex H2; subst e0. - inj_ex H4; subst v. - by rewrite (IH _ _ H3). +- move=> g a b p p1 {}v {}mv. + inversion 1. subst g0 a0 b0. + inj_ex H2; subst v. + inj_ex H4. + have -> : p1 = p2 by []. - move=> g t e0 k ev IH {}v {}mv. inversion 1; subst g0 t0. inj_ex H2; subst e0. diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index 43c69487aa..426d90091f 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -8,6 +8,10 @@ Require Import lebesgue_measure numfun lebesgue_integral kernel prob_lang. Require Import lang_syntax_util lang_syntax. From mathcomp Require Import ring lra. +(******************************************************************************) +(* Casino example *) +(* some steps *) + Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. @@ -19,12 +23,64 @@ Local Open Scope classical_set_scope. Local Open Scope ring_scope. (* Local Open Scope ereal_scope. *) +Section trunc_lemmas. +Open Scope ring_scope. +Open Scope lang_scope. +Context (R : realType). + +Lemma bernoulli_truncE (p : R) U : + (0 <= p <= 1)%R -> + (bernoulli_trunc p U = + p%:E * \d_true U + (`1-p)%:E * \d_false U)%E. +Proof. +move=> /andP[p0 p1]. +rewrite /bernoulli_trunc. +case: (sumbool_ler 0 p) => [{}p0/=|]. + case: (sumbool_ler p 1) => [{}p1/=|]. + by rewrite /bernoulli/= measure_addE. + by rewrite ltNge p1. +by rewrite ltNge p0. +Qed. + +Lemma __ p p1 : + @execD R [::] _ (exp_bernoulli p p1) = + execD (exp_bernoulli_trunc [{p%:num}:R]). +Proof. +apply: eq_execD. +rewrite execD_bernoulli execD_bernoulli_trunc execD_real/=. +apply: funext=> x /=. +rewrite /bernoulli_trunc. +case: (sumbool_ler 0 p%:num) => [{}p0/=|]. + case: (sumbool_ler p%:num 1) => [{}p1'/=|]. + admit. +Abort. + + +End trunc_lemmas. + Section casino_example. Open Scope ring_scope. Open Scope lang_scope. Context (R : realType). Lemma a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. +(* guard test *) +Definition test_guard : @exp R _ [::] _ := [ + let "p" := Sample {exp_bernoulli (1 / 3)%:nng (p1S 2)} in + let "_" := if #{"p"} then return TT else Score {0}:R in + return #{"p"} +]. + +Lemma exec_guard t U : execP test_guard t U = ((1 / 3)%:E * \d_true U)%E. +Proof. +rewrite /test_guard 2!execP_letin execP_sample execD_bernoulli execP_if/=. +rewrite !execP_return !exp_var'E !(execD_var_erefl "p") execD_unit execP_score execD_real/=. +rewrite letin'E ge0_integral_measure_sum//. +rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//= !integral_dirac//. +rewrite !letin'E !iteE/= integral_dirac// ge0_integral_mscale//=. +by rewrite normr0 mul0e !mule0 !adde0 !diracT !mul1e. +Qed. + (* Definition ex : exp _ [::] _ := @exp_bernoulli R [::] (1 / 2)%:nng (p1S 1). Example ex1 : projT1 (execD ex) tt = 1%:E. *) @@ -60,20 +116,6 @@ Definition casino1' : @exp R _ [::] _ := let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. -Lemma bernoulli_truncE (p : R) U : - (0 <= p <= 1)%R -> - (bernoulli_trunc p U = - p%:E * \d_true U + (`1-p)%:E * \d_false U)%E. -Proof. -move=> /andP[p0 p1]. -rewrite /bernoulli_trunc. -case: (sumbool_ler 0 p) => [{}p0/=|]. - case: (sumbool_ler p 1) => [{}p1/=|]. - by rewrite /bernoulli/= measure_addE. - by rewrite ltNge p1. -by rewrite ltNge p0. -Qed. - Lemma binomial_le1' n p U : 0 <= p <= 1 -> (\int[binomial_probability_trunc n p]_y0 \d_(0 < y0)%N U = @@ -170,94 +212,6 @@ exact: p0. apply/onem_ge0/p1. Qed. -(* exec [let x := sample (binomial n e) in *) -(* return (x >= 1)] = *) -(* exec [let y := sample (bernoulli (1 - e^n)) in *) -(* return y] *) -(* Lemma binomial_le1' x y n g e t U : - (0 <= (projT1 (execD e) t : mtyp Real) <= 1) -> - execP ([let x := Sample {exp_binomial_trunc n e} in - return {1}:N <= #x] : @exp R _ g _) t U = - execP [let y := Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - e]} ^+ n]} in - return #y] t U. -Proof. -rewrite !execP_letin !execP_sample execD_binomial_trunc execD_bernoulli_trunc/=. -rewrite !exp_var'E/=. - exact/ctx_prf_head. - exact/ctx_prf_head. -move=> H0 H1. -rewrite !execP_return !execD_rel. - have /= := @execD_var R ((y, Bool) :: g) y. - rewrite eqxx => /(_ H0) ->. - have /= := @execD_var R ((x, Nat) :: g) x. - rewrite eqxx => /(_ H1) -> /=. -rewrite (@execD_bin _ _ binop_minus) execD_pow (@execD_bin _ _ binop_minus). -rewrite !execD_real execD_nat/=. -rewrite 2!letin'E/=. -set p := projT1 (execD e) t. -move => /andP[p0 p1]. -rewrite (@integral_bernoulli_trunc _ _ (fun x => \d_x U))//; last first. - apply/andP; split. - apply/onemX_ge0; lra. - apply/onem_le1/exprn_ge0; lra. -rewrite (@integral_binomial_probabilty_trunc _ n p _ _ (fun y => \d_(1 <= y)%N U))//; last first. - (* apply: measurable_fun_dirac. - have := @subsetT _ U; rewrite setT_bool => UT. - have [->|->|->|->] /= := subset_set2 UT. - exact: measurable0. - rewrite [X in measurable X](_ : _ = `[1, +oo[%classic) //. - apply/seteqP. - admit. - rewrite [X in measurable X](_ : _ = `]-oo, 1[%classic) //. - apply/seteqP. - admit. - admit. - rewrite diracE. *) -rewrite !big_ord_recl/=. -rewrite addeC /bump/=. -congr (_ + _)%E. - under eq_bigr => i _. - have -> : (0 < 1 + i)%N => //. - over. -congr (_ * _)%E. -have -> : (0 < 0)%N = false. -rewrite /bump. -under eq_bigr => i _. - rewrite /= natrD. - have -> : 1 <= 1 + i%:R :> R. - by rewrite lerDl. - over. -rewrite addeC -ge0_sume_distrl. - congr (_ + _)%E; congr (_ * _)%E. - have -> : (\sum_(i < n) (p ^+ (n - (1 + i)) * `1-p ^+ (1 + i) *+ 'C(n, 1 + i))%:E)%E = - (\sum_(i < n.+1) (p ^+ (n - i) * `1-p ^+ i *+ 'C(n, i))%:E - (p ^+ n)%:E)%E. - rewrite big_ord_recl/= subn0 addeC addeA. - rewrite bin0 mulr1 mulr1n. - have <- : 0%E = ((- p ^+ n)%:E + (p ^+ n)%:E)%E. - rewrite EFinN. - congr _%:E. - lra. - by rewrite add0e. - congr _%E. - rewrite sumEFin. - rewrite !EFinB EFin_expe. - congr (_ - _)%E. - rewrite -(@exprDn_comm _ p `1-p n); last first. - by rewrite /GRing.comm/onem; lra. - rewrite /onem addrC. - have -> : 1 - p + p = 1 by lra. - by rewrite expr1n. - rewrite subn0 expr0 bin0 mulr1 mulr1n. - rewrite /onem. - congr _%:E. - set pn := p ^+ n. - lra. -move=> i _. -apply/mulrn_wge0/mulr_ge0; apply/exprn_ge0. -exact: p0. -apply/onem_ge0/p1. -Admitted. *) - Lemma __ : uniform_probability a01 `[0, (1 / 2)] = (1 / 2)%:E. Proof. rewrite /uniform_probability /mscale/= /mrestr. @@ -284,21 +238,18 @@ Qed. Let weak_head fl g {t1 t2} x (e : @exp R fl g t2) (xg : x \notin dom g) := exp_weak fl [::] _ (x, t1) e xg. -Lemma execP_letin_uniform g u p - (* (Hap : infer (a != p)) (Hbp : infer (b != p)) *) - (* (bg : b \notin dom ((p, Real) :: g)) (ag : b \notin dom ((b, Nat) :: (p, Real) :: g)) *) - (s0 s1 : exp P ((p, Real) :: g) u) : - (forall (t : R) x U, 0 <= t <= 1 -> - execP s0 (t, x) U = execP s1 (t, x) U) -> +Lemma execP_letin_uniform g t str (s0 s1 : exp P ((str, Real) :: g) t) : + (forall (p : R) x U, 0 <= p <= 1 -> + execP s0 (p, x) U = execP s1 (p, x) U) -> forall x U, measurable U -> - execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s0}] x U = - execP [let p := Sample {@exp_uniform _ g 0 1 a01} in {s1}] x U. + execP [let str := Sample {@exp_uniform _ g 0 1 a01} in {s0}] x U = + execP [let str := Sample {@exp_uniform _ g 0 1 a01} in {s1}] x U. Proof. move=> s01 x U mU. rewrite !execP_letin execP_sample execD_uniform/=. rewrite !letin'_sample_uniform//. congr (_ * _)%E. -apply: eq_integral => t t01. +apply: eq_integral => p p01. apply: s01. by rewrite inE in t01. admit. @@ -360,21 +311,10 @@ Qed. *) Lemma casino01' y V : measurable V -> execP casino0 y V = execP casino1' y V. Proof. -move=> mV. +move=> mV //. rewrite /casino0 /casino1. -pose s0 := - [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in - return {1}:N <= #{"a2"}]. -(* pose s1 := - [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - return {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%N}]}]. *) -have := (@execP_letin_uniform [::] Bool "p" (s0 R (found "p" Real [::]) _ _) _). -apply => //. +apply: execP_letin_uniform => //. move=> p x U r01. -rewrite /s0/=. rewrite !execP_letin !execP_sample !execD_binomial_trunc /=. rewrite execP_if execP_score !execP_return !execD_bernoulli_trunc/=. rewrite !execD_rel (@execD_bin _ _ binop_minus) execD_pow. @@ -394,23 +334,6 @@ rewrite !letin'E/=. by apply/binomial_le1'/andP. Qed. -(* guard test *) -Definition test_guard : @exp R _ [::] _ := [ - let "p" := Sample {exp_bernoulli (1 / 3)%:nng (p1S 2)} in - let "_" := if #{"p"} then return TT else Score {0}:R in - return #{"p"} -]. - -Lemma exec_guard t U : execP test_guard t U = ((1 / 3)%:E * \d_true U)%E. -Proof. -rewrite /test_guard 2!execP_letin execP_sample execD_bernoulli execP_if/=. -rewrite !execP_return !exp_var'E !(execD_var_erefl "p") execD_unit execP_score execD_real/=. -rewrite letin'E ge0_integral_measure_sum//. -rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//= !integral_dirac//. -rewrite !letin'E !iteE/= integral_dirac// ge0_integral_mscale//=. -by rewrite normr0 mul0e !mule0 !adde0 !diracT !mul1e. -Qed. - Lemma exec_casino t U : execP casino0 t U = ((10 / 99)%:E * \d_true U + (1 / 99)%:E * \d_false U)%E . Proof. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 4f409d920d..f2b6dcbabc 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -190,9 +190,9 @@ HB.instance Definition _ := Probability.on bernoulli0. Definition bernoulli_trunc (p : R) := match sumbool_ler 0%R p with | inl l0p => match sumbool_ler (NngNum l0p)%:num 1%R with | inl lp1 => [the probability _ _ of @bernoulli R (NngNum l0p) lp1] - | inr _ => [the probability _ _ of bernoulli0] + | inr _ => bernoulli0 end -| inr _ => [the probability _ _ of bernoulli0] +| inr _ => bernoulli0 end. (*HB.instance Definition _ (p : R) := Probability.on (bernoulli_trunc p).*) @@ -275,7 +275,7 @@ Section binomial_probability. Context {R : realType} (n : nat) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). Local Open Scope ring_scope. -(* C(n, k) * p^(n-k) * (1-p)^k *) +(* C(n, k) * p^k * (1-p)^(n-k) *) Definition bino_term (k : nat) : {nonneg R} := (p%:num^+k * (NngNum (onem_ge0 p1))%:num^+(n-k)%N *+ 'C(n, k))%:nng. @@ -301,9 +301,7 @@ Import Notations. [the measure _ _ of \d_k]]). *) (* \sum_(k < n.+1) (bino_coef p n k) * \d_k. *) Definition binomial_probability : set nat -> \bar R := - @msum _ _ R - (fun k => [the measure _ _ of mscale (bino_term k) - [the measure _ _ of \d_k]]) n.+1. + @msum _ _ R (fun k => mscale (bino_term k) \d_k) n.+1. HB.instance Definition _ := Measure.on binomial_probability. @@ -602,6 +600,95 @@ Qed. End integral_uniform. +Section beta_probability. +Context {R : realType}. +Local Open Scope ring_scope. +Variables (a b : nat). + +Definition Beta t : R := + t^+(a-1) * (1-t)^+(b-1). + +Definition B : R := + (a-1)`!%:R * (b-1)`!%:R / (a+b-1)`!%:R. + (* \int[lebesgue_measure]_(t in `[0%R, 1%R]) Beta t. *) + +Axiom beta1 : \int[lebesgue_measure]_(x in `[0, 1]) (x^+(a-1) * (1-x)^+(b-1)) = B. + +(* apply: fine_ge0. +apply: integral_ge0 => x x01. +have /andP[x0 x1] : (0 <= x <= 1) by apply: x01. +rewrite lee_fin. +apply: mulr_ge0; apply/exprn_ge0; lra. *) + +Set Printing All. + +Lemma beta_ge0 : 0 <= B. +Proof. +rewrite /B. +Admitted. + +(* Definition betaPDF (p : R) := + if 0 <= p <= 1 then Beta p / B else 0. + +Lemma measurable_beta : measurable_fun setT betaPDF. +Proof. +rewrite /betaPDF; apply: measurable_fun_if => //. + apply: (measurable_fun_bool true). + by rewrite (_ : _ @^-1` _ = `[0, 1]%classic). +apply: measurable_funM=> //; apply /measurable_funM => //. +apply/measurable_fun_pow/measurable_funB => //. +Qed. *) + +(* TODO: beta_ge0 a b : is_true (0 <= beta a b) (to remove %E) *) +Let p01 : 0 < 1 - 0 :> R. +Proof. lra. Qed. + +(* 1/B(a, b) * \int_U p^(a-1) * (1-p)^(b-1) dx = beta *) +Definition beta (U : set _) : \bar R := + (1 / B)%:E * \int[uniform_probability p01]_(t in U) (t^+(a-1) * (1-t)^+(b-1))%:E. + (* \int[lebesgue_measure]_(t in U) + @mscale _ _ R (t^+(a-1) + (* * (NngNum (onem_ge0 p1))%:num^+(b-1) *) + * + (invr_nonneg (NngNum beta_ge0))%:num)%:nng + (mrestr lebesgue_measure (measurable_itv `[0, 1])) U. *) + +Example __ : beta `[0, 1] = 1%:E. +Proof. +rewrite /beta integral_mkcond. +rewrite integral_uniform//=. +rewrite oppr0 addr0 invr1 mul1e. +Admitted. + +HB.instance Definition _ (p : {nonneg R}) (p1 : p%:num <= 1) + := Measure.on (beta p1). + +Let beta_setT (p : {nonneg R}) (p1 : p%:num <= 1) + : beta p1 [set: _] = 1%:E. +Proof. +rewrite /beta /mscale/= /mrestr/=. +rewrite setTI lebesgue_measure_itv/= lte_fin. +rewrite ltr01 oppr0 adde0 mule1 /B /Beta. +(* rewrite -subr_gt0 -EFinD . -EFinM mulVf// gt_eqF// subr_gt0. *) +Admitted. + +HB.instance Definition _ (p : {nonneg R}) (p1 : p%:num <= 1) := @Measure_isProbability.Build _ _ R + (beta p1) (beta_setT p1). + +(* Lemma __ : beta_probability 6 4 (p1S 2) `[0, 1] = 1%:E. +Proof. +rewrite /beta_probability/beta/= /mscale/=. +rewrite lebesgue_measure_itv. *) + +Let H01 : (0 < 1 - 0 :> R). +Proof. lra. Qed. + +(* Lemma uniform_beta U : (\int[uniform_probability H01]_x + \int[mscale (NngNum (normr_ge0 (56 * x ^+ 5 * (1 - x) ^+ 3)%R)) (measure_dirac__canonical__measure_Measure tt R)]__ + bernoulli_trunc (1 - (1 - x) ^+ 3) U)%E = 1%:E. *) + +End beta_probability. + Section mscore. Context d (T : measurableType d) (R : realType). Variable f : T -> R. From 74e8f124b82deb3d79af61ea00fd9f4663e8dc51 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 15 Feb 2024 13:21:17 +0900 Subject: [PATCH 13/48] fix --- theories/lang_syntax_examples_wip.v | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index 426d90091f..4edd345267 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -251,10 +251,8 @@ rewrite !letin'_sample_uniform//. congr (_ * _)%E. apply: eq_integral => p p01. apply: s01. -by rewrite inE in t01. -admit. -admit. -Admitted. +by rewrite inE in p01. +Qed. (* Lemma casino01 : execP casino0 = execP casino1. Proof. From a60848c04e4fa956a70548c48d7e76a66f5c45a3 Mon Sep 17 00:00:00 2001 From: AyumuSaito Date: Fri, 16 Feb 2024 10:48:05 +0900 Subject: [PATCH 14/48] write casino (wip) - add exp_beta --- theories/lang_syntax.v | 108 ++++++++++++++++++++++++---- theories/lang_syntax_examples_wip.v | 60 ++++++++++++++++ theories/prob_lang.v | 55 +++++++++++--- 3 files changed, 201 insertions(+), 22 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 4d838f2c28..e6eaf50e6b 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -346,6 +346,75 @@ End accessor_functions. Arguments acc_typ {R} s n. Arguments measurable_acc_typ {R} s n. + +Section beta. +Context {R : realType}. +(* Check sample_cst (beta 6 4) : R.-sfker _ ~> R. *) +(* Check sample_cst (beta 6 4) : R.-ker _ ~> measurableTypeR R. *) +Check sample_cst (uniform_probability _) : R.-ker _ ~> measurableTypeR R. + +Open Scope ring_scope. +Lemma a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. +Import Notations. + +(* Lemma beta11_uniform : + beta 1 1 `[0, 1] = uniform_probability a01 `[0, 1]. +Proof. +rewrite /beta /uniform_probability. +congr mscale. +congr invr_nonneg. +admit. +rewrite /mscale. +apply: funext=> x. *) + +Definition beta_bern : R.-sfker munit ~> mbool := + @letin' _ _ _ munit _ mbool R + (sample_cst (beta 6 4)) + (* (sample_cst (uniform_probability a01)) *) + (sample (bernoulli_trunc \o (@fst (measurableTypeR R) _)) (measurableT_comp measurable_bernoulli_trunc (measurable_acc_typ [:: Real] 0))). + +Lemma letin'_sample_uniform d d' (T : measurableType d) + (T' : measurableType d') (a b : R) (ab0 : (0 < b - a)%R) + (u : R.-sfker [the measurableType _ of (_ * T)%type] ~> T') x y : + measurable y -> + letin' (sample_cst (uniform_probability ab0)) u x y = + ((b - a)^-1%:E * \int[lebesgue_measure]_(x0 in `[a, b]) u (x0, x) y)%E. +Admitted. + +Definition uni_bern : R.-sfker munit ~> mbool := + @letin' _ _ _ munit (measurableTypeR R) mbool R + (sample_cst (uniform_probability a01)) + (* (sample_cst (uniform_probability a01)) *) + (sample (bernoulli_trunc \o (@fst (measurableTypeR R) _)) (measurableT_comp measurable_bernoulli_trunc (measurable_acc_typ [:: Real] 0))). + +Lemma ex_beta_bern : beta_bern tt [set true] = (6 / 10)%:E. +Proof. +rewrite /beta_bern /uni_bern. +rewrite [LHS]letin'E. +rewrite /beta. +rewrite ge0_integral_mscale//. +rewrite /mscale/=/B. +rewrite /prebeta/=. +Search integral lebesgue_measure. + +(* rewrite ge0_integral_mscale//=. +rewrite EFinM. +congr (_ * _)%E. +rewrite /prebeta/=. +Search "pow". +rewrite subn1/=. +transitivity (\int[(integral (uniform_probability p01))^~ (@cst R _ 1%:E)]_x bernoulli_trunc x U)%E. + admit. +rewrite /bernoulli_trunc/=. +rewrite integral_bernoulli_trunc. +(* rewrite /B invr1 !mulr1 fact0 invr1 mul1e. *) +rewrite /prebeta/=. +under eq_integral. +Search integral lebesgue_measure. *) +Abort. + +End beta. + Section context. Variables (R : realType). Definition ctx := seq (string * typ). @@ -451,7 +520,7 @@ Inductive exp : flag -> ctx -> typ -> Type := | exp_binomial_trunc g (n : nat) : exp D g Real -> exp D g (Prob Nat) | exp_uniform g (a b : R) (ab0 : (0 < b - a)%R) : exp D g (Prob Real) -| exp_beta g (a b : nat) : exp D g (Prob Real) +| exp_beta g (a b : nat) (* NB: shound be R *) : exp D g (Prob Real) | exp_poisson g : nat -> exp D g Real -> exp D g Real | exp_normalize g t : exp P g t -> exp D g (Prob t) | exp_letin g t1 t2 str : exp P g t1 -> exp P ((str, t1) :: g) t2 -> @@ -758,8 +827,8 @@ Inductive evalD : forall g t, exp D g t -> (exp_uniform a b ab0 : exp D g _) -D> cst (uniform_probability ab0) ; measurable_cst _ -| eval_beta g (a b : nat) (p : {nonneg R}) (p1 : (p%:num <= 1)%R) : - (exp_beta a b : exp D g _) -D> cst (beta a b p1) ; measurable_cst _ +| eval_beta g (a b : nat) : + (exp_beta a b : exp D g _) -D> cst (beta a b) ; measurable_cst _ | eval_poisson g n (e : exp D g _) f mf : e -D> f ; mf -> @@ -896,7 +965,7 @@ all: (rewrite {g t e u v mu mv hu}). by rewrite (IH _ _ H2). - move=> g n p p1 {}v {}mv. inversion 1; subst g0 n0 p0. - inj_ex H4; subst v. + inj_ex H2; subst v. by have -> : p1 = p3 by []. - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. @@ -906,17 +975,21 @@ all: (rewrite {g t e u v mu mv hu}). by rewrite (IH _ _ H3). - move=> g a b ab0 {}v {}mv. inversion 1; subst g0 a0 b0. - inj_ex H4; subst v. + inj_ex H2; subst v. by have -> : ab0 = ab2. -- move=> g a b p p1 {}v {}mv. +- move=> g a b {}v {}mv. inversion 1. subst g0 a0 b0. - inj_ex H2; subst v. - inj_ex H4. - have -> : p1 = p2 by []. -- move=> g t e0 k ev IH {}v {}mv. - inversion 1; subst g0 t0. + by inj_ex H2; subst v. (* TODO: beta *) +- move=> g t e k mk ev IH {}v {}mv. + inversion 1; subst g0 t. inj_ex H2; subst e0. inj_ex H4; subst v. + by rewrite (IH _ _ H3). +- move=> g t e k ev IH f mf. + inversion 1; subst g0 t0. + inj_ex H2; subst e0. + inj_ex H4; subst f. + inj_ex H5; subst mf. by rewrite (IH _ H3). - move=> g t e f mf e1 f1 mf1 e2 f2 mf2 ev ih ev1 ih1 ev2 ih2 v m. inversion 1; subst g0 t0. @@ -1050,7 +1123,7 @@ all: rewrite {g t e u v eu}. by rewrite (IH _ _ H2). - move=> g n p p1 {}v {}mv. inversion 1; subst g0 n0 p0. - inj_ex H4; subst v. + inj_ex H2; subst v. by have -> : p1 = p3 by []. - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. @@ -1060,8 +1133,11 @@ all: rewrite {g t e u v eu}. by rewrite (IH _ _ H3). - move=> g a b ab0 {}v {}mv. inversion 1; subst g0 a0 b0. - inj_ex H4; subst v. + inj_ex H2; subst v. by have -> : ab0 = ab2. +- move=> g a b {}v {}mv. + inversion 1; subst g0 a0 b0. + by inj_ex H2; subst v. - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. inj_ex H2; subst e0. @@ -1163,6 +1239,7 @@ all: rewrite {z g t}. exists (binomial_probability_trunc n \o p). eexists; exact: (eval_binomial_trunc n). - by eexists; eexists; exact: eval_uniform. +- by eexists; eexists; exact: eval_beta. - move=> g h e [f [mf H]]. by exists (poisson h \o f); eexists; exact: eval_poisson. - move=> g t e [k ek]. @@ -1358,6 +1435,11 @@ Lemma execD_uniform g a b ab0 : existT _ (cst [the probability _ _ of uniform_probability ab0]) (measurable_cst _). Proof. exact/execD_evalD/eval_uniform. Qed. +Lemma execD_beta g a b : + @execD g _ (exp_beta a b) = + existT _ (cst [the probability _ _ of beta a b]) (measurable_cst _). +Proof. exact/execD_evalD/eval_beta. Qed. + Lemma execD_normalize_pt g t (e : exp P g t) : @execD g _ [Normalize e] = existT _ (normalize_pt (execP e) : _ -> pprobability _ _) diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index 4edd345267..d96074b1c8 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -414,4 +414,64 @@ congr (_%:E * _)%E. lra. Qed. +Definition casino2 : @exp R _ [::] _ := + [let "p" := Sample {exp_uniform 0 1 a01} in + let "_" := Score {[{56}:R * #{"p"} ^+ {5%nat} * {[{1}:R - #{"p"}]} ^+ {3%nat}]} in + Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. + +Lemma casino12 y V : + measurable V -> + execP casino1' y V = execP casino2 y V. +Proof. +move=> mV. +rewrite /casino1' /casino2. +apply: execP_letin_uniform => //. +move=> p x U /andP[p0 p1]. +rewrite !execP_letin !execP_sample execP_if execD_rel/=. +rewrite !execP_score !(@execD_bin _ _ binop_mult). +rewrite !execD_bernoulli_trunc/= !(@execD_bin _ _ binop_minus) !execD_pow. +rewrite !(@execD_bin _ _ binop_minus)/=. +rewrite !execD_real !execD_nat/= execP_return execD_unit. +rewrite !execD_binomial_trunc/=. +rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a1")/=. +rewrite !letin'E/=. +rewrite integral_binomial_probabilty_trunc//=. +(* set s := letin' _ _. *) +rewrite (bigD1 (inord 5))//=. + rewrite big1; last first. + move=> [[|[|[|[|[|[|[|[|[|//]]]]]]]]]]//= Hi Hi5; rewrite letin'E iteE; + rewrite ?ge0_integral_mscale//= ?normr0 ?mul0e ?mule0 ?add0e//. + suff: false by []. + move/negbTE: Hi5 => <-. + by apply/eqP/val_inj => /=; rewrite inordK. +rewrite letin'E iteE ge0_integral_mscale//= inordK//= adde0 /onem. +congr (_ * _)%E. +rewrite ger0_norm. + by rewrite -mulrA mulr_natl. +apply/mulr_ge0. + exact/mulr_ge0/exprn_ge0. +apply/exprn_ge0. +by rewrite subr_ge0. +Qed. + +Definition casino3 : @exp R _ [::] _ := + [let "_" := Score {1 / 9}:R in + let "p" := Sample {exp_beta 6 4} in + Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. + +Definition casino4 : @exp R _ [::] _ := + [let "_" := Score {1 / 9}:R in + Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. + +Definition casino5 : @exp R _ [::] _ := + [Normalize Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. + +(* Lemma casino5_ans U : projT1 (execD casino5) tt U = + ((10 / 11)%:E * \d_true U + (1 / 11)%:E * \d_false U)%E. +Proof. +rewrite /casino5 execD_normalize_pt execP_sample execD_bernoulli_trunc/=. +rewrite execD_real/=. +rewrite normalizeE/= bernoulli_truncE. +rewrite /mnormalize/=. *) + End casino_example. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index f2b6dcbabc..629c3aed2e 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -622,7 +622,7 @@ apply: mulr_ge0; apply/exprn_ge0; lra. *) Set Printing All. -Lemma beta_ge0 : 0 <= B. +Lemma B_ge0 : 0 <= B. Proof. rewrite /B. Admitted. @@ -640,12 +640,39 @@ apply/measurable_fun_pow/measurable_funB => //. Qed. *) (* TODO: beta_ge0 a b : is_true (0 <= beta a b) (to remove %E) *) -Let p01 : 0 < 1 - 0 :> R. -Proof. lra. Qed. +Definition p01 : 0 < 1 - 0 :> R. +Proof. by []. Qed. + +Definition prebeta U : \bar R := + (* \int[lebesgue_measure]_(x in U `&` `[0, 1]) (x^+(a-1) * (1-x)^+(b-1))%:E. *) + \int[uniform_probability p01]_(t in U) (t^+(a-1) * (1-t)^+(b-1))%:E. + +Let prebeta0 : prebeta set0 = 0%:E. +Proof. +(* by rewrite /prebeta integral_setI_indic// integral_set0. *) +by rewrite /prebeta integral_set0. +Qed. + +Let prebeta_ge0 U : (0 <= prebeta U)%E. +Proof. +rewrite /prebeta. +rewrite integral_ge0// => x Ux. +Admitted. + +Let prebeta_sigma_additive : semi_sigma_additive prebeta. +Proof. move=> /= F mF tF mUF. +rewrite /prebeta. +Admitted. + +HB.instance Definition _ := isMeasure.Build _ _ _ prebeta + prebeta0 prebeta_ge0 prebeta_sigma_additive. + +(* HB.instance Definition _ := Measure.on prebeta. *) (* 1/B(a, b) * \int_U p^(a-1) * (1-p)^(b-1) dx = beta *) -Definition beta (U : set _) : \bar R := - (1 / B)%:E * \int[uniform_probability p01]_(t in U) (t^+(a-1) * (1-t)^+(b-1))%:E. +Definition beta : set _ -> \bar R := + @mscale _ _ R (invr_nonneg (NngNum B_ge0)) prebeta. + (* \int[lebesgue_measure]_(t in U) @mscale _ _ R (t^+(a-1) (* * (NngNum (onem_ge0 p1))%:num^+(b-1) *) @@ -653,14 +680,24 @@ Definition beta (U : set _) : \bar R := (invr_nonneg (NngNum beta_ge0))%:num)%:nng (mrestr lebesgue_measure (measurable_itv `[0, 1])) U. *) +HB.instance Definition _ := Measure.on beta. + +Let beta_setT : beta [set: _] = 1%:E. +Proof. +rewrite /beta /mscale/= /mrestr/=. +Admitted. + +HB.instance Definition _ := @Measure_isProbability.Build _ _ R + beta beta_setT. + Example __ : beta `[0, 1] = 1%:E. Proof. -rewrite /beta integral_mkcond. +(* rewrite /beta integral_mkcond. rewrite integral_uniform//=. -rewrite oppr0 addr0 invr1 mul1e. +rewrite oppr0 addr0 invr1 mul1e. *) Admitted. -HB.instance Definition _ (p : {nonneg R}) (p1 : p%:num <= 1) +(* HB.instance Definition _ (p : {nonneg R}) (p1 : p%:num <= 1) := Measure.on (beta p1). Let beta_setT (p : {nonneg R}) (p1 : p%:num <= 1) @@ -673,7 +710,7 @@ rewrite ltr01 oppr0 adde0 mule1 /B /Beta. Admitted. HB.instance Definition _ (p : {nonneg R}) (p1 : p%:num <= 1) := @Measure_isProbability.Build _ _ R - (beta p1) (beta_setT p1). + (beta p1) (beta_setT p1). *) (* Lemma __ : beta_probability 6 4 (p1S 2) `[0, 1] = 1%:E. Proof. From fb91523e9eabcf5de5e977c757b83a5a21a0bded Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 20 Feb 2024 12:42:59 +0900 Subject: [PATCH 15/48] WIP - clean up def of beta - beta_nat_bern_bern --- _CoqProject | 1 + theories/Make | 1 + theories/lang_syntax.v | 203 ++++++++++++++++++---------- theories/lang_syntax_examples_wip.v | 13 +- theories/prob_lang.v | 196 +++++++++++++++------------ 5 files changed, 250 insertions(+), 164 deletions(-) diff --git a/_CoqProject b/_CoqProject index 4c99e218df..0936ccbc72 100644 --- a/_CoqProject +++ b/_CoqProject @@ -142,3 +142,4 @@ theories/lang_syntax_util.v theories/lang_syntax_toy.v theories/lang_syntax.v theories/lang_syntax_examples.v +theories/lang_syntax_examples_wip.v diff --git a/theories/Make b/theories/Make index 43b739a23d..80466f4fc1 100644 --- a/theories/Make +++ b/theories/Make @@ -103,3 +103,4 @@ lang_syntax_util.v lang_syntax_toy.v lang_syntax.v lang_syntax_examples.v +lang_syntax_examples_wip.v diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index e6eaf50e6b..6afbcce8bd 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -3,10 +3,11 @@ From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. From mathcomp Require Import interval_inference. From mathcomp Require Import mathcomp_extra boolp classical_sets. -From mathcomp Require Import functions cardinality fsbigop. -From mathcomp Require Import reals ereal topology normedtype sequences esum. -From mathcomp Require Import measure lebesgue_measure numfun lebesgue_integral. -From mathcomp Require Import kernel prob_lang lang_syntax_util. +From mathcomp Require Import functions cardinality fsbigop interval_inference. +From mathcomp Require Import reals ereal topology normedtype sequences. +From mathcomp Require Import esum measure lebesgue_measure numfun. +From mathcomp Require Import lebesgue_integral kernel prob_lang. +From mathcomp Require Import lang_syntax_util. From mathcomp Require Import ring lra. (******************************************************************************) @@ -346,74 +347,132 @@ End accessor_functions. Arguments acc_typ {R} s n. Arguments measurable_acc_typ {R} s n. - -Section beta. +(* part of the step from casino3 to casino4 *) +Section casino3_casino4. Context {R : realType}. -(* Check sample_cst (beta 6 4) : R.-sfker _ ~> R. *) -(* Check sample_cst (beta 6 4) : R.-ker _ ~> measurableTypeR R. *) -Check sample_cst (uniform_probability _) : R.-ker _ ~> measurableTypeR R. +Variables a b a' b' : nat. + +Local Notation mu := lebesgue_measure. Open Scope ring_scope. -Lemma a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. Import Notations. -(* Lemma beta11_uniform : - beta 1 1 `[0, 1] = uniform_probability a01 `[0, 1]. +Program Definition beta_nat_bern : R.-sfker munit ~> mbool := + @letin' _ _ _ munit _ mbool _ + (sample_cst [the probability _ _ of (beta_nat a b)]) + (sample (bernoulli_trunc \o ubeta_nat_pdf a' b' \o @fst (measurableTypeR R) _ ) _). +Next Obligation. +apply: measurableT_comp. + apply: measurableT_comp. + exact: measurable_bernoulli_trunc. + exact: measurable_ubeta_nat_pdf. +exact: (measurable_acc_typ [:: Real] 0). +Qed. + +Local Notation B := beta_nat_norm. + +Definition Baa'bb'Bab : R := (B (a + a') (b + b')) / B a b. + +Lemma Baa'bb'Bab_ge0 : 0 <= Baa'bb'Bab. +Proof. by rewrite /Baa'bb'Bab divr_ge0// beta_nat_norm_ge0. Qed. + +Definition Baa'bb'Bab_nneg : {nonneg R} := NngNum Baa'bb'Bab_ge0. + +Lemma Baa'bb'Bab_le1 : Baa'bb'Bab_nneg%:num <= 1. Proof. -rewrite /beta /uniform_probability. -congr mscale. -congr invr_nonneg. -admit. -rewrite /mscale. -apply: funext=> x. *) - -Definition beta_bern : R.-sfker munit ~> mbool := - @letin' _ _ _ munit _ mbool R - (sample_cst (beta 6 4)) - (* (sample_cst (uniform_probability a01)) *) - (sample (bernoulli_trunc \o (@fst (measurableTypeR R) _)) (measurableT_comp measurable_bernoulli_trunc (measurable_acc_typ [:: Real] 0))). - -Lemma letin'_sample_uniform d d' (T : measurableType d) - (T' : measurableType d') (a b : R) (ab0 : (0 < b - a)%R) - (u : R.-sfker [the measurableType _ of (_ * T)%type] ~> T') x y : - measurable y -> - letin' (sample_cst (uniform_probability ab0)) u x y = - ((b - a)^-1%:E * \int[lebesgue_measure]_(x0 in `[a, b]) u (x0, x) y)%E. +rewrite /Baa'bb'Bab_nneg/= /Baa'bb'Bab. +rewrite ler_pdivrMr// ?mul1r ?beta_nat_norm_gt0//. +rewrite /B /beta_nat_norm. +rewrite ler_pdivrMr ?ltr0n ?fact_gt0//. +rewrite mulrAC. +rewrite ler_pdivlMr ?ltr0n ?fact_gt0//. +(* TODO: plausible *) Admitted. -Definition uni_bern : R.-sfker munit ~> mbool := - @letin' _ _ _ munit (measurableTypeR R) mbool R - (sample_cst (uniform_probability a01)) - (* (sample_cst (uniform_probability a01)) *) - (sample (bernoulli_trunc \o (@fst (measurableTypeR R) _)) (measurableT_comp measurable_bernoulli_trunc (measurable_acc_typ [:: Real] 0))). +Lemma onem_Baa'bb'Bab_ge0 : 0 <= B a b - (Baa'bb'Bab_nneg%:num). +Admitted. + +Lemma ubeta_nat_pdf_ge0' t : 0 <= ubeta_nat_pdf a' b' t :> R. +Admitted. + +Lemma ubeta_nat_pdf_le1' t : (NngNum (ubeta_nat_pdf_ge0' t))%:num <= 1 :> R. +Admitted. + +Lemma integral_ubeta_nat : + (\int[ubeta_nat a b]_x (ubeta_nat_pdf a' b' x)%:E = + \int[mu]_(x in `[0%R, 1%R]) + (x ^+ a'.-1 * `1-x ^+ b'.-1 * x ^+ a * `1-x ^+ b)%:E :> \bar R)%E. +Admitted. -Lemma ex_beta_bern : beta_bern tt [set true] = (6 / 10)%:E. +Lemma beta_nat_bern_bern U : + beta_nat_bern tt U = + sample_cst (bernoulli Baa'bb'Bab_le1) tt U. Proof. -rewrite /beta_bern /uni_bern. -rewrite [LHS]letin'E. -rewrite /beta. -rewrite ge0_integral_mscale//. -rewrite /mscale/=/B. -rewrite /prebeta/=. -Search integral lebesgue_measure. - -(* rewrite ge0_integral_mscale//=. -rewrite EFinM. -congr (_ * _)%E. -rewrite /prebeta/=. -Search "pow". -rewrite subn1/=. -transitivity (\int[(integral (uniform_probability p01))^~ (@cst R _ 1%:E)]_x bernoulli_trunc x U)%E. +rewrite /beta_nat_bern. +rewrite [LHS]letin'E/=. +transitivity ((\int[beta_nat a b]_y + (@bernoulli R (NngNum (ubeta_nat_pdf_ge0' y)) (ubeta_nat_pdf_le1' y) U))%E). + apply: eq_integral => /= r _. + rewrite /bernoulli_trunc/=. + case: sumbool_ler; last admit. + move=> H_ge0. + case: sumbool_ler; last admit. + move=> H_le1. + rewrite /=. admit. -rewrite /bernoulli_trunc/=. -rewrite integral_bernoulli_trunc. -(* rewrite /B invr1 !mulr1 fact0 invr1 mul1e. *) -rewrite /prebeta/=. -under eq_integral. -Search integral lebesgue_measure. *) -Abort. +rewrite /beta_nat/=. +transitivity ((((B a b)^-1)%:E * \int[ubeta_nat a b]_x bernoulli (ubeta_nat_pdf_le1' x) U)%E). + apply: ge0_integral_mscale => //=. + admit. +rewrite {2}/bernoulli. +suff: (\int[ubeta_nat a b]_x bernoulli (ubeta_nat_pdf_le1' x) U)%E = + measure_add + (mscale (NngNum (beta_nat_norm_ge0 (a + a') (b + b'))) \d_true) + (mscale (NngNum (onem_Baa'bb'Bab_ge0)) \d_false) U. + admit. +transitivity (( + \int[ubeta_nat a b]_x (mscale (NngNum (ubeta_nat_pdf_ge0' x)) \d_true U) + + + \int[ubeta_nat a b]_x (mscale (onem_nonneg (ubeta_nat_pdf_le1' x)) \d_false) U))%E. + admit. +rewrite measure_addE. +congr adde. + rewrite [in LHS]/mscale/= [in RHS]/mscale/=. + suff: (\int[ubeta_nat a b]_x ((ubeta_nat_pdf a' b' x)%:E) + = (B (a + a') (b + b'))%:E :> \bar R)%E. + admit. + transitivity ( (* calculating distribution (13) *) + \int[mu]_(x in `[(0:R)%R, (1:R)%R]%classic) + ((x ^+ a'.-1 * (`1- x) ^+ b'.-1 * x ^+ a * (`1- x) ^+ b)%:E) + )%E. + by rewrite integral_ubeta_nat. + transitivity ( + (\int[mu]_(x in `[0%R, 1%R]) + ((x ^+ (a + a').-1 * `1-x ^+ (b + b').-1)%:E))%E : \bar R). + admit. + rewrite beta_nat_normE//. + rewrite /ubeta_nat_pdf//. + admit. +rewrite [in LHS]/mscale/= [in RHS]/mscale/=. +suff: (\int[ubeta_nat a b]_x (`1-(ubeta_nat_pdf a' b' x))%:E = + (B a b - B (a + a') (b + b'))%:E :> \bar R)%E. + admit. +under eq_integral do rewrite EFinB/=. +rewrite integralB_EFin//=; last 2 first. + admit. + admit. +rewrite integral_cst//= mul1e. +rewrite beta_nat_normE. +rewrite {1}/ubeta_nat setTI. +rewrite EFinB. +congr (_ - _)%E. + admit. +rewrite integral_ubeta_nat//. +rewrite beta_nat_normE /ubeta_nat_pdf. +admit. +Admitted. -End beta. +End casino3_casino4. Section context. Variables (R : realType). @@ -828,7 +887,7 @@ Inductive evalD : forall g t, exp D g t -> measurable_cst _ | eval_beta g (a b : nat) : - (exp_beta a b : exp D g _) -D> cst (beta a b) ; measurable_cst _ + (exp_beta a b : exp D g _) -D> (cst (beta_nat a b)) ; measurable_cst _ | eval_poisson g n (e : exp D g _) f mf : e -D> f ; mf -> @@ -965,7 +1024,7 @@ all: (rewrite {g t e u v mu mv hu}). by rewrite (IH _ _ H2). - move=> g n p p1 {}v {}mv. inversion 1; subst g0 n0 p0. - inj_ex H2; subst v. + inj_ex H4; subst v. by have -> : p1 = p3 by []. - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. @@ -975,11 +1034,11 @@ all: (rewrite {g t e u v mu mv hu}). by rewrite (IH _ _ H3). - move=> g a b ab0 {}v {}mv. inversion 1; subst g0 a0 b0. - inj_ex H2; subst v. + inj_ex H4; subst v. by have -> : ab0 = ab2. -- move=> g a b {}v {}mv. - inversion 1. subst g0 a0 b0. - by inj_ex H2; subst v. (* TODO: beta *) +- (* TODO: beta *) move=> g a b {}v {}mv. + inversion 1; subst g0 a0 b0. + by inj_ex H4; subst v. - move=> g t e k mk ev IH {}v {}mv. inversion 1; subst g0 t. inj_ex H2; subst e0. @@ -1123,7 +1182,7 @@ all: rewrite {g t e u v eu}. by rewrite (IH _ _ H2). - move=> g n p p1 {}v {}mv. inversion 1; subst g0 n0 p0. - inj_ex H2; subst v. + inj_ex H4; subst v. by have -> : p1 = p3 by []. - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. @@ -1133,11 +1192,11 @@ all: rewrite {g t e u v eu}. by rewrite (IH _ _ H3). - move=> g a b ab0 {}v {}mv. inversion 1; subst g0 a0 b0. - inj_ex H2; subst v. + inj_ex H4; subst v. by have -> : ab0 = ab2. -- move=> g a b {}v {}mv. +- (* TODO: beta case*) move=> g a b {}v {}mv. inversion 1; subst g0 a0 b0. - by inj_ex H2; subst v. + by inj_ex H4; subst v. - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. inj_ex H2; subst e0. @@ -1435,9 +1494,9 @@ Lemma execD_uniform g a b ab0 : existT _ (cst [the probability _ _ of uniform_probability ab0]) (measurable_cst _). Proof. exact/execD_evalD/eval_uniform. Qed. -Lemma execD_beta g a b : +Lemma execD_beta_nat g a b : @execD g _ (exp_beta a b) = - existT _ (cst [the probability _ _ of beta a b]) (measurable_cst _). + existT _ (cst [the probability _ _ of beta_nat a b]) (measurable_cst _). Proof. exact/execD_evalD/eval_beta. Qed. Lemma execD_normalize_pt g t (e : exp P g t) : diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index d96074b1c8..8af7c301c3 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -1,12 +1,12 @@ Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. -From mathcomp.classical Require Import mathcomp_extra boolp classical_sets. -From mathcomp.classical Require Import functions cardinality fsbigop. -Require Import signed reals ereal topology normedtype sequences esum measure. -Require Import lebesgue_measure numfun lebesgue_integral kernel prob_lang. -Require Import lang_syntax_util lang_syntax. From mathcomp Require Import ring lra. +From mathcomp Require Import mathcomp_extra boolp classical_sets functions. +From mathcomp Require Import cardinality fsbigop interval_inference reals. +From mathcomp Require Import ereal topology normedtype sequences esum. +From mathcomp Require Import measure lebesgue_measure numfun lebesgue_integral. +From mathcomp Require Import kernel prob_lang lang_syntax_util lang_syntax. (******************************************************************************) (* Casino example *) @@ -22,6 +22,7 @@ Import numFieldTopology.Exports. Local Open Scope classical_set_scope. Local Open Scope ring_scope. (* Local Open Scope ereal_scope. *) +Local Open Scope string_scope. Section trunc_lemmas. Open Scope ring_scope. @@ -227,7 +228,7 @@ Proof. move=> my; rewrite letin'E/=. rewrite integral_uniform//= => _ /= Y mY /=. have /= := measurable_kernel u _ my measurableT _ mY. -move/measurable_ysection => /(_ R x) /=. +move/measurable_ysection => /(_ x) /=. set A := (X in measurable X). set B := (X in _ -> measurable X). suff : A = B by move=> ->. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 629c3aed2e..0fb8687ed3 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -565,7 +565,7 @@ apply: eq_fsbigr => r _; rewrite ge0_integralZl//. - by rewrite lee_fin invr_ge0// ltW. Qed. -Lemma integral_uniform (f : measurableTypeR R -> \bar R) (a b : R) +Lemma integral_uniform (f : _ -> \bar R) (a b : R) (ab0 : (0 < b - a)%R) : measurable_fun setT f -> (forall x, 0 <= f x) -> let m := uniform_probability ab0 in \int[m]_x f x = @@ -600,131 +600,155 @@ Qed. End integral_uniform. +(* definition of the beta probability specialized to natural numbers *) Section beta_probability. -Context {R : realType}. Local Open Scope ring_scope. -Variables (a b : nat). +Context {R : realType}. +Variables a b : nat. -Definition Beta t : R := - t^+(a-1) * (1-t)^+(b-1). +(* unnormalized pdf for beta specialized to nat *) +Definition ubeta_nat_pdf (t : R) := t ^+ a.-1 * (`1- t) ^+ b.-1. -Definition B : R := - (a-1)`!%:R * (b-1)`!%:R / (a+b-1)`!%:R. - (* \int[lebesgue_measure]_(t in `[0%R, 1%R]) Beta t. *) +Lemma ubeta_nat_pdf_ge0 t : 0 <= t <= 1 -> 0 <= ubeta_nat_pdf t. +Proof. by move=> /andP[t0 t1]; rewrite mulr_ge0// exprn_ge0// onem_ge0. Qed. -Axiom beta1 : \int[lebesgue_measure]_(x in `[0, 1]) (x^+(a-1) * (1-x)^+(b-1)) = B. +Lemma ubeta_nat_pdf_le1 t : 0 <= t <= 1 -> ubeta_nat_pdf t <= 1. +Proof. +move=> /andP[t0 t1]; rewrite /ubeta_nat_pdf. +by rewrite mulr_ile1// ?(exprn_ge0,onem_ge0,exprn_ile1,onem_le1). +Qed. -(* apply: fine_ge0. -apply: integral_ge0 => x x01. -have /andP[x0 x1] : (0 <= x <= 1) by apply: x01. -rewrite lee_fin. -apply: mulr_ge0; apply/exprn_ge0; lra. *) +Lemma measurable_ubeta_nat_pdf : measurable_fun setT ubeta_nat_pdf. +Proof. +by apply /measurable_funM => //; exact/measurable_fun_pow/measurable_funB. +Qed. -Set Printing All. +(* normalization constant *) +Definition beta_nat_norm : R := a.-1`!%:R * b.-1`!%:R / (a + b).-1`!%:R. -Lemma B_ge0 : 0 <= B. +Lemma beta_nat_normE : + beta_nat_norm = \int[lebesgue_measure]_(x in `[0, 1]) ubeta_nat_pdf x. Proof. -rewrite /B. Admitted. -(* Definition betaPDF (p : R) := - if 0 <= p <= 1 then Beta p / B else 0. +Lemma beta_nat_norm_gt0 : 0 < beta_nat_norm. +Proof. +by rewrite /beta_nat_norm divr_gt0// ?mulr_gt0 ?ltr0n ?fact_gt0. +Qed. + +Lemma beta_nat_norm_ge0 : 0 <= beta_nat_norm. +Proof. exact/ltW/beta_nat_norm_gt0. Qed. -Lemma measurable_beta : measurable_fun setT betaPDF. +(* normalized pdf for beta specialized to nat *) +Definition beta_nat_pdf t := ubeta_nat_pdf t / beta_nat_norm. + +Lemma measurable_beta_nat_pdf : measurable_fun setT beta_nat_pdf. +Proof. by apply: measurable_funM => //; exact: measurable_ubeta_nat_pdf. Qed. + +Lemma beta_nat_pdf_ge0 t : 0 <= t <= 1 -> 0 <= beta_nat_pdf t. Proof. -rewrite /betaPDF; apply: measurable_fun_if => //. - apply: (measurable_fun_bool true). - by rewrite (_ : _ @^-1` _ = `[0, 1]%classic). -apply: measurable_funM=> //; apply /measurable_funM => //. -apply/measurable_fun_pow/measurable_funB => //. -Qed. *) +move=> t01; rewrite /beta_nat_pdf divr_ge0//. + exact: ubeta_nat_pdf_ge0. +exact: beta_nat_norm_ge0. +Qed. -(* TODO: beta_ge0 a b : is_true (0 <= beta a b) (to remove %E) *) -Definition p01 : 0 < 1 - 0 :> R. -Proof. by []. Qed. +Local Notation mu := lebesgue_measure. -Definition prebeta U : \bar R := - (* \int[lebesgue_measure]_(x in U `&` `[0, 1]) (x^+(a-1) * (1-x)^+(b-1))%:E. *) - \int[uniform_probability p01]_(t in U) (t^+(a-1) * (1-t)^+(b-1))%:E. +(* unnormalized beta specialized to nat *) +Definition ubeta_nat (U : set (measurableTypeR R)) : \bar R := + \int[mu]_(x in U `&` `[0, 1](*NB: is this correct?*)) (ubeta_nat_pdf x)%:E. +(* TODO: define as \int[uniform_probability p01]_(t in U) (ubeta_nat_pdf t)%:E ? *) -Let prebeta0 : prebeta set0 = 0%:E. +Let ubeta_nat0 : ubeta_nat set0 = 0%:E. +Proof. by rewrite /ubeta_nat set0I integral_set0. Qed. + +Let ubeta_nat_ge0 U : (0 <= ubeta_nat U)%E. Proof. -(* by rewrite /prebeta integral_setI_indic// integral_set0. *) -by rewrite /prebeta integral_set0. +rewrite /ubeta_nat integral_ge0//= => x [Ux]. +by rewrite in_itv/= => x01; rewrite lee_fin ubeta_nat_pdf_ge0. Qed. -Let prebeta_ge0 U : (0 <= prebeta U)%E. +(* TODO: doable *) +Let ubeta_nat_sigma_additive : semi_sigma_additive ubeta_nat. Proof. -rewrite /prebeta. -rewrite integral_ge0// => x Ux. +move=> /= F mF tF mUF. +rewrite /ubeta_nat. +rewrite setI_bigcupl. +apply: cvg_toP. + apply: ereal_nondecreasing_is_cvgn => m n mn. + admit. +rewrite ge0_integral_bigcup//=. +admit. +admit. +admit. +admit. Admitted. -Let prebeta_sigma_additive : semi_sigma_additive prebeta. -Proof. move=> /= F mF tF mUF. -rewrite /prebeta. -Admitted. +HB.instance Definition _ := isMeasure.Build _ _ _ ubeta_nat + ubeta_nat0 ubeta_nat_ge0 ubeta_nat_sigma_additive. -HB.instance Definition _ := isMeasure.Build _ _ _ prebeta - prebeta0 prebeta_ge0 prebeta_sigma_additive. +Definition beta_nat (*: set [the measurableType (R.-ocitv.-measurable).-sigma of + salgebraType R.-ocitv.-measurable] -> \bar R*) := + @mscale _ _ _ (invr_nonneg (NngNum beta_nat_norm_ge0)) ubeta_nat. -(* HB.instance Definition _ := Measure.on prebeta. *) +HB.instance Definition _ := Measure.on beta_nat. -(* 1/B(a, b) * \int_U p^(a-1) * (1-p)^(b-1) dx = beta *) -Definition beta : set _ -> \bar R := - @mscale _ _ R (invr_nonneg (NngNum B_ge0)) prebeta. +(*Let beta_nat0 : beta_nat set0 = 0. +Proof. exact: measure0. Qed. - (* \int[lebesgue_measure]_(t in U) - @mscale _ _ R (t^+(a-1) - (* * (NngNum (onem_ge0 p1))%:num^+(b-1) *) - * - (invr_nonneg (NngNum beta_ge0))%:num)%:nng - (mrestr lebesgue_measure (measurable_itv `[0, 1])) U. *) +Let beta_nat_ge0 U : (0 <= beta_nat U)%E. +Proof. exact: measure_ge0. Qed. -HB.instance Definition _ := Measure.on beta. +Let beta_nat_sigma_additive : semi_sigma_additive beta_nat. +Proof. move=> /= F mF tF mUF; exact: measure_semi_sigma_additive. Qed. + +HB.instance Definition _ := isMeasure.Build _ _ _ beta_nat beta_nat0 beta_nat_ge0 beta_nat_sigma_additive.*) -Let beta_setT : beta [set: _] = 1%:E. +Let beta_nat_setT : beta_nat setT = 1%:E. Proof. -rewrite /beta /mscale/= /mrestr/=. +rewrite /beta_nat. +rewrite /=. +rewrite /mscale /=. +rewrite beta_nat_normE /ubeta_nat setTI. +(* TODO: doable *) Admitted. -HB.instance Definition _ := @Measure_isProbability.Build _ _ R - beta beta_setT. +HB.instance Definition _ := @Measure_isProbability.Build _ _ _ -Example __ : beta `[0, 1] = 1%:E. + beta_nat beta_nat_setT. + +Lemma beta_nat01 : beta_nat `[0, 1] = 1%:E. Proof. -(* rewrite /beta integral_mkcond. -rewrite integral_uniform//=. -rewrite oppr0 addr0 invr1 mul1e. *) +rewrite /beta_nat. +rewrite /mscale/= /beta_nat /ubeta_nat. +rewrite beta_nat_normE setIid. Admitted. -(* HB.instance Definition _ (p : {nonneg R}) (p1 : p%:num <= 1) - := Measure.on (beta p1). +End beta_probability. -Let beta_setT (p : {nonneg R}) (p1 : p%:num <= 1) - : beta p1 [set: _] = 1%:E. -Proof. -rewrite /beta /mscale/= /mrestr/=. -rewrite setTI lebesgue_measure_itv/= lte_fin. -rewrite ltr01 oppr0 adde0 mule1 /B /Beta. -(* rewrite -subr_gt0 -EFinD . -EFinM mulVf// gt_eqF// subr_gt0. *) -Admitted. +Section beta_probability11. +Local Open Scope ring_scope. +Context {R : realType}. -HB.instance Definition _ (p : {nonneg R}) (p1 : p%:num <= 1) := @Measure_isProbability.Build _ _ R - (beta p1) (beta_setT p1). *) +Lemma ubeta_nat_pdf11 : ubeta_nat_pdf 1 1 = @cst R _ 1. +Proof. by apply/funext => r; rewrite /ubeta_nat_pdf/= !expr0 mulr1. Qed. -(* Lemma __ : beta_probability 6 4 (p1S 2) `[0, 1] = 1%:E. -Proof. -rewrite /beta_probability/beta/= /mscale/=. -rewrite lebesgue_measure_itv. *) +Lemma beta_nat_norm11 : beta_nat_norm 1 1 = 1 :> R. +Proof. by rewrite /beta_nat_norm/= fact0 mulr1/= divff. Qed. -Let H01 : (0 < 1 - 0 :> R). -Proof. lra. Qed. +Let a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. -(* Lemma uniform_beta U : (\int[uniform_probability H01]_x - \int[mscale (NngNum (normr_ge0 (56 * x ^+ 5 * (1 - x) ^+ 3)%R)) (measure_dirac__canonical__measure_Measure tt R)]__ - bernoulli_trunc (1 - (1 - x) ^+ 3) U)%E = 1%:E. *) - -End beta_probability. +Lemma beta11_uniform U : measurable U -> + beta_nat 1 1 U = uniform_probability a01 U. +Proof. +move=> mU; rewrite /beta_nat /uniform_probability. +rewrite /mscale/= beta_nat_norm11 subr0 invr1 !mul1e. +rewrite /ubeta_nat /mrestr/=. +rewrite ubeta_nat_pdf11/= integral_cst/= ?mul1e//. +exact: measurableI. +Qed. + +End beta_probability11. Section mscore. Context d (T : measurableType d) (R : realType). From 9a5999c2e3909a146747fbb180b65dcd8fe5b4b3 Mon Sep 17 00:00:00 2001 From: AyumuSaito Date: Tue, 20 Feb 2024 17:31:59 +0900 Subject: [PATCH 16/48] progress in proving the casino exampl - rewrite casino01 - proved casino34' - casino34 (remain admit) - casino45 - add Normalize --- theories/lang_syntax.v | 15 +- theories/lang_syntax_examples_wip.v | 345 ++++++++++++++++------------ theories/prob_lang.v | 24 +- 3 files changed, 231 insertions(+), 153 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 6afbcce8bd..58605b51f2 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -371,7 +371,7 @@ Qed. Local Notation B := beta_nat_norm. -Definition Baa'bb'Bab : R := (B (a + a') (b + b')) / B a b. +Definition Baa'bb'Bab : R := (beta_nat_norm (a + a') (b + b')) / beta_nat_norm a b. Lemma Baa'bb'Bab_ge0 : 0 <= Baa'bb'Bab. Proof. by rewrite /Baa'bb'Bab divr_ge0// beta_nat_norm_ge0. Qed. @@ -418,8 +418,8 @@ transitivity ((\int[beta_nat a b]_y move=> H_ge0. case: sumbool_ler; last admit. move=> H_le1. - rewrite /=. - admit. + rewrite /=/bernoulli !measure_addE/= /mscale/=. + congr _%E. rewrite /beta_nat/=. transitivity ((((B a b)^-1)%:E * \int[ubeta_nat a b]_x bernoulli (ubeta_nat_pdf_le1' x) U)%E). apply: ge0_integral_mscale => //=. @@ -455,7 +455,7 @@ congr adde. admit. rewrite [in LHS]/mscale/= [in RHS]/mscale/=. suff: (\int[ubeta_nat a b]_x (`1-(ubeta_nat_pdf a' b' x))%:E = - (B a b - B (a + a') (b + b'))%:E :> \bar R)%E. + (B a b - B (a + a'.-1) (b + b'.-1))%:E :> \bar R)%E. admit. under eq_integral do rewrite EFinB/=. rewrite integralB_EFin//=; last 2 first. @@ -466,9 +466,14 @@ rewrite beta_nat_normE. rewrite {1}/ubeta_nat setTI. rewrite EFinB. congr (_ - _)%E. + (* rewrite fineK//. *) admit. rewrite integral_ubeta_nat//. rewrite beta_nat_normE /ubeta_nat_pdf. +under eq_integral => x _. + rewrite -mulrA -mulrCA !mulrA/= -exprD -mulrA [X in _ * X]mulrC -exprD. +over. +rewrite /=. admit. Admitted. @@ -887,7 +892,7 @@ Inductive evalD : forall g t, exp D g t -> measurable_cst _ | eval_beta g (a b : nat) : - (exp_beta a b : exp D g _) -D> (cst (beta_nat a b)) ; measurable_cst _ + (exp_beta a b : exp D g _) -D> cst (beta_nat a b) ; measurable_cst _ | eval_poisson g n (e : exp D g _) f mf : e -D> f ; mf -> diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index 8af7c301c3..8908c7ee41 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -56,9 +56,33 @@ case: (sumbool_ler 0 p%:num) => [{}p0/=|]. admit. Abort. - End trunc_lemmas. +Section beta_example. +Open Scope ring_scope. +Open Scope lang_scope. +Context (R : realType). + +Let p610 : ((6 / 10%:R)%:nng : {nonneg R})%:num <= 1. +Proof. admit. Admitted. + +Lemma beta_bernoulli : + @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in Sample {exp_bernoulli_trunc [#{"p"}]}] = + execP [Sample {@exp_bernoulli _ _ (6 / 10%:R)%:nng p610}]. +Proof. +rewrite execP_letin !execP_sample !execD_beta_nat !execD_bernoulli/=. +rewrite execD_bernoulli_trunc exp_var'E !(execD_var_erefl "p")/=. +apply: eq_sfkernel=> x U. +rewrite letin'E/=. +rewrite /beta_nat/mscale/=. +transitivity (bernoulli_trunc ((@beta_nat_norm R 7 4) / (@beta_nat_norm R 6 4)) U); last first. + (* congr (bernoulli_trunc _ _). *) + (* rewrite /beta_nat_norm/= factE/=; lra. *) +(* apply: beta_nat_bern_bern. *) +Admitted. + +End beta_example. + Section casino_example. Open Scope ring_scope. Open Scope lang_scope. @@ -82,42 +106,7 @@ rewrite !letin'E !iteE/= integral_dirac// ge0_integral_mscale//=. by rewrite normr0 mul0e !mule0 !adde0 !diracT !mul1e. Qed. -(* Definition ex : exp _ [::] _ := @exp_bernoulli R [::] (1 / 2)%:nng (p1S 1). -Example ex1 : projT1 (execD ex) tt = 1%:E. *) - -Definition casino0 : exp _ [::] Bool := - [let "p" := Sample {exp_uniform 0 1 a01} in - let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in - return {1}:N <= #{"a2"}]. - -Example e1 : @exp R _ [::] _ := [{1}:R + {2}:R * {2}:R ^+ {3%nat}]. - -Lemma exe1 : projT1 (execD e1) = (fun x => 17). -Proof. -rewrite /e1 (@execD_bin _ _ binop_add) (@execD_bin _ _ binop_mult)/=. -rewrite execD_pow/= !execD_real /=. -apply: funext => x /=. -lra. -Qed. - -(* Arguments exp_bin {R g b} &. *) -Definition casino1 : @exp R _ [::] _ := - [let "p" := Sample {exp_uniform 0 1 a01} in - let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - let "a2" := - Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]} in - return #{"a2"}]. - -Definition casino1' : @exp R _ [::] _ := - [let "p" := Sample {exp_uniform 0 1 a01} in - let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. - -Lemma binomial_le1' n p U : +Lemma binomial_over1 n p U : 0 <= p <= 1 -> (\int[binomial_probability_trunc n p]_y0 \d_(0 < y0)%N U = bernoulli_trunc (1 - `1-p ^+ n) U :> \bar R)%E. @@ -236,9 +225,6 @@ rewrite {}/A {}/B !setTI /ysection/= (*TODO: lemma?*) /preimage/=. by apply/seteqP; split => [z|z] /=; rewrite inE/=. Qed. -Let weak_head fl g {t1 t2} x (e : @exp R fl g t2) (xg : x \notin dom g) := - exp_weak fl [::] _ (x, t1) e xg. - Lemma execP_letin_uniform g t str (s0 s1 : exp P ((str, Real) :: g) t) : (forall (p : R) x U, 0 <= p <= 1 -> execP s0 (p, x) U = execP s1 (p, x) U) -> @@ -255,96 +241,27 @@ apply: s01. by rewrite inE in p01. Qed. -(* Lemma casino01 : execP casino0 = execP casino1. -Proof. -rewrite /casino0 /casino1. -rewrite !execP_letin !execP_sample execD_uniform !execD_binomial_trunc /=. -rewrite execP_if execP_score !execP_return !execD_bernoulli_trunc/=. -rewrite !execD_rel (@execD_bin _ _ binop_minus) !execD_real execD_pow/=. -rewrite !execD_nat execD_unit/=. -rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a1")/=. -rewrite !(execD_var_erefl "a2")/=. -do 3 congr letin'. -apply: eq_sfkernel => x U. -rewrite !letin'E/=. -apply: binomial_le1. - -rewrite /=. -Abort. *) - -(* Lemma casino01 : execP casino0 = execP casino1. -Proof. -rewrite /casino0 /casino1. -pose s0 := - [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in - {exp_return [{1}:N <= #{"a2"}]}]. -pose s1 := - [let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - let "a2" := Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%N}]} in - {exp_return [#{"a2"}]}]. -have := (@execP_letin_uniform [::] Bool "p" (s0 R (found "p" Real [::]) _ _) (s1 R (found "p" Real [::]) _ _)). -apply. -move=> p x U r01. -rewrite /s0/s1/=. -rewrite !execP_letin !execP_sample !execD_binomial_trunc /=. -rewrite execP_if execP_score !execP_return !execD_bernoulli_trunc/=. -rewrite !execD_rel (@execD_bin _ _ binop_minus) execD_pow. -rewrite (@execD_bin _ _ binop_minus) !execD_real/=. -rewrite !execD_nat execD_unit/=. -rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a1")/=. -rewrite !(execD_var_erefl "a2")/=. -rewrite !letin'E/=. -move: r01 => /andP[r0 r1]. -rewrite !integral_binomial_probabilty_trunc//=. -apply: eq_bigr => i _. -congr (_ * _)%E. -rewrite !letin'E iteE/=. -congr (\int[_]_y _)%E. -apply: funext => x0. -rewrite !letin'E/=. -by apply/binomial_le1/andP. -Qed. *) - -Lemma casino01' y V : measurable V -> execP casino0 y V = execP casino1' y V. +Lemma congr_letin g t1 t2 str (e : @exp _ _ _ t1) (e1 e2 : @exp _ _ (_ :: g) t2) x U : + (forall y V, execP e1 (y, x) V = execP e2 (y, x) V) -> + @execP R g t2 [let str := e in e1] x U = @execP R g t2 [let str := e in e2] x U. Proof. -move=> mV //. -rewrite /casino0 /casino1. -apply: execP_letin_uniform => //. -move=> p x U r01. -rewrite !execP_letin !execP_sample !execD_binomial_trunc /=. -rewrite execP_if execP_score !execP_return !execD_bernoulli_trunc/=. -rewrite !execD_rel (@execD_bin _ _ binop_minus) execD_pow. -rewrite (@execD_bin _ _ binop_minus) !execD_real/=. -rewrite !execD_nat execD_unit/=. -rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a1")/=. -rewrite !(execD_var_erefl "a2")/=. -rewrite !letin'E/=. -move: r01 => /andP[r0 r1]. -rewrite !integral_binomial_probabilty_trunc//=. -apply: eq_bigr => i _. -congr (_ * _)%E. -rewrite !letin'E iteE/=. -congr (\int[_]_y _)%E. -apply: funext => x0. -rewrite !letin'E/=. -by apply/binomial_le1'/andP. +move=> He. +rewrite !execP_letin !letin'E. +apply: eq_integral => ? _. +apply: He. Qed. -Lemma exec_casino t U : - execP casino0 t U = ((10 / 99)%:E * \d_true U + (1 / 99)%:E * \d_false U)%E . +Lemma congr_normalize g t (e1 e2 : @exp R _ g t) : + (forall x U, execP e1 x U = execP e2 x U) -> + execD [Normalize e1] = execD [Normalize e2]. Proof. -rewrite /casino0 !execP_letin !execP_sample execD_uniform/=. -rewrite !execD_binomial_trunc execP_if !execP_return !execP_score/=. -rewrite !execD_rel !execD_real execD_unit/=. -rewrite !exp_var'E (execD_var_erefl "p") (execD_var_erefl "a1"). -rewrite (execD_var_erefl "p") (execD_var_erefl "a2")/=. -rewrite letin'E/= /uniform_probability ge0_integral_mscale//=. -rewrite subr0 invr1 mul1e. -under eq_integral. -Admitted. +move=> He. +apply: eq_execD. +rewrite !execD_normalize_pt /=. +f_equal. +apply: eq_kernel => y V. +apply: He. +Qed. Definition uniform_syntax : @exp R _ [::] _ := [let "p" := Sample {exp_uniform 0 1 a01} in @@ -415,17 +332,72 @@ congr (_%:E * _)%E. lra. Qed. +Lemma exec_beta_a1 U : + @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in + Sample {exp_bernoulli_trunc [#{"p"}]}] tt U = + @execP R [::] _ [Sample {exp_bernoulli_trunc [{3 / 5}:R]}] tt U. +Proof. +rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. +rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. +transitivity (beta_nat_bern 6 4 1 0 tt U : \bar R). + rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. + apply: eq_integral => x _. + do 2 f_equal. + by rewrite (mul1r, mulr1). +rewrite beta_nat_bern_bern/= /bernoulli/= bernoulli_truncE; last by lra. +rewrite measure_addE/= /mscale/=. +congr (_ * _ + _ * _)%:E; rewrite /onem; rewrite /Baa'bb'Bab /beta_nat_norm/=; +by rewrite !factS/= fact0; field. +Qed. + +Definition casino0 : @exp R _ [::] _ := + [Normalize + let "p" := Sample {exp_uniform 0 1 a01} in + let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in + return {1}:N <= #{"a2"}]. + +Definition casino1 : @exp R _ [::] _ := + [Normalize + let "p" := Sample {exp_uniform 0 1 a01} in + let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. + +Lemma casino01 : + execD casino0 = execD casino1. +Proof. +rewrite /casino0 /casino1. +apply: eq_execD. +f_equal. +apply: congr_normalize => y V. +apply: execP_letin_uniform => //. +move=> p x U r01. +apply: congr_letin => y0 V0. +apply: congr_letin => y1 V1. +rewrite !execP_letin !execP_sample !execD_binomial_trunc /=. +rewrite !execP_return !execD_bernoulli_trunc/=. +rewrite !execD_rel (@execD_bin _ _ binop_minus) execD_pow. +rewrite (@execD_bin _ _ binop_minus) !execD_real/= !execD_nat. +rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a2")/=. +rewrite !letin'E/=. +move: r01 => /andP[r0 r1]. +by apply/binomial_over1/andP. +Qed. + Definition casino2 : @exp R _ [::] _ := - [let "p" := Sample {exp_uniform 0 1 a01} in + [Normalize + let "p" := Sample {exp_uniform 0 1 a01} in let "_" := Score {[{56}:R * #{"p"} ^+ {5%nat} * {[{1}:R - #{"p"}]} ^+ {3%nat}]} in Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. -Lemma casino12 y V : - measurable V -> - execP casino1' y V = execP casino2 y V. +Lemma casino12 : + execD casino1 = execD casino2. Proof. -move=> mV. -rewrite /casino1' /casino2. +apply: eq_execD. +f_equal. +apply: congr_normalize => y V. apply: execP_letin_uniform => //. move=> p x U /andP[p0 p1]. rewrite !execP_letin !execP_sample execP_if execD_rel/=. @@ -437,7 +409,6 @@ rewrite !execD_binomial_trunc/=. rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a1")/=. rewrite !letin'E/=. rewrite integral_binomial_probabilty_trunc//=. -(* set s := letin' _ _. *) rewrite (bigD1 (inord 5))//=. rewrite big1; last first. move=> [[|[|[|[|[|[|[|[|[|//]]]]]]]]]]//= Hi Hi5; rewrite letin'E iteE; @@ -456,23 +427,113 @@ by rewrite subr_ge0. Qed. Definition casino3 : @exp R _ [::] _ := - [let "_" := Score {1 / 9}:R in + [Normalize + let "_" := Score {1 / 9}:R in let "p" := Sample {exp_beta 6 4} in Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. +Lemma casino34' U : + @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in + Sample {exp_bernoulli_trunc [{[{1}:R - #{"p"}]} ^+ {3%nat}]}] tt U = + @execP R [::] _ [Sample {exp_bernoulli_trunc [{1 / 11}:R]}] tt U. +Proof. +rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. +rewrite execD_pow/= (@execD_bin _ _ binop_minus) !execD_real/=. +rewrite exp_var'E (execD_var_erefl "p")/=. +transitivity (beta_nat_bern 6 4 0 3 tt U : \bar R). + rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. + apply: eq_integral => x _. + do 2 f_equal. + by rewrite mul1r. +rewrite bernoulli_truncE; last by lra. +rewrite beta_nat_bern_bern/= /bernoulli/=. +rewrite measure_addE/= /mscale/=. +by congr (_ * _ + _ * _)%:E; rewrite /onem; +rewrite /Baa'bb'Bab /beta_nat_norm/=; rewrite !factS/= fact0; field. +Qed. + +Lemma bern_onem (f : _ -> R) U p : + (\int[beta_nat 6 4]_y bernoulli_trunc (f y) U = p%:E * \d_true U + `1-p%:E * \d_false U)%E -> + (\int[beta_nat 6 4]_y bernoulli_trunc (1 - f y) U = `1-p%:E * \d_true U + p%:E * \d_false U)%E. +Proof. +under eq_integral => x _. + rewrite bernoulli_truncE. + over. +admit. +rewrite /= /beta_nat/mscale/= /beta_nat_norm/= /ubeta_nat/ubeta_nat_pdf. +Admitted. + Definition casino4 : @exp R _ [::] _ := - [let "_" := Score {1 / 9}:R in + [Normalize + let "_" := Score {1 / 9}:R in Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. +Lemma casino34 : + execD casino3 = execD casino4. +Proof. +apply: eq_execD. +f_equal. +apply: congr_normalize => y V. +apply: congr_letin => x U. +rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. +rewrite (@execD_bin _ _ binop_minus) execD_pow/= (@execD_bin _ _ binop_minus). +rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. +transitivity (\int[beta_nat 6 4]_y bernoulli_trunc (1 - (1 - y) ^+ 3) U : \bar R)%E. + by rewrite /beta_nat_bern !letin'E/= /onem. +rewrite bernoulli_truncE; last by lra. +have -> := (@bern_onem (fun x => (1 - x) ^+ 3) U (1 / 11) _). + congr (_ * _ + _ * _)%E; congr _%:E; rewrite /onem; lra. +transitivity (beta_nat_bern 6 4 0 3 tt U : \bar R). + rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. + apply: eq_integral => y0 _. + do 2 f_equal. + by rewrite mul1r. +rewrite beta_nat_bern_bern/= /bernoulli/=. +rewrite measure_addE/= /mscale/=. +by congr (_ * _ + _ * _)%:E; rewrite /onem; +rewrite /Baa'bb'Bab /beta_nat_norm/=; rewrite !factS/= fact0; field. +Qed. + +Lemma norm_score_bern g p1 p2 (p10 : p1 != 0) (p1_ge0 : 0 <= p1) +(p201 : 0 <= p2 <= 1) : + @execD R g _ + [Normalize let "_" := Score {p1}:R in + Sample {exp_bernoulli_trunc [{p2}:R]}] = + execD [Normalize Sample {exp_bernoulli_trunc [{p2}:R]}]. +Proof. +apply: eq_execD. +rewrite !execD_normalize_pt/= !execP_letin !execP_score. +rewrite !execP_sample !execD_bernoulli_trunc !execD_real/=. +apply: funext=> x. +apply: eq_probability=> /= y. +rewrite /normalize_pt !normalizeE/=. +rewrite !bernoulli_truncE; last lra; last lra. +rewrite !diracT !mule1 /onem -EFinD addrCA subrr addr0. +rewrite !letin'E. +under eq_integral. + move=> x0 _ /=. + rewrite !bernoulli_truncE; last lra. + rewrite !diracT !mule1 /onem -EFinD addrCA subrr addr0. + over. +rewrite !ge0_integral_mscale//= ger0_norm//. +rewrite integral_dirac// diracT !mule1. +rewrite !ifF; last first. + rewrite eqe. + apply/negbTE/negP => /orP[/eqP|//]. + by rewrite /onem; lra. + rewrite eqe. + apply/negbTE/negP => /orP[/eqP|//]. + by rewrite /onem; lra. +rewrite !bernoulli_truncE; last lra. +rewrite integral_dirac//= diracT !diracE. +by rewrite muleC muleA -EFinM mulVf// invr1 /onem !(mul1r, mule1). +Qed. + Definition casino5 : @exp R _ [::] _ := [Normalize Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. -(* Lemma casino5_ans U : projT1 (execD casino5) tt U = - ((10 / 11)%:E * \d_true U + (1 / 11)%:E * \d_false U)%E. -Proof. -rewrite /casino5 execD_normalize_pt execP_sample execD_bernoulli_trunc/=. -rewrite execD_real/=. -rewrite normalizeE/= bernoulli_truncE. -rewrite /mnormalize/=. *) +Lemma casino45 : + execD casino4 = execD casino5. +Proof. by rewrite norm_score_bern//; lra. Qed. End casino_example. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 0fb8687ed3..d907b0c0f8 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -491,7 +491,7 @@ End binomial_example. Section uniform_probability. Context (R : realType) (a b : R) (ab0 : (0 < b - a)%R). -Definition uniform_probability (* : set _ -> \bar R *) := +Definition uniform_probability : set _ -> \bar R := @mscale _ _ R (invr_nonneg (NngNum (ltW ab0))) (mrestr lebesgue_measure (measurable_itv `[a, b])). @@ -607,7 +607,7 @@ Context {R : realType}. Variables a b : nat. (* unnormalized pdf for beta specialized to nat *) -Definition ubeta_nat_pdf (t : R) := t ^+ a.-1 * (`1- t) ^+ b.-1. +Definition ubeta_nat_pdf (t : R) := t ^+ a * (`1- t) ^+ b. Lemma ubeta_nat_pdf_ge0 t : 0 <= t <= 1 -> 0 <= ubeta_nat_pdf t. Proof. by move=> /andP[t0 t1]; rewrite mulr_ge0// exprn_ge0// onem_ge0. Qed. @@ -691,7 +691,19 @@ Definition beta_nat (*: set [the measurableType (R.-ocitv.-measurable).-sigma of salgebraType R.-ocitv.-measurable] -> \bar R*) := @mscale _ _ _ (invr_nonneg (NngNum beta_nat_norm_ge0)) ubeta_nat. -HB.instance Definition _ := Measure.on beta_nat. +Let beta_nat0 : beta_nat set0 = 0. +Proof. exact: measure0. Qed. + +Let beta_nat_ge0 U : (0 <= beta_nat U)%E. +Proof. exact: measure_ge0. Qed. + +Let beta_nat_sigma_additive : semi_sigma_additive beta_nat. +Proof. move=> /= F mF tF mUF; exact: measure_semi_sigma_additive. Qed. + +HB.instance Definition _ := isMeasure.Build _ _ _ beta_nat + beta_nat0 beta_nat_ge0 beta_nat_sigma_additive. + +(* HB.instance Definition _ := Measure.on beta_nat. *) (*Let beta_nat0 : beta_nat set0 = 0. Proof. exact: measure0. Qed. @@ -730,16 +742,16 @@ Section beta_probability11. Local Open Scope ring_scope. Context {R : realType}. -Lemma ubeta_nat_pdf11 : ubeta_nat_pdf 1 1 = @cst R _ 1. +Lemma ubeta_nat_pdf11 : ubeta_nat_pdf 0 0 = @cst R _ 1. Proof. by apply/funext => r; rewrite /ubeta_nat_pdf/= !expr0 mulr1. Qed. -Lemma beta_nat_norm11 : beta_nat_norm 1 1 = 1 :> R. +Lemma beta_nat_norm11 : beta_nat_norm 0 0 = 1 :> R. Proof. by rewrite /beta_nat_norm/= fact0 mulr1/= divff. Qed. Let a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. Lemma beta11_uniform U : measurable U -> - beta_nat 1 1 U = uniform_probability a01 U. + beta_nat 0 0 U = uniform_probability a01 U. Proof. move=> mU; rewrite /beta_nat /uniform_probability. rewrite /mscale/= beta_nat_norm11 subr0 invr1 !mul1e. From fbe2441e068294725f52088d41ee30041aff8325 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 27 Feb 2024 16:24:15 +0900 Subject: [PATCH 17/48] admit about fact and adjustment of exp but casino broken --- theories/lang_syntax.v | 181 +++++++++++++++++++++++++--- theories/lang_syntax_examples_wip.v | 55 ++++++--- theories/prob_lang.v | 70 ++++++----- 3 files changed, 244 insertions(+), 62 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 58605b51f2..5b3a26eeff 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -347,6 +347,69 @@ End accessor_functions. Arguments acc_typ {R} s n. Arguments measurable_acc_typ {R} s n. + +Lemma factD n m : (n`! * m`! <= (n + m).+1`!)%N. +Proof. +elim: n m => /= [m|n ih m]. + by rewrite fact0 mul1n add0n factS leq_pmull. +rewrite 2!factS [in X in (_ <= _ * X)%N]addSn -mulnA leq_mul//. +by rewrite ltnS addSnnS leq_addr. +Qed. + +Lemma factD' n m : (n`! * m.-1`! <= (n + m)`!)%N. +Proof. +case: m => //= [|m]. + by rewrite fact0 muln1 addn0. +by rewrite addnS factD. +Qed. + +Lemma leq_prod2 (x y n m : nat) : (n <= x)%N -> (m <= y)%N -> + (\prod_(m <= i < y) i * \prod_(n <= i < x) i <= \prod_(n + m <= i < x + y) i)%N. +Proof. +move=> nx my. +rewrite big_addn. +rewrite -addnBA//. +rewrite {3}/index_iota. +rewrite -addnBAC//. +rewrite iotaD. +rewrite big_cat/=. +rewrite mulnC. +rewrite leq_mul//. + rewrite /index_iota. + apply: leq_prod. + by move=> i _; rewrite leq_addr. +rewrite subnKC//. +rewrite -{1}(add0n m). +rewrite big_addn. +rewrite {2}(_ : (y - m) = ((y - m + x) - x))%N; last first. + by rewrite -addnBA// subnn addn0. +rewrite -{1}(add0n x). +rewrite big_addn. +rewrite -addnBA// subnn addn0. +apply: leq_prod => i _. +by rewrite leq_add2r leq_addr. +Qed. + +Lemma leq_fact2 (x y n m : nat) : + (n <= x) %N -> (m <= y)%N -> + (x`! * y`! * ((n + m).+1)`! <= n`! * m`! * ((x + y).+1)`!)%N. +Proof. +move=> nx my. +rewrite (_ : x`! = n`! * \prod_(n.+1 <= i < x.+1) i)%N; last first. + by rewrite -fact_split. +rewrite -!mulnA leq_mul2l; apply/orP; right. +rewrite (_ : y`! = m`! * \prod_(m.+1 <= i < y.+1) i)%N; last first. + by rewrite -fact_split. +rewrite mulnCA -!mulnA leq_mul2l; apply/orP; right. +rewrite (_ : (x + y).+1`! = (n + m).+1`! * \prod_((n + m).+2 <= i < (x + y).+2) i)%N; last first. + rewrite -fact_split//. + by rewrite ltnS leq_add. +rewrite mulnA mulnC leq_mul2l; apply/orP; right. +rewrite -addSn -addnS. +rewrite -addSn -addnS. +exact: leq_prod2. +Qed. + (* part of the step from casino3 to casino4 *) Section casino3_casino4. Context {R : realType}. @@ -359,7 +422,7 @@ Import Notations. Program Definition beta_nat_bern : R.-sfker munit ~> mbool := @letin' _ _ _ munit _ mbool _ - (sample_cst [the probability _ _ of (beta_nat a b)]) + (sample_cst [the probability _ _ of beta_nat a b]) (sample (bernoulli_trunc \o ubeta_nat_pdf a' b' \o @fst (measurableTypeR R) _ ) _). Next Obligation. apply: measurableT_comp. @@ -376,7 +439,7 @@ Definition Baa'bb'Bab : R := (beta_nat_norm (a + a') (b + b')) / beta_nat_norm a Lemma Baa'bb'Bab_ge0 : 0 <= Baa'bb'Bab. Proof. by rewrite /Baa'bb'Bab divr_ge0// beta_nat_norm_ge0. Qed. -Definition Baa'bb'Bab_nneg : {nonneg R} := NngNum Baa'bb'Bab_ge0. +Definition Baa'bb'Bab_nneg : {nonneg R} := NngNum Baa'bb'Bab_ge0. Lemma Baa'bb'Bab_le1 : Baa'bb'Bab_nneg%:num <= 1. Proof. @@ -386,22 +449,67 @@ rewrite /B /beta_nat_norm. rewrite ler_pdivrMr ?ltr0n ?fact_gt0//. rewrite mulrAC. rewrite ler_pdivlMr ?ltr0n ?fact_gt0//. -(* TODO: plausible *) -Admitted. +rewrite -!natrM ler_nat. +case: a. + rewrite /= fact0 mul1n !add0n. + case: b => /=. + case: a' => //. + case: b' => //= m. + by rewrite fact0 !mul1n muln1. + move=> n/=. + by rewrite fact0 add0n muln1 mul1n factD'. + move=> m. + rewrite mulnC leq_mul// mulnC. + by rewrite (leq_trans (factD' _ _))// addSn addnS//= addnC. +move=> n. +rewrite addSn. +case: b. + rewrite !fact0 add0n muln1 [leqRHS]mulnC addn0/= leq_mul//. + by rewrite factD'. +move=> m. +clear a b. +rewrite [(n + a').+1.-1]/=. +rewrite [n.+1.-1]/=. +rewrite [m.+1.-1]/=. +rewrite addnS. +rewrite [(_ + m).+1.-1]/=. +rewrite (addSn m b'). +rewrite [(m + _).+1.-1]/=. +rewrite (addSn (n + a')). +rewrite [_.+1.-1]/=. +rewrite addSn addnS. +by rewrite leq_fact2// leq_addr. +Qed. -Lemma onem_Baa'bb'Bab_ge0 : 0 <= B a b - (Baa'bb'Bab_nneg%:num). -Admitted. +Lemma onem_Baa'bb'Bab_ge0 : 0 <= 1 - (Baa'bb'Bab_nneg%:num). +Proof. by rewrite subr_ge0 Baa'bb'Bab_le1. Qed. + +Lemma onem_Baa'bb'Bab_ge0_fix : 0 <= B a b * (1 - Baa'bb'Bab_nneg%:num). +Proof. +rewrite mulr_ge0//. + rewrite /B. + exact: beta_nat_norm_ge0. +rewrite subr_ge0. +exact: Baa'bb'Bab_le1. +Qed. Lemma ubeta_nat_pdf_ge0' t : 0 <= ubeta_nat_pdf a' b' t :> R. +Proof. +apply: ubeta_nat_pdf_ge0. (* TODO: needs 0 <= t <= 1 *) Admitted. Lemma ubeta_nat_pdf_le1' t : (NngNum (ubeta_nat_pdf_ge0' t))%:num <= 1 :> R. +Proof. +rewrite /=. +rewrite /ubeta_nat_pdf. +rewrite /ubeta_nat_pdf'. (* TODO: needs 0 <= t <= 1 *) Admitted. Lemma integral_ubeta_nat : (\int[ubeta_nat a b]_x (ubeta_nat_pdf a' b' x)%:E = \int[mu]_(x in `[0%R, 1%R]) (x ^+ a'.-1 * `1-x ^+ b'.-1 * x ^+ a * `1-x ^+ b)%:E :> \bar R)%E. +Proof. Admitted. Lemma beta_nat_bern_bern U : @@ -414,32 +522,65 @@ transitivity ((\int[beta_nat a b]_y (@bernoulli R (NngNum (ubeta_nat_pdf_ge0' y)) (ubeta_nat_pdf_le1' y) U))%E). apply: eq_integral => /= r _. rewrite /bernoulli_trunc/=. - case: sumbool_ler; last admit. + case: sumbool_ler; last first. + admit. move=> H_ge0. - case: sumbool_ler; last admit. + case: sumbool_ler; last first. + admit. move=> H_le1. rewrite /=/bernoulli !measure_addE/= /mscale/=. - congr _%E. + by congr _%E. rewrite /beta_nat/=. -transitivity ((((B a b)^-1)%:E * \int[ubeta_nat a b]_x bernoulli (ubeta_nat_pdf_le1' x) U)%E). - apply: ge0_integral_mscale => //=. +transitivity (((B a b)^-1%:E * \int[ubeta_nat a b]_x bernoulli (ubeta_nat_pdf_le1' x) U)%E). + rewrite ge0_integral_mscale //=. admit. rewrite {2}/bernoulli. suff: (\int[ubeta_nat a b]_x bernoulli (ubeta_nat_pdf_le1' x) U)%E = measure_add (mscale (NngNum (beta_nat_norm_ge0 (a + a') (b + b'))) \d_true) - (mscale (NngNum (onem_Baa'bb'Bab_ge0)) \d_false) U. - admit. -transitivity (( + (mscale (NngNum (onem_Baa'bb'Bab_ge0_fix)) \d_false) U. + rewrite /= => ->. + rewrite !measure_addE/= /mscale/=. + rewrite /Baa'bb'Bab/=. + rewrite muleDr//. + congr (_ + _)%E. + rewrite -EFinM. + rewrite mulrA. + rewrite (mulrC (_^-1)). + by rewrite EFinM. + rewrite muleA. + congr (_ * _)%E. + rewrite -EFinM. + rewrite mulrA mulVr ?mul1r ?unitfE//. + rewrite gt_eqF//. + by rewrite beta_nat_norm_gt0. +transitivity ( \int[ubeta_nat a b]_x (mscale (NngNum (ubeta_nat_pdf_ge0' x)) \d_true U) + - \int[ubeta_nat a b]_x (mscale (onem_nonneg (ubeta_nat_pdf_le1' x)) \d_false) U))%E. - admit. + \int[ubeta_nat a b]_x (mscale (onem_nonneg (ubeta_nat_pdf_le1' x)) \d_false) U)%E. + rewrite /bernoulli/=. + under eq_integral do rewrite measure_addE/=. + rewrite ge0_integralD//=. + rewrite /mscale/=. + apply: emeasurable_funM => //. + by apply/EFin_measurable_fun; exact: measurable_ubeta_nat_pdf. + apply/EFin_measurable_fun => /=. + apply: measurable_funM => //. + apply: measurable_funB => //. + exact: measurable_ubeta_nat_pdf. rewrite measure_addE. congr adde. rewrite [in LHS]/mscale/= [in RHS]/mscale/=. suff: (\int[ubeta_nat a b]_x ((ubeta_nat_pdf a' b' x)%:E) = (B (a + a') (b + b'))%:E :> \bar R)%E. + move=> <-. + rewrite muleC. + rewrite -ge0_integralZl//=. + by under eq_integral do rewrite muleC. + apply/EFin_measurable_fun. + exact: measurable_ubeta_nat_pdf. + move=> x _. + rewrite lee_fin ubeta_nat_pdf_ge0//. admit. transitivity ( (* calculating distribution (13) *) \int[mu]_(x in `[(0:R)%R, (1:R)%R]%classic) @@ -449,13 +590,18 @@ congr adde. transitivity ( (\int[mu]_(x in `[0%R, 1%R]) ((x ^+ (a + a').-1 * `1-x ^+ (b + b').-1)%:E))%E : \bar R). + apply: eq_integral => /= t. + rewrite inE/= in_itv/= => /andP[t0 t1]. + congr EFin. admit. rewrite beta_nat_normE//. rewrite /ubeta_nat_pdf//. + rewrite /ubeta_nat_pdf'. admit. rewrite [in LHS]/mscale/= [in RHS]/mscale/=. suff: (\int[ubeta_nat a b]_x (`1-(ubeta_nat_pdf a' b' x))%:E = - (B a b - B (a + a'.-1) (b + b'.-1))%:E :> \bar R)%E. + (B a b - B ((a + a')) ((b + b')))%:E :> \bar R)%E. + rewrite /Baa'bb'Bab. admit. under eq_integral do rewrite EFinB/=. rewrite integralB_EFin//=; last 2 first. @@ -470,6 +616,7 @@ congr (_ - _)%E. admit. rewrite integral_ubeta_nat//. rewrite beta_nat_normE /ubeta_nat_pdf. +rewrite /ubeta_nat_pdf'. under eq_integral => x _. rewrite -mulrA -mulrCA !mulrA/= -exprD -mulrA [X in _ * X]mulrC -exprD. over. diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index 8908c7ee41..e15941f361 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -335,23 +335,29 @@ Qed. Lemma exec_beta_a1 U : @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in Sample {exp_bernoulli_trunc [#{"p"}]}] tt U = - @execP R [::] _ [Sample {exp_bernoulli_trunc [{3 / 5}:R]}] tt U. + @execP R [::] _ [Sample {exp_bernoulli_trunc [{7 / 55 (* TODO: 3 / 5 *)}:R]}] tt U. Proof. rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. -transitivity (beta_nat_bern 6 4 1 0 tt U : \bar R). +transitivity (beta_nat_bern 6 4 2 1 tt U : \bar R). rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. apply: eq_integral => x _. + rewrite /=. do 2 f_equal. - by rewrite (mul1r, mulr1). + by rewrite /ubeta_nat_pdf' expr1 expr0 mulr1. rewrite beta_nat_bern_bern/= /bernoulli/= bernoulli_truncE; last by lra. rewrite measure_addE/= /mscale/=. -congr (_ * _ + _ * _)%:E; rewrite /onem; rewrite /Baa'bb'Bab /beta_nat_norm/=; -by rewrite !factS/= fact0; field. +congr (_ * _ + _ * _)%:E. + rewrite /Baa'bb'Bab /beta_nat_norm/=. + rewrite !factS/= fact0. + by field. +rewrite /onem; rewrite /Baa'bb'Bab /beta_nat_norm/=; +rewrite !factS/= fact0. +by field. Qed. Definition casino0 : @exp R _ [::] _ := - [Normalize + [Normalize let "p" := Sample {exp_uniform 0 1 a01} in let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in @@ -435,24 +441,29 @@ Definition casino3 : @exp R _ [::] _ := Lemma casino34' U : @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in Sample {exp_bernoulli_trunc [{[{1}:R - #{"p"}]} ^+ {3%nat}]}] tt U = - @execP R [::] _ [Sample {exp_bernoulli_trunc [{1 / 11}:R]}] tt U. + @execP R [::] _ [Sample {exp_bernoulli_trunc [{3 / 143 (* TODO: 1 / 11*)}:R]}] tt U. Proof. rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. rewrite execD_pow/= (@execD_bin _ _ binop_minus) !execD_real/=. rewrite exp_var'E (execD_var_erefl "p")/=. -transitivity (beta_nat_bern 6 4 0 3 tt U : \bar R). +transitivity (beta_nat_bern 6 4 1 4 tt U : \bar R). rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. apply: eq_integral => x _. do 2 f_equal. + rewrite /ubeta_nat_pdf'. + rewrite expr0. by rewrite mul1r. rewrite bernoulli_truncE; last by lra. rewrite beta_nat_bern_bern/= /bernoulli/=. rewrite measure_addE/= /mscale/=. -by congr (_ * _ + _ * _)%:E; rewrite /onem; -rewrite /Baa'bb'Bab /beta_nat_norm/=; rewrite !factS/= fact0; field. +congr (_ * _ + _ * _)%:E; rewrite /onem. + rewrite /Baa'bb'Bab /beta_nat_norm/=. + by rewrite !factS/= fact0; field. +rewrite /Baa'bb'Bab /beta_nat_norm/=. +by rewrite !factS/= fact0; field. Qed. -Lemma bern_onem (f : _ -> R) U p : +Lemma bern_onem (f : _ -> R) U p : (\int[beta_nat 6 4]_y bernoulli_trunc (f y) U = p%:E * \d_true U + `1-p%:E * \d_false U)%E -> (\int[beta_nat 6 4]_y bernoulli_trunc (1 - f y) U = `1-p%:E * \d_true U + p%:E * \d_false U)%E. Proof. @@ -466,7 +477,7 @@ Admitted. Definition casino4 : @exp R _ [::] _ := [Normalize let "_" := Score {1 / 9}:R in - Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. + Sample {exp_bernoulli_trunc [{140 / 143(*TODO: 10 / 11*)}:R]}]. Lemma casino34 : execD casino3 = execD casino4. @@ -481,17 +492,22 @@ rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. transitivity (\int[beta_nat 6 4]_y bernoulli_trunc (1 - (1 - y) ^+ 3) U : \bar R)%E. by rewrite /beta_nat_bern !letin'E/= /onem. rewrite bernoulli_truncE; last by lra. -have -> := (@bern_onem (fun x => (1 - x) ^+ 3) U (1 / 11) _). +have -> := (@bern_onem (fun x => (1 - x) ^+ 3) U (3 / 143) (*TODO: (1 / 11)*) _). congr (_ * _ + _ * _)%E; congr _%:E; rewrite /onem; lra. -transitivity (beta_nat_bern 6 4 0 3 tt U : \bar R). +transitivity (beta_nat_bern 6 4 1 4 tt U : \bar R). rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. apply: eq_integral => y0 _. do 2 f_equal. + rewrite /ubeta_nat_pdf'/=. + rewrite expr0. by rewrite mul1r. rewrite beta_nat_bern_bern/= /bernoulli/=. rewrite measure_addE/= /mscale/=. -by congr (_ * _ + _ * _)%:E; rewrite /onem; -rewrite /Baa'bb'Bab /beta_nat_norm/=; rewrite !factS/= fact0; field. +congr (_ * _ + _ * _)%:E; rewrite /onem. + rewrite /Baa'bb'Bab /beta_nat_norm/=. + by rewrite !factS/= fact0; field. +rewrite /Baa'bb'Bab /beta_nat_norm/=. +by rewrite !factS/= fact0; field. Qed. Lemma norm_score_bern g p1 p2 (p10 : p1 != 0) (p1_ge0 : 0 <= p1) @@ -530,10 +546,13 @@ by rewrite muleC muleA -EFinM mulVf// invr1 /onem !(mul1r, mule1). Qed. Definition casino5 : @exp R _ [::] _ := - [Normalize Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. + [Normalize Sample {exp_bernoulli_trunc [{(*TODO: 10 / 11*) 140 /143}:R]}]. Lemma casino45 : execD casino4 = execD casino5. -Proof. by rewrite norm_score_bern//; lra. Qed. +Proof. +rewrite norm_score_bern//. +lra. +Qed. End casino_example. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index d907b0c0f8..ae3ecb733e 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -600,14 +600,41 @@ Qed. End integral_uniform. +(* normalization constant *) +Section beta_nat_norm. +Context {R : realType} (a b : nat). + +Definition beta_nat_norm : R := + (a.-1`!%:R * b.-1`!%:R / (a + b).-1`!%:R)%R. + +Lemma beta_nat_norm_gt0 : (0 < beta_nat_norm :> R)%R. +Proof. +by rewrite /beta_nat_norm divr_gt0// ?mulr_gt0 ?ltr0n ?fact_gt0. +Qed. + +Lemma beta_nat_norm_ge0 : (0 <= beta_nat_norm :> R)%R. +Proof. exact/ltW/beta_nat_norm_gt0. Qed. + +End beta_nat_norm. + +Lemma beta_nat_norm00 {R : realType} : beta_nat_norm 0 0 = 1%R :> R. +Proof. by rewrite /beta_nat_norm/= fact0 mulr1/= divff. Qed. + +Lemma beta_nat_norm11 {R : realType} : beta_nat_norm 1 1 = 1%R :> R. +Proof. by rewrite /beta_nat_norm/= fact0 mulr1/= divff. Qed. + (* definition of the beta probability specialized to natural numbers *) + +Definition ubeta_nat_pdf' {R : realType} (a b : nat) (t : R) := + (t ^+ a * (`1- t) ^+ b)%R. + Section beta_probability. Local Open Scope ring_scope. Context {R : realType}. Variables a b : nat. - (* unnormalized pdf for beta specialized to nat *) -Definition ubeta_nat_pdf (t : R) := t ^+ a * (`1- t) ^+ b. + +Definition ubeta_nat_pdf (t : R) := ubeta_nat_pdf' a.-1 b.-1 t. Lemma ubeta_nat_pdf_ge0 t : 0 <= t <= 1 -> 0 <= ubeta_nat_pdf t. Proof. by move=> /andP[t0 t1]; rewrite mulr_ge0// exprn_ge0// onem_ge0. Qed. @@ -623,24 +650,15 @@ Proof. by apply /measurable_funM => //; exact/measurable_fun_pow/measurable_funB. Qed. -(* normalization constant *) -Definition beta_nat_norm : R := a.-1`!%:R * b.-1`!%:R / (a + b).-1`!%:R. - Lemma beta_nat_normE : - beta_nat_norm = \int[lebesgue_measure]_(x in `[0, 1]) ubeta_nat_pdf x. + beta_nat_norm a b = \int[lebesgue_measure]_(x in `[0, 1]) ubeta_nat_pdf x. Proof. +rewrite /beta_nat_norm. +rewrite /ubeta_nat_pdf /ubeta_nat_pdf'. Admitted. -Lemma beta_nat_norm_gt0 : 0 < beta_nat_norm. -Proof. -by rewrite /beta_nat_norm divr_gt0// ?mulr_gt0 ?ltr0n ?fact_gt0. -Qed. - -Lemma beta_nat_norm_ge0 : 0 <= beta_nat_norm. -Proof. exact/ltW/beta_nat_norm_gt0. Qed. - (* normalized pdf for beta specialized to nat *) -Definition beta_nat_pdf t := ubeta_nat_pdf t / beta_nat_norm. +Definition beta_nat_pdf t := ubeta_nat_pdf t / (beta_nat_norm a b). Lemma measurable_beta_nat_pdf : measurable_fun setT beta_nat_pdf. Proof. by apply: measurable_funM => //; exact: measurable_ubeta_nat_pdf. Qed. @@ -678,10 +696,12 @@ apply: cvg_toP. apply: ereal_nondecreasing_is_cvgn => m n mn. admit. rewrite ge0_integral_bigcup//=. -admit. -admit. -admit. -admit. +- by move=> k; exact: measurableI. +- admit. +- move=> /= r [i _ [Fir]] /=. + rewrite in_itv/= => r01. + by rewrite lee_fin ubeta_nat_pdf_ge0. +- exact: trivIset_setIr. Admitted. HB.instance Definition _ := isMeasure.Build _ _ _ ubeta_nat @@ -689,7 +709,7 @@ HB.instance Definition _ := isMeasure.Build _ _ _ ubeta_nat Definition beta_nat (*: set [the measurableType (R.-ocitv.-measurable).-sigma of salgebraType R.-ocitv.-measurable] -> \bar R*) := - @mscale _ _ _ (invr_nonneg (NngNum beta_nat_norm_ge0)) ubeta_nat. + @mscale _ _ _ (invr_nonneg (NngNum (beta_nat_norm_ge0 a b))) ubeta_nat. Let beta_nat0 : beta_nat set0 = 0. Proof. exact: measure0. Qed. @@ -726,7 +746,6 @@ rewrite beta_nat_normE /ubeta_nat setTI. Admitted. HB.instance Definition _ := @Measure_isProbability.Build _ _ _ - beta_nat beta_nat_setT. Lemma beta_nat01 : beta_nat `[0, 1] = 1%:E. @@ -742,16 +761,13 @@ Section beta_probability11. Local Open Scope ring_scope. Context {R : realType}. -Lemma ubeta_nat_pdf11 : ubeta_nat_pdf 0 0 = @cst R _ 1. -Proof. by apply/funext => r; rewrite /ubeta_nat_pdf/= !expr0 mulr1. Qed. - -Lemma beta_nat_norm11 : beta_nat_norm 0 0 = 1 :> R. -Proof. by rewrite /beta_nat_norm/= fact0 mulr1/= divff. Qed. +Lemma ubeta_nat_pdf11 : ubeta_nat_pdf 1 1 = @cst R _ 1. +Proof. by apply/funext => r; rewrite /ubeta_nat_pdf/= /ubeta_nat_pdf' !expr0 mulr1. Qed. Let a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. Lemma beta11_uniform U : measurable U -> - beta_nat 0 0 U = uniform_probability a01 U. + beta_nat 1 1 U = uniform_probability a01 U. Proof. move=> mU; rewrite /beta_nat /uniform_probability. rewrite /mscale/= beta_nat_norm11 subr0 invr1 !mul1e. From 1f325b8ed3f2914f72d8a281a2cdee6912abb2a0 Mon Sep 17 00:00:00 2001 From: AyumuSaito Date: Tue, 27 Feb 2024 20:43:09 +0900 Subject: [PATCH 18/48] fix --- theories/lang_syntax.v | 11 ++++++----- theories/lang_syntax_examples_wip.v | 16 ++++++++-------- 2 files changed, 14 insertions(+), 13 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 5b3a26eeff..79084e18e3 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -423,7 +423,7 @@ Import Notations. Program Definition beta_nat_bern : R.-sfker munit ~> mbool := @letin' _ _ _ munit _ mbool _ (sample_cst [the probability _ _ of beta_nat a b]) - (sample (bernoulli_trunc \o ubeta_nat_pdf a' b' \o @fst (measurableTypeR R) _ ) _). + (sample (bernoulli_trunc \o ubeta_nat_pdf a'.+1 b'.+1 \o @fst (measurableTypeR R) _ ) _). Next Obligation. apply: measurableT_comp. apply: measurableT_comp. @@ -493,7 +493,7 @@ rewrite subr_ge0. exact: Baa'bb'Bab_le1. Qed. -Lemma ubeta_nat_pdf_ge0' t : 0 <= ubeta_nat_pdf a' b' t :> R. +Lemma ubeta_nat_pdf_ge0' t : 0 <= ubeta_nat_pdf a'.+1 b'.+1 t :> R. Proof. apply: ubeta_nat_pdf_ge0. (* TODO: needs 0 <= t <= 1 *) Admitted. @@ -506,10 +506,11 @@ rewrite /ubeta_nat_pdf'. (* TODO: needs 0 <= t <= 1 *) Admitted. Lemma integral_ubeta_nat : - (\int[ubeta_nat a b]_x (ubeta_nat_pdf a' b' x)%:E = + (\int[ubeta_nat a b]_x (ubeta_nat_pdf a'.+1 b'.+1 x)%:E = \int[mu]_(x in `[0%R, 1%R]) (x ^+ a'.-1 * `1-x ^+ b'.-1 * x ^+ a * `1-x ^+ b)%:E :> \bar R)%E. Proof. +rewrite /ubeta_nat/ubeta_nat_pdf. Admitted. Lemma beta_nat_bern_bern U : @@ -571,7 +572,7 @@ transitivity ( rewrite measure_addE. congr adde. rewrite [in LHS]/mscale/= [in RHS]/mscale/=. - suff: (\int[ubeta_nat a b]_x ((ubeta_nat_pdf a' b' x)%:E) + suff: (\int[ubeta_nat a b]_x ((ubeta_nat_pdf a'.+1 b'.+1 x)%:E) = (B (a + a') (b + b'))%:E :> \bar R)%E. move=> <-. rewrite muleC. @@ -599,7 +600,7 @@ congr adde. rewrite /ubeta_nat_pdf'. admit. rewrite [in LHS]/mscale/= [in RHS]/mscale/=. -suff: (\int[ubeta_nat a b]_x (`1-(ubeta_nat_pdf a' b' x))%:E = +suff: (\int[ubeta_nat a b]_x (`1-(ubeta_nat_pdf a'.+1 b'.+1 x))%:E = (B a b - B ((a + a')) ((b + b')))%:E :> \bar R)%E. rewrite /Baa'bb'Bab. admit. diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index e15941f361..3a1ef935d1 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -335,11 +335,11 @@ Qed. Lemma exec_beta_a1 U : @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in Sample {exp_bernoulli_trunc [#{"p"}]}] tt U = - @execP R [::] _ [Sample {exp_bernoulli_trunc [{7 / 55 (* TODO: 3 / 5 *)}:R]}] tt U. + @execP R [::] _ [Sample {exp_bernoulli_trunc [{3 / 5}:R]}] tt U. Proof. rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. -transitivity (beta_nat_bern 6 4 2 1 tt U : \bar R). +transitivity (beta_nat_bern 6 4 1 0 tt U : \bar R). rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. apply: eq_integral => x _. rewrite /=. @@ -441,12 +441,12 @@ Definition casino3 : @exp R _ [::] _ := Lemma casino34' U : @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in Sample {exp_bernoulli_trunc [{[{1}:R - #{"p"}]} ^+ {3%nat}]}] tt U = - @execP R [::] _ [Sample {exp_bernoulli_trunc [{3 / 143 (* TODO: 1 / 11*)}:R]}] tt U. + @execP R [::] _ [Sample {exp_bernoulli_trunc [{1 / 11}:R]}] tt U. Proof. rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. rewrite execD_pow/= (@execD_bin _ _ binop_minus) !execD_real/=. rewrite exp_var'E (execD_var_erefl "p")/=. -transitivity (beta_nat_bern 6 4 1 4 tt U : \bar R). +transitivity (beta_nat_bern 6 4 0 3 tt U : \bar R). rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. apply: eq_integral => x _. do 2 f_equal. @@ -477,7 +477,7 @@ Admitted. Definition casino4 : @exp R _ [::] _ := [Normalize let "_" := Score {1 / 9}:R in - Sample {exp_bernoulli_trunc [{140 / 143(*TODO: 10 / 11*)}:R]}]. + Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. Lemma casino34 : execD casino3 = execD casino4. @@ -492,9 +492,9 @@ rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. transitivity (\int[beta_nat 6 4]_y bernoulli_trunc (1 - (1 - y) ^+ 3) U : \bar R)%E. by rewrite /beta_nat_bern !letin'E/= /onem. rewrite bernoulli_truncE; last by lra. -have -> := (@bern_onem (fun x => (1 - x) ^+ 3) U (3 / 143) (*TODO: (1 / 11)*) _). +have -> := (@bern_onem (fun x => (1 - x) ^+ 3) U (1 / 11) _). congr (_ * _ + _ * _)%E; congr _%:E; rewrite /onem; lra. -transitivity (beta_nat_bern 6 4 1 4 tt U : \bar R). +transitivity (beta_nat_bern 6 4 0 3 tt U : \bar R). rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. apply: eq_integral => y0 _. do 2 f_equal. @@ -546,7 +546,7 @@ by rewrite muleC muleA -EFinM mulVf// invr1 /onem !(mul1r, mule1). Qed. Definition casino5 : @exp R _ [::] _ := - [Normalize Sample {exp_bernoulli_trunc [{(*TODO: 10 / 11*) 140 /143}:R]}]. + [Normalize Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. Lemma casino45 : execD casino4 = execD casino5. From 2882daf8a1cdd08eb54e528b60bdd84023bfd290 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Fri, 1 Mar 2024 16:46:29 +0900 Subject: [PATCH 19/48] minor progress --- theories/lang_syntax.v | 12 +++------- theories/prob_lang.v | 54 +++++++++++++++++++++++++----------------- 2 files changed, 35 insertions(+), 31 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 79084e18e3..4ba97bd7cd 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -582,7 +582,7 @@ congr adde. exact: measurable_ubeta_nat_pdf. move=> x _. rewrite lee_fin ubeta_nat_pdf_ge0//. - admit. + admit. (* 0 <= x <= 1 *) transitivity ( (* calculating distribution (13) *) \int[mu]_(x in `[(0:R)%R, (1:R)%R]%classic) ((x ^+ a'.-1 * (`1- x) ^+ b'.-1 * x ^+ a * (`1- x) ^+ b)%:E) @@ -595,10 +595,7 @@ congr adde. rewrite inE/= in_itv/= => /andP[t0 t1]. congr EFin. admit. - rewrite beta_nat_normE//. - rewrite /ubeta_nat_pdf//. - rewrite /ubeta_nat_pdf'. - admit. + by rewrite beta_nat_normE. rewrite [in LHS]/mscale/= [in RHS]/mscale/=. suff: (\int[ubeta_nat a b]_x (`1-(ubeta_nat_pdf a'.+1 b'.+1 x))%:E = (B a b - B ((a + a')) ((b + b')))%:E :> \bar R)%E. @@ -609,12 +606,10 @@ rewrite integralB_EFin//=; last 2 first. admit. admit. rewrite integral_cst//= mul1e. -rewrite beta_nat_normE. rewrite {1}/ubeta_nat setTI. rewrite EFinB. congr (_ - _)%E. - (* rewrite fineK//. *) - admit. + by rewrite -beta_nat_normE. rewrite integral_ubeta_nat//. rewrite beta_nat_normE /ubeta_nat_pdf. rewrite /ubeta_nat_pdf'. @@ -622,7 +617,6 @@ under eq_integral => x _. rewrite -mulrA -mulrCA !mulrA/= -exprD -mulrA [X in _ * X]mulrC -exprD. over. rewrite /=. -admit. Admitted. End casino3_casino4. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index ae3ecb733e..5fbd431703 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -651,10 +651,11 @@ by apply /measurable_funM => //; exact/measurable_fun_pow/measurable_funB. Qed. Lemma beta_nat_normE : - beta_nat_norm a b = \int[lebesgue_measure]_(x in `[0, 1]) ubeta_nat_pdf x. + (beta_nat_norm a b)%:E = (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (ubeta_nat_pdf x)%:E)%E. Proof. rewrite /beta_nat_norm. -rewrite /ubeta_nat_pdf /ubeta_nat_pdf'. +rewrite /ubeta_nat_pdf. +rewrite /ubeta_nat_pdf'. Admitted. (* normalized pdf for beta specialized to nat *) @@ -686,23 +687,34 @@ rewrite /ubeta_nat integral_ge0//= => x [Ux]. by rewrite in_itv/= => x01; rewrite lee_fin ubeta_nat_pdf_ge0. Qed. -(* TODO: doable *) +(* TODO: should be shorter *) Let ubeta_nat_sigma_additive : semi_sigma_additive ubeta_nat. Proof. -move=> /= F mF tF mUF. -rewrite /ubeta_nat. -rewrite setI_bigcupl. -apply: cvg_toP. +move=> /= F mF tF mUF; rewrite /ubeta_nat setI_bigcupl; apply: cvg_toP. apply: ereal_nondecreasing_is_cvgn => m n mn. - admit. + apply: lee_sum_nneg_natr => // k _ _. + apply: integral_ge0 => /= x [_]; rewrite in_itv => x01. + by rewrite lee_fin; exact: ubeta_nat_pdf_ge0. rewrite ge0_integral_bigcup//=. - by move=> k; exact: measurableI. -- admit. -- move=> /= r [i _ [Fir]] /=. - rewrite in_itv/= => r01. +- apply/integrableP; split. + by apply/EFin_measurable_fun; exact: measurable_funTS measurable_ubeta_nat_pdf. + apply: le_lt_trans => /=. + apply: (@subset_integral _ _ _ mu _ `[0%R, 1%R]) => //=. + - rewrite -setI_bigcupl; apply: measurableI => //. + - apply/measurableT_comp => //; apply/measurableT_comp => //. + exact: measurable_funTS measurable_ubeta_nat_pdf. + - by apply: bigcup_sub => k _; exact: subIsetr. + rewrite /=. + under eq_integral. + move=> /= x; rewrite inE/= in_itv/= => x01. + rewrite ger0_norm//; last by rewrite ubeta_nat_pdf_ge0. + over. + by rewrite -beta_nat_normE ltry. +- move=> x [k _ [_]]; rewrite /= in_itv/= => x01. by rewrite lee_fin ubeta_nat_pdf_ge0. - exact: trivIset_setIr. -Admitted. +Qed. HB.instance Definition _ := isMeasure.Build _ _ _ ubeta_nat ubeta_nat0 ubeta_nat_ge0 ubeta_nat_sigma_additive. @@ -738,22 +750,20 @@ HB.instance Definition _ := isMeasure.Build _ _ _ beta_nat beta_nat0 beta_nat_ge Let beta_nat_setT : beta_nat setT = 1%:E. Proof. -rewrite /beta_nat. -rewrite /=. -rewrite /mscale /=. -rewrite beta_nat_normE /ubeta_nat setTI. -(* TODO: doable *) -Admitted. +rewrite /beta_nat /= /mscale /=. +rewrite /ubeta_nat/= setTI. +by rewrite -beta_nat_normE -EFinM mulVr// unitfE gt_eqF// beta_nat_norm_gt0. +Qed. HB.instance Definition _ := @Measure_isProbability.Build _ _ _ beta_nat beta_nat_setT. Lemma beta_nat01 : beta_nat `[0, 1] = 1%:E. Proof. -rewrite /beta_nat. -rewrite /mscale/= /beta_nat /ubeta_nat. -rewrite beta_nat_normE setIid. -Admitted. +rewrite /beta_nat /= /mscale/=. +rewrite /beta_nat /ubeta_nat setIidr//. +by rewrite -beta_nat_normE -EFinM mulVr// unitfE gt_eqF// beta_nat_norm_gt0. +Qed. End beta_probability. From 4547a3c5431d3497b4e27a9213f6f3facf73ad56 Mon Sep 17 00:00:00 2001 From: saito ayumu Date: Wed, 6 Mar 2024 18:13:46 +0900 Subject: [PATCH 20/48] casino example (wip) - use change of variables from charge.v to prove lemma integral_beta_nat (wip) - beta_nat_bern_bern --- theories/lang_syntax.v | 274 ---------------- theories/lang_syntax_examples_wip.v | 348 ++++++++++++++------- theories/prob_lang.v | 463 ++++++++++++++++++++++++---- 3 files changed, 652 insertions(+), 433 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 4ba97bd7cd..aaf541d2f0 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -347,280 +347,6 @@ End accessor_functions. Arguments acc_typ {R} s n. Arguments measurable_acc_typ {R} s n. - -Lemma factD n m : (n`! * m`! <= (n + m).+1`!)%N. -Proof. -elim: n m => /= [m|n ih m]. - by rewrite fact0 mul1n add0n factS leq_pmull. -rewrite 2!factS [in X in (_ <= _ * X)%N]addSn -mulnA leq_mul//. -by rewrite ltnS addSnnS leq_addr. -Qed. - -Lemma factD' n m : (n`! * m.-1`! <= (n + m)`!)%N. -Proof. -case: m => //= [|m]. - by rewrite fact0 muln1 addn0. -by rewrite addnS factD. -Qed. - -Lemma leq_prod2 (x y n m : nat) : (n <= x)%N -> (m <= y)%N -> - (\prod_(m <= i < y) i * \prod_(n <= i < x) i <= \prod_(n + m <= i < x + y) i)%N. -Proof. -move=> nx my. -rewrite big_addn. -rewrite -addnBA//. -rewrite {3}/index_iota. -rewrite -addnBAC//. -rewrite iotaD. -rewrite big_cat/=. -rewrite mulnC. -rewrite leq_mul//. - rewrite /index_iota. - apply: leq_prod. - by move=> i _; rewrite leq_addr. -rewrite subnKC//. -rewrite -{1}(add0n m). -rewrite big_addn. -rewrite {2}(_ : (y - m) = ((y - m + x) - x))%N; last first. - by rewrite -addnBA// subnn addn0. -rewrite -{1}(add0n x). -rewrite big_addn. -rewrite -addnBA// subnn addn0. -apply: leq_prod => i _. -by rewrite leq_add2r leq_addr. -Qed. - -Lemma leq_fact2 (x y n m : nat) : - (n <= x) %N -> (m <= y)%N -> - (x`! * y`! * ((n + m).+1)`! <= n`! * m`! * ((x + y).+1)`!)%N. -Proof. -move=> nx my. -rewrite (_ : x`! = n`! * \prod_(n.+1 <= i < x.+1) i)%N; last first. - by rewrite -fact_split. -rewrite -!mulnA leq_mul2l; apply/orP; right. -rewrite (_ : y`! = m`! * \prod_(m.+1 <= i < y.+1) i)%N; last first. - by rewrite -fact_split. -rewrite mulnCA -!mulnA leq_mul2l; apply/orP; right. -rewrite (_ : (x + y).+1`! = (n + m).+1`! * \prod_((n + m).+2 <= i < (x + y).+2) i)%N; last first. - rewrite -fact_split//. - by rewrite ltnS leq_add. -rewrite mulnA mulnC leq_mul2l; apply/orP; right. -rewrite -addSn -addnS. -rewrite -addSn -addnS. -exact: leq_prod2. -Qed. - -(* part of the step from casino3 to casino4 *) -Section casino3_casino4. -Context {R : realType}. -Variables a b a' b' : nat. - -Local Notation mu := lebesgue_measure. - -Open Scope ring_scope. -Import Notations. - -Program Definition beta_nat_bern : R.-sfker munit ~> mbool := - @letin' _ _ _ munit _ mbool _ - (sample_cst [the probability _ _ of beta_nat a b]) - (sample (bernoulli_trunc \o ubeta_nat_pdf a'.+1 b'.+1 \o @fst (measurableTypeR R) _ ) _). -Next Obligation. -apply: measurableT_comp. - apply: measurableT_comp. - exact: measurable_bernoulli_trunc. - exact: measurable_ubeta_nat_pdf. -exact: (measurable_acc_typ [:: Real] 0). -Qed. - -Local Notation B := beta_nat_norm. - -Definition Baa'bb'Bab : R := (beta_nat_norm (a + a') (b + b')) / beta_nat_norm a b. - -Lemma Baa'bb'Bab_ge0 : 0 <= Baa'bb'Bab. -Proof. by rewrite /Baa'bb'Bab divr_ge0// beta_nat_norm_ge0. Qed. - -Definition Baa'bb'Bab_nneg : {nonneg R} := NngNum Baa'bb'Bab_ge0. - -Lemma Baa'bb'Bab_le1 : Baa'bb'Bab_nneg%:num <= 1. -Proof. -rewrite /Baa'bb'Bab_nneg/= /Baa'bb'Bab. -rewrite ler_pdivrMr// ?mul1r ?beta_nat_norm_gt0//. -rewrite /B /beta_nat_norm. -rewrite ler_pdivrMr ?ltr0n ?fact_gt0//. -rewrite mulrAC. -rewrite ler_pdivlMr ?ltr0n ?fact_gt0//. -rewrite -!natrM ler_nat. -case: a. - rewrite /= fact0 mul1n !add0n. - case: b => /=. - case: a' => //. - case: b' => //= m. - by rewrite fact0 !mul1n muln1. - move=> n/=. - by rewrite fact0 add0n muln1 mul1n factD'. - move=> m. - rewrite mulnC leq_mul// mulnC. - by rewrite (leq_trans (factD' _ _))// addSn addnS//= addnC. -move=> n. -rewrite addSn. -case: b. - rewrite !fact0 add0n muln1 [leqRHS]mulnC addn0/= leq_mul//. - by rewrite factD'. -move=> m. -clear a b. -rewrite [(n + a').+1.-1]/=. -rewrite [n.+1.-1]/=. -rewrite [m.+1.-1]/=. -rewrite addnS. -rewrite [(_ + m).+1.-1]/=. -rewrite (addSn m b'). -rewrite [(m + _).+1.-1]/=. -rewrite (addSn (n + a')). -rewrite [_.+1.-1]/=. -rewrite addSn addnS. -by rewrite leq_fact2// leq_addr. -Qed. - -Lemma onem_Baa'bb'Bab_ge0 : 0 <= 1 - (Baa'bb'Bab_nneg%:num). -Proof. by rewrite subr_ge0 Baa'bb'Bab_le1. Qed. - -Lemma onem_Baa'bb'Bab_ge0_fix : 0 <= B a b * (1 - Baa'bb'Bab_nneg%:num). -Proof. -rewrite mulr_ge0//. - rewrite /B. - exact: beta_nat_norm_ge0. -rewrite subr_ge0. -exact: Baa'bb'Bab_le1. -Qed. - -Lemma ubeta_nat_pdf_ge0' t : 0 <= ubeta_nat_pdf a'.+1 b'.+1 t :> R. -Proof. -apply: ubeta_nat_pdf_ge0. (* TODO: needs 0 <= t <= 1 *) -Admitted. - -Lemma ubeta_nat_pdf_le1' t : (NngNum (ubeta_nat_pdf_ge0' t))%:num <= 1 :> R. -Proof. -rewrite /=. -rewrite /ubeta_nat_pdf. -rewrite /ubeta_nat_pdf'. (* TODO: needs 0 <= t <= 1 *) -Admitted. - -Lemma integral_ubeta_nat : - (\int[ubeta_nat a b]_x (ubeta_nat_pdf a'.+1 b'.+1 x)%:E = - \int[mu]_(x in `[0%R, 1%R]) - (x ^+ a'.-1 * `1-x ^+ b'.-1 * x ^+ a * `1-x ^+ b)%:E :> \bar R)%E. -Proof. -rewrite /ubeta_nat/ubeta_nat_pdf. -Admitted. - -Lemma beta_nat_bern_bern U : - beta_nat_bern tt U = - sample_cst (bernoulli Baa'bb'Bab_le1) tt U. -Proof. -rewrite /beta_nat_bern. -rewrite [LHS]letin'E/=. -transitivity ((\int[beta_nat a b]_y - (@bernoulli R (NngNum (ubeta_nat_pdf_ge0' y)) (ubeta_nat_pdf_le1' y) U))%E). - apply: eq_integral => /= r _. - rewrite /bernoulli_trunc/=. - case: sumbool_ler; last first. - admit. - move=> H_ge0. - case: sumbool_ler; last first. - admit. - move=> H_le1. - rewrite /=/bernoulli !measure_addE/= /mscale/=. - by congr _%E. -rewrite /beta_nat/=. -transitivity (((B a b)^-1%:E * \int[ubeta_nat a b]_x bernoulli (ubeta_nat_pdf_le1' x) U)%E). - rewrite ge0_integral_mscale //=. - admit. -rewrite {2}/bernoulli. -suff: (\int[ubeta_nat a b]_x bernoulli (ubeta_nat_pdf_le1' x) U)%E = - measure_add - (mscale (NngNum (beta_nat_norm_ge0 (a + a') (b + b'))) \d_true) - (mscale (NngNum (onem_Baa'bb'Bab_ge0_fix)) \d_false) U. - rewrite /= => ->. - rewrite !measure_addE/= /mscale/=. - rewrite /Baa'bb'Bab/=. - rewrite muleDr//. - congr (_ + _)%E. - rewrite -EFinM. - rewrite mulrA. - rewrite (mulrC (_^-1)). - by rewrite EFinM. - rewrite muleA. - congr (_ * _)%E. - rewrite -EFinM. - rewrite mulrA mulVr ?mul1r ?unitfE//. - rewrite gt_eqF//. - by rewrite beta_nat_norm_gt0. -transitivity ( - \int[ubeta_nat a b]_x (mscale (NngNum (ubeta_nat_pdf_ge0' x)) \d_true U) - + - \int[ubeta_nat a b]_x (mscale (onem_nonneg (ubeta_nat_pdf_le1' x)) \d_false) U)%E. - rewrite /bernoulli/=. - under eq_integral do rewrite measure_addE/=. - rewrite ge0_integralD//=. - rewrite /mscale/=. - apply: emeasurable_funM => //. - by apply/EFin_measurable_fun; exact: measurable_ubeta_nat_pdf. - apply/EFin_measurable_fun => /=. - apply: measurable_funM => //. - apply: measurable_funB => //. - exact: measurable_ubeta_nat_pdf. -rewrite measure_addE. -congr adde. - rewrite [in LHS]/mscale/= [in RHS]/mscale/=. - suff: (\int[ubeta_nat a b]_x ((ubeta_nat_pdf a'.+1 b'.+1 x)%:E) - = (B (a + a') (b + b'))%:E :> \bar R)%E. - move=> <-. - rewrite muleC. - rewrite -ge0_integralZl//=. - by under eq_integral do rewrite muleC. - apply/EFin_measurable_fun. - exact: measurable_ubeta_nat_pdf. - move=> x _. - rewrite lee_fin ubeta_nat_pdf_ge0//. - admit. (* 0 <= x <= 1 *) - transitivity ( (* calculating distribution (13) *) - \int[mu]_(x in `[(0:R)%R, (1:R)%R]%classic) - ((x ^+ a'.-1 * (`1- x) ^+ b'.-1 * x ^+ a * (`1- x) ^+ b)%:E) - )%E. - by rewrite integral_ubeta_nat. - transitivity ( - (\int[mu]_(x in `[0%R, 1%R]) - ((x ^+ (a + a').-1 * `1-x ^+ (b + b').-1)%:E))%E : \bar R). - apply: eq_integral => /= t. - rewrite inE/= in_itv/= => /andP[t0 t1]. - congr EFin. - admit. - by rewrite beta_nat_normE. -rewrite [in LHS]/mscale/= [in RHS]/mscale/=. -suff: (\int[ubeta_nat a b]_x (`1-(ubeta_nat_pdf a'.+1 b'.+1 x))%:E = - (B a b - B ((a + a')) ((b + b')))%:E :> \bar R)%E. - rewrite /Baa'bb'Bab. - admit. -under eq_integral do rewrite EFinB/=. -rewrite integralB_EFin//=; last 2 first. - admit. - admit. -rewrite integral_cst//= mul1e. -rewrite {1}/ubeta_nat setTI. -rewrite EFinB. -congr (_ - _)%E. - by rewrite -beta_nat_normE. -rewrite integral_ubeta_nat//. -rewrite beta_nat_normE /ubeta_nat_pdf. -rewrite /ubeta_nat_pdf'. -under eq_integral => x _. - rewrite -mulrA -mulrCA !mulrA/= -exprD -mulrA [X in _ * X]mulrC -exprD. -over. -rewrite /=. -Admitted. - -End casino3_casino4. - Section context. Variables (R : realType). Definition ctx := seq (string * typ). diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v index 3a1ef935d1..eabf4b4fab 100644 --- a/theories/lang_syntax_examples_wip.v +++ b/theories/lang_syntax_examples_wip.v @@ -24,61 +24,86 @@ Local Open Scope ring_scope. (* Local Open Scope ereal_scope. *) Local Open Scope string_scope. -Section trunc_lemmas. -Open Scope ring_scope. -Open Scope lang_scope. -Context (R : realType). - -Lemma bernoulli_truncE (p : R) U : - (0 <= p <= 1)%R -> - (bernoulli_trunc p U = - p%:E * \d_true U + (`1-p)%:E * \d_false U)%E. -Proof. -move=> /andP[p0 p1]. -rewrite /bernoulli_trunc. -case: (sumbool_ler 0 p) => [{}p0/=|]. - case: (sumbool_ler p 1) => [{}p1/=|]. - by rewrite /bernoulli/= measure_addE. - by rewrite ltNge p1. -by rewrite ltNge p0. -Qed. - -Lemma __ p p1 : - @execD R [::] _ (exp_bernoulli p p1) = - execD (exp_bernoulli_trunc [{p%:num}:R]). -Proof. -apply: eq_execD. -rewrite execD_bernoulli execD_bernoulli_trunc execD_real/=. -apply: funext=> x /=. -rewrite /bernoulli_trunc. -case: (sumbool_ler 0 p%:num) => [{}p0/=|]. - case: (sumbool_ler p%:num 1) => [{}p1'/=|]. - admit. -Abort. - -End trunc_lemmas. - Section beta_example. Open Scope ring_scope. Open Scope lang_scope. -Context (R : realType). - -Let p610 : ((6 / 10%:R)%:nng : {nonneg R})%:num <= 1. -Proof. admit. Admitted. +Context {R : realType}. -Lemma beta_bernoulli : - @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in Sample {exp_bernoulli_trunc [#{"p"}]}] = - execP [Sample {@exp_bernoulli _ _ (6 / 10%:R)%:nng p610}]. +Lemma beta_bern610 : + @execP R [::] _ + [let "p" := + Sample {exp_beta 6 4} in Sample {exp_bernoulli_trunc [#{"p"}]}] = + execP [Sample {exp_bernoulli_trunc [{6 / 10}:R]}]. Proof. -rewrite execP_letin !execP_sample !execD_beta_nat !execD_bernoulli/=. -rewrite execD_bernoulli_trunc exp_var'E !(execD_var_erefl "p")/=. +rewrite execP_letin !execP_sample !execD_beta_nat !execD_bernoulli_trunc/=. +rewrite execD_real exp_var'E !(execD_var_erefl "p")/=. apply: eq_sfkernel=> x U. rewrite letin'E/=. -rewrite /beta_nat/mscale/=. -transitivity (bernoulli_trunc ((@beta_nat_norm R 7 4) / (@beta_nat_norm R 6 4)) U); last first. - (* congr (bernoulli_trunc _ _). *) - (* rewrite /beta_nat_norm/= factE/=; lra. *) -(* apply: beta_nat_bern_bern. *) +transitivity ((\int[beta_nat 6 4]_(y in `[0%R, 1%R]) bernoulli_trunc y U)%E : \bar R). + rewrite [in RHS]integral_mkcond /=. + apply: eq_integral => y _. + rewrite patchE. + case: ifPn => //. + (* Unset Printing Notations. *) + simpl in *. + rewrite mem_setE /= in_itv /= negb_and -!ltNge => /orP[y0|y1]; + rewrite /bernoulli_trunc/=. + case: sumbool_ler. + move=> a. + by rewrite ltNge a in y0. + rewrite /prob_lang.bernoulli0 /bernoulli => _. + rewrite [LHS]measure_addE/= /mscale/=. + (* match default value *) + admit. + admit. + rewrite integral_beta_nat. +under eq_integral => y y01. + rewrite bernoulli_truncE. + rewrite muleC muleDr// !muleA muleC [X in _ + X]muleC. + over. + admit. +rewrite /= bernoulli_truncE. +rewrite integralD//. +congr (_ + _)%E. +rewrite integralZl//. +rewrite muleC. +congr (_ * _)%E. +(* rewrite /= setTI. *) +rewrite /beta_nat_pdf. +rewrite /ubeta_nat_pdf/ubeta_nat_pdf'/=. +transitivity (\int[@lebesgue_measure R]_(x0 in `[0%R, 1%R]) + ((ubeta_nat_pdf 7 4 x0 / beta_nat_norm 6 4)%:E))%E. + apply: eq_integral => p _. + rewrite muleC -EFinM. + rewrite -[X in X * _]divr1 mulf_div. + congr (_ / _)%:E; last by rewrite mul1r. + by rewrite /ubeta_nat_pdf/= /ubeta_nat_pdf' mulrA -exprS. +under eq_integral do rewrite mulrC EFinM. +rewrite integralZl//=. +rewrite -beta_nat_normE /beta_nat_norm/= !factE. +rewrite -!EFinM/=. +congr _%:E; lra. + admit. + admit. +rewrite integralZl//. +rewrite muleC. +congr (_ * _)%E. +rewrite /beta_nat_pdf. +rewrite /ubeta_nat_pdf/ubeta_nat_pdf'/=. +transitivity (\int[@lebesgue_measure R]_(x0 in `[0%R, 1%R]) + ((ubeta_nat_pdf 6 5 x0 / beta_nat_norm 6 4)%:E))%E. + apply: eq_integral => p _. + by rewrite mulrC -EFinM -2!mulrA -exprSr mulrC. +under eq_integral do rewrite mulrC EFinM. +rewrite integralZl//=. +rewrite -beta_nat_normE /beta_nat_norm/= !factE. +rewrite -!EFinM/onem/=. +congr _%:E; lra. + admit. + admit. + admit. + admit. + lra. Admitted. End beta_example. @@ -241,6 +266,22 @@ apply: s01. by rewrite inE in p01. Qed. +(* Lemma measurable_mtyp (t : typ) (U : set (@mtyp R t)) : measurable U. +Proof. +induction t => //. *) + +Lemma congr_letinl g t1 t2 str (e1 e2 : @exp _ _ g t1) +(e : @exp _ _ (_ :: g) t2) x U : + (forall y V, execP e1 y V = execP e2 y V) -> + measurable U -> + @execP R g t2 [let str := e1 in e] x U = + @execP R g t2 [let str := e2 in e] x U. +Proof. +move=> He mU. +apply eq_sfkernel in He. +by rewrite !execP_letin He. +Qed. + Lemma congr_letin g t1 t2 str (e : @exp _ _ _ t1) (e1 e2 : @exp _ _ (_ :: g) t2) x U : (forall y V, execP e1 (y, x) V = execP e2 (y, x) V) -> @execP R g t2 [let str := e in e1] x U = @execP R g t2 [let str := e in e2] x U. @@ -339,14 +380,17 @@ Lemma exec_beta_a1 U : Proof. rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. -transitivity (beta_nat_bern 6 4 1 0 tt U : \bar R). +transitivity (beta_nat_bern 6 4 1 0 U : \bar R). rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. apply: eq_integral => x _. rewrite /=. do 2 f_equal. by rewrite /ubeta_nat_pdf' expr1 expr0 mulr1. -rewrite beta_nat_bern_bern/= /bernoulli/= bernoulli_truncE; last by lra. -rewrite measure_addE/= /mscale/=. +rewrite beta_nat_bern_bern// !bernoulli_truncE; last 2 first. + by lra. + apply/andP; split. + by apply/Baa'bb'Bab_ge0. + by apply/Baa'bb'Bab_le1. congr (_ * _ + _ * _)%:E. rewrite /Baa'bb'Bab /beta_nat_norm/=. rewrite !factS/= fact0. @@ -357,19 +401,47 @@ by field. Qed. Definition casino0 : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_uniform 0 1 a01} in - let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in - return {1}:N <= #{"a2"}]. + [Normalize + let "p" := Sample {exp_uniform 0 1 a01} in + let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in + return {1}:N <= #{"a2"}]. Definition casino1 : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_uniform 0 1 a01} in - let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. + [Normalize + let "p" := Sample {exp_uniform 0 1 a01} in + let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in + let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in + Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. + +Definition casino2 : @exp R _ [::] _ := + [Normalize + let "p" := Sample {exp_uniform 0 1 a01} in + let "_" := + Score {[{56}:R * #{"p"} ^+ {5%nat} * {[{1}:R - #{"p"}]} ^+ {3%nat}]} in + Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. + +Definition casino2' : @exp R _ [::] _ := + [Normalize + let "p" := Sample {exp_beta 1 1} in + let "_" := Score + {[{56}:R * #{"p"} ^+ {5%nat} * {[{1}:R - #{"p"}]} ^+ {3%nat}]} in + Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. + +Definition casino3 : @exp R _ [::] _ := + [Normalize + let "_" := Score {1 / 9}:R in + let "p" := Sample {exp_beta 6 4} in + Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. + +Definition casino4 : @exp R _ [::] _ := + [Normalize + let "_" := Score {1 / 9}:R in + Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. + +Definition casino5 : @exp R _ [::] _ := + [Normalize Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. Lemma casino01 : execD casino0 = execD casino1. @@ -392,12 +464,6 @@ move: r01 => /andP[r0 r1]. by apply/binomial_over1/andP. Qed. -Definition casino2 : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_uniform 0 1 a01} in - let "_" := Score {[{56}:R * #{"p"} ^+ {5%nat} * {[{1}:R - #{"p"}]} ^+ {3%nat}]} in - Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. - Lemma casino12 : execD casino1 = execD casino2. Proof. @@ -432,11 +498,88 @@ apply/exprn_ge0. by rewrite subr_ge0. Qed. -Definition casino3 : @exp R _ [::] _ := - [Normalize - let "_" := Score {1 / 9}:R in - let "p" := Sample {exp_beta 6 4} in - Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. +Lemma casino22' : + execD casino2 = execD casino2'. +Proof. +apply: eq_execD. +f_equal. +apply: congr_normalize => //= x U. +apply: congr_letinl => //= y V. +rewrite !execP_sample execD_uniform execD_beta_nat/=. +rewrite beta11_uniform//. +rewrite /=. +apply: sub_sigma_algebra. +rewrite //. +admit. +Admitted. + +Lemma casino23 : + execD casino2' = execD casino3. +Proof. +apply: eq_execD. +f_equal. +apply: congr_normalize => x U. +rewrite !execP_letin !execP_sample !execP_score !execD_beta_nat. +rewrite !execD_bernoulli_trunc/= !(@execD_bin _ _ binop_mult). +rewrite !execD_pow !(@execD_bin _ _ binop_minus) !execD_real/=. +rewrite !execD_pow !(@execD_bin _ _ binop_minus) !execD_real/=. +rewrite !exp_var'E !(execD_var_erefl "p")/=. +rewrite !letin'E/= ![in RHS]ge0_integral_mscale//=. +rewrite /=. +under eq_integral => y _. + rewrite letin'E/= /mscale/=. + over. +rewrite /=. +(* set f := letin' _ _. *) +(* oops... +transitivity (\int[beta_nat 1 1]_(y in `[0%R, 1%R]) f (y, x) U)%E. + rewrite [in RHS]integral_mkcond /=. + apply: eq_integral => y _. + rewrite patchE. + case: ifPn => //. + simpl in *. + rewrite mem_setE /= in_itv /= negb_and -!ltNge => /orP[y0|y1]; + rewrite /bernoulli_trunc/=. + case: sumbool_ler. + move=> a. + by rewrite ltNge a in y0. + rewrite /prob_lang.bernoulli0 /bernoulli => _. + rewrite [LHS]measure_addE/= /mscale/=. + (* match default value *) + admit. + admit. + rewrite integral_beta_nat. + admit. +rewrite (integral_beta_nat 6 4). + rewrite ger0_norm// integral_dirac// diracT mul1e letin'E/=. + transitivity (((1 / 9)%:E * \int[beta_nat 6 4]_(y in `[0%R, 1%R]) + bernoulli_trunc (1 - (1 - y) ^+ 3) U)%E : \bar R); last first. + admit. +rewrite (integral_beta_nat 6 4)//=. +rewrite -integralZl//=. + apply: eq_integral => y y01. + rewrite /f letin'E /= !ge0_integral_mscale//= integral_dirac// diracT mul1e. + rewrite -muleAC muleC -[in RHS]muleCA. + congr (_ * _)%E. + rewrite ger0_norm. + rewrite /beta_nat_pdf ubeta_nat_pdf11/= muleC -!EFinM. + rewrite !div1r. + rewrite /beta_nat_norm/= /ubeta_nat_pdf/ubeta_nat_pdf' factE/=/onem. + congr _%:E; lra. + rewrite inE/= in_itv/= in y01. + move: y01 => /andP[y0 y1]. + apply/mulr_ge0/exprn_ge0 => //. + apply/mulr_ge0/exprn_ge0 => //. + lra. +admit. +admit. +admit. +apply: (@measurableT_comp _ _ _ _ _ _ (fun x0 => f x0 U)). + apply: measurable_kernel. + done. +apply: measurable_pair2. +admit.*) +Admitted. Lemma casino34' U : @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in @@ -446,39 +589,34 @@ Proof. rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. rewrite execD_pow/= (@execD_bin _ _ binop_minus) !execD_real/=. rewrite exp_var'E (execD_var_erefl "p")/=. -transitivity (beta_nat_bern 6 4 0 3 tt U : \bar R). - rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. - apply: eq_integral => x _. - do 2 f_equal. - rewrite /ubeta_nat_pdf'. - rewrite expr0. - by rewrite mul1r. -rewrite bernoulli_truncE; last by lra. -rewrite beta_nat_bern_bern/= /bernoulli/=. -rewrite measure_addE/= /mscale/=. -congr (_ * _ + _ * _)%:E; rewrite /onem. - rewrite /Baa'bb'Bab /beta_nat_norm/=. - by rewrite !factS/= fact0; field. -rewrite /Baa'bb'Bab /beta_nat_norm/=. -by rewrite !factS/= fact0; field. +(* TODO: generalize *) +rewrite letin'E/=. +have := (@beta_nat_bern_bern R 6 4 0 3 U). +rewrite /beta_nat_bern /ubeta_nat_pdf/ubeta_nat_pdf'/=. +under eq_integral do rewrite expr0 mul1r. +move=> ->//. +rewrite /Baa'bb'Bab addn0 /beta_nat_norm/= !factE/= !factE. +by congr (bernoulli_trunc _ _); field. Qed. Lemma bern_onem (f : _ -> R) U p : - (\int[beta_nat 6 4]_y bernoulli_trunc (f y) U = p%:E * \d_true U + `1-p%:E * \d_false U)%E -> - (\int[beta_nat 6 4]_y bernoulli_trunc (1 - f y) U = `1-p%:E * \d_true U + p%:E * \d_false U)%E. + (forall x, 0 <= f x <= 1) -> + (\int[beta_nat 6 4]_y bernoulli_trunc (f y) U = + p%:E * \d_true U + `1-p%:E * \d_false U)%E -> + (\int[beta_nat 6 4]_y bernoulli_trunc (1 - f y) U + = `1-p%:E * \d_true U + p%:E * \d_false U)%E. Proof. +move=> f01. under eq_integral => x _. rewrite bernoulli_truncE. over. -admit. -rewrite /= /beta_nat/mscale/= /beta_nat_norm/= /ubeta_nat/ubeta_nat_pdf. +done. +move=> h1. +rewrite /= in h1. +rewrite /bernoulli_trunc. +(* /beta_nat/mscale/= /beta_nat_norm/= /ubeta_nat/ubeta_nat_pdf/=. *) Admitted. -Definition casino4 : @exp R _ [::] _ := - [Normalize - let "_" := Score {1 / 9}:R in - Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. - Lemma casino34 : execD casino3 = execD casino4. Proof. @@ -494,21 +632,24 @@ transitivity (\int[beta_nat 6 4]_y bernoulli_trunc (1 - (1 - y) ^+ 3) U : \bar R rewrite bernoulli_truncE; last by lra. have -> := (@bern_onem (fun x => (1 - x) ^+ 3) U (1 / 11) _). congr (_ * _ + _ * _)%E; congr _%:E; rewrite /onem; lra. -transitivity (beta_nat_bern 6 4 0 3 tt U : \bar R). - rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. + admit. +transitivity (beta_nat_bern 6 4 0 3 U : \bar R). + rewrite /beta_nat_bern /ubeta_nat_pdf/= /onem. apply: eq_integral => y0 _. do 2 f_equal. rewrite /ubeta_nat_pdf'/=. rewrite expr0. by rewrite mul1r. -rewrite beta_nat_bern_bern/= /bernoulli/=. -rewrite measure_addE/= /mscale/=. +rewrite beta_nat_bern_bern//= bernoulli_truncE; last first. + apply/andP; split. + apply/Baa'bb'Bab_ge0. + apply/Baa'bb'Bab_le1. congr (_ * _ + _ * _)%:E; rewrite /onem. rewrite /Baa'bb'Bab /beta_nat_norm/=. by rewrite !factS/= fact0; field. rewrite /Baa'bb'Bab /beta_nat_norm/=. by rewrite !factS/= fact0; field. -Qed. +Admitted. Lemma norm_score_bern g p1 p2 (p10 : p1 != 0) (p1_ge0 : 0 <= p1) (p201 : 0 <= p2 <= 1) : @@ -545,9 +686,6 @@ rewrite integral_dirac//= diracT !diracE. by rewrite muleC muleA -EFinM mulVf// invr1 /onem !(mul1r, mule1). Qed. -Definition casino5 : @exp R _ [::] _ := - [Normalize Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. - Lemma casino45 : execD casino4 = execD casino5. Proof. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 5fbd431703..54aa27cf6f 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -6,14 +6,21 @@ From mathcomp Require Import mathcomp_extra boolp classical_sets functions. From mathcomp Require Import cardinality fsbigop reals ereal. From mathcomp Require Import topology normedtype sequences. From mathcomp Require Import esum measure lebesgue_measure numfun. -From mathcomp Require Import lebesgue_integral exp kernel. +From mathcomp Require Import lebesgue_integral exp kernel charge. From mathcomp Require Import ring lra. -(******************************************************************************) -(* Semantics of a probabilistic programming language using s-finite kernels *) -(* *) -(* bernoulli r1 == Bernoulli probability with r1 a proof that *) -(* r : {nonneg R} is smaller than 1 *) +(**md**************************************************************************) +(* # Semantics of a probabilistic programming language using s-finite kernels *) +(* bernoulli p1 == Bernoulli probability with p1 a proof that *) +(* p : {nonneg R} is smaller than 1 *) +(* bernoulli_trunc r == Bernoulli probability with real number r *) +(* bino_term n k p == $\binom{n}{k}p^k (1-p)^(n-k)$ *) +(* Computes a binomial distribution term for *) +(* k successes in n trials with success rate p *) +(* binomial_probability n p1 == binomial probability with n and p1 a proof *) +(* that p : {nonneg R} is smaller than 1 *) +(* binomial_probability_trunc n r == binomial probability with n and real *) +(* number r *) (* uniform_probability a b ab0 == uniform probability over the interval [a,b] *) (* sample mP == sample according to the probability P where mP is a *) (* proof that P is a measurable function *) @@ -195,7 +202,21 @@ Definition bernoulli_trunc (p : R) := match sumbool_ler 0%R p with | inr _ => bernoulli0 end. -(*HB.instance Definition _ (p : R) := Probability.on (bernoulli_trunc p).*) +Lemma bernoulli_truncE (p : R) U : + (0 <= p <= 1)%R -> + (bernoulli_trunc p U = + p%:E * \d_true U + (`1-p)%:E * \d_false U)%E. +Proof. +move=> /andP[p0 p1]. +rewrite /bernoulli_trunc. +case: (sumbool_ler 0 p) => [{}p0/=|]. + case: (sumbool_ler p 1) => [{}p1/=|]. + by rewrite /bernoulli/= measure_addE. + by rewrite ltNge p1. +by rewrite ltNge p0. +Qed. + +(* HB.instance Definition _ (p : R) := Probability.on (bernoulli_trunc p). *) Let simpe := (@mule0 R, @adde0 R, @mule1 R, @add0e R). @@ -225,7 +246,7 @@ have [->|->|->|->] /= := subset_set2 UT. (fun x => if 0 <= x <= 1 then (`1-x)%:E else 1%E))//. apply: measurable_fun_ifT => //=. by apply: measurable_and => //; exact: measurable_fun_ler. - by apply/EFin_measurable_fun; apply/measurable_funB. + by apply/measurable_EFinP; apply/measurable_funB. apply/funext => x/=; case: sumbool_ler => /= x0. case: sumbool_ler => /= x1. rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. @@ -421,16 +442,16 @@ rewrite (_ : f = fun x => if 0 <= x <= 1 then (\sum_(m < n.+1) else \d_0%N Ys)//. move=> ?; apply: measurable_fun_ifT => //=. by apply: measurable_and => //; exact: measurable_fun_ler. - apply: emeasurable_fun_sum => m /=. + apply: emeasurable_sum => m /=. rewrite /mscale/= [X in measurable_fun _ X](_ : _ = (fun x => ((x ^+ m * `1-x ^+ (n - m) *+ 'C(n, m))%:E * \d_(nat_of_ord m) Ys)%E)); last first. by apply:funext => x; case: sumbool_ler => // x0; case: sumbool_ler. - apply: emeasurable_funM => //; apply/EFin_measurable_fun => //. + apply: emeasurable_funM => //; apply/measurable_EFinP => //. under eq_fun do rewrite -mulrnAr. apply: measurable_funM => //. under eq_fun do rewrite -mulr_natr. apply: measurable_funM => //=. - apply: measurable_fun_pow. + apply: measurable_funX. exact: measurable_funB. rewrite {}/f. apply/funext => x. @@ -544,10 +565,10 @@ Proof. move=> m. under [LHS]eq_integral do rewrite fimfunE -fsumEFin//. rewrite [LHS]ge0_integral_fsum//; last 2 first. - - by move=> r; exact/EFin_measurable_fun/measurableT_comp. + - by move=> r; exact/measurable_EFinP/measurableT_comp. - by move=> n x _; rewrite EFinM nnfun_muleindic_ge0. rewrite -[RHS]ge0_integralZl//; last 3 first. - - exact/EFin_measurable_fun/measurable_funTS. + - exact/measurable_EFinP/measurable_funTS. - by move=> x _; rewrite lee_fin. - by rewrite lee_fin invr_ge0// ltW. under [RHS]eq_integral. @@ -555,12 +576,12 @@ under [RHS]eq_integral. by move=> r; rewrite EFinM nnfun_muleindic_ge0. over. rewrite [RHS]ge0_integral_fsum//; last 2 first. - - by move=> r; apply/EFin_measurable_fun; do 2 apply/measurableT_comp => //. + - by move=> r; apply/measurable_EFinP; do 2 apply/measurableT_comp => //. - move=> n x _; rewrite EFinM mule_ge0//; last by rewrite nnfun_muleindic_ge0. by rewrite lee_fin invr_ge0// ltW. apply: eq_fsbigr => r _; rewrite ge0_integralZl//. - by rewrite !integralZl_indic_nnsfun//= integral_uniform_indic// muleCA. -- exact/EFin_measurable_fun/measurableT_comp. +- exact/measurable_EFinP/measurableT_comp. - by move=> t _; rewrite nnfun_muleindic_ge0. - by rewrite lee_fin invr_ge0// ltW. Qed. @@ -572,30 +593,30 @@ Lemma integral_uniform (f : _ -> \bar R) (a b : R) (b - a)^-1%:E * \int[lebesgue_measure]_(x in `[a, b]) f x. Proof. move=> mf f0 m. -have [f_ [ndf_ f_f]] := approximation measurableT mf (fun x0 _ => f0 x0). +pose f_ := nnsfun_approx measurableT mf. transitivity (lim (\int[m]_(x in setT) (f_ n x)%:E @[n --> \oo])). rewrite -monotone_convergence//=. - apply: eq_integral => ? /[!inE] xD; apply/esym/cvg_lim => //=. - exact: f_f. - - by move=> n; exact/EFin_measurable_fun/measurable_funTS. + exact/cvg_nnsfun_approx. + - by move=> n; exact/measurable_EFinP/measurable_funTS. - by move=> n ? _; rewrite lee_fin. - - by move=> ? _ ? ? mn; rewrite lee_fin; exact/lefP/ndf_. + - by move=> ? _ ? ? mn; rewrite lee_fin; exact/lefP/nd_nnsfun_approx. rewrite [X in _ = _ * X](_ : _ = lim (\int[lebesgue_measure]_(x in `[a, b]) (f_ n x)%:E @[n --> \oo])); last first. rewrite -monotone_convergence//=. - - by apply: eq_integral => ? /[!inE] xD; apply/esym/cvg_lim => //; exact: f_f. - - by move=> n; exact/EFin_measurable_fun/measurable_funTS. + - by apply: eq_integral => ? /[!inE] xD; apply/esym/cvg_lim => //; exact: cvg_nnsfun_approx. + - by move=> n; exact/measurable_EFinP/measurable_funTS. - by move=> n ? _; rewrite lee_fin. - - by move=> ? _ ? ? /ndf_ /lefP; rewrite lee_fin. + - by move=> ? _ ? ? ab; rewrite lee_fin; exact/lefP/nd_nnsfun_approx. rewrite -limeMl//. apply: congr_lim. by apply/funext => n /=; exact: integral_uniform_nnsfun. apply/ereal_nondecreasing_is_cvgn => ? ? ab; apply: ge0_le_integral => //=. - by move=> ? _; rewrite lee_fin. -- exact/EFin_measurable_fun/measurable_funTS. +- exact/measurable_EFinP/measurable_funTS. - by move=> ? _; rewrite lee_fin. -- exact/EFin_measurable_fun/measurable_funTS. -- by move=> ? _; rewrite lee_fin; move/ndf_ : ab => /lefP. +- exact/measurable_EFinP/measurable_funTS. +- by move=> ? _; rewrite lee_fin; exact/lefP/nd_nnsfun_approx. Qed. End integral_uniform. @@ -647,7 +668,7 @@ Qed. Lemma measurable_ubeta_nat_pdf : measurable_fun setT ubeta_nat_pdf. Proof. -by apply /measurable_funM => //; exact/measurable_fun_pow/measurable_funB. +by apply /measurable_funM => //; exact/measurable_funX/measurable_funB. Qed. Lemma beta_nat_normE : @@ -678,6 +699,15 @@ Definition ubeta_nat (U : set (measurableTypeR R)) : \bar R := \int[mu]_(x in U `&` `[0, 1](*NB: is this correct?*)) (ubeta_nat_pdf x)%:E. (* TODO: define as \int[uniform_probability p01]_(t in U) (ubeta_nat_pdf t)%:E ? *) +Lemma ubeta_natE U : + (ubeta_nat U = + \int[mu]_(x in U `&` `[0%R, 1%R]) (ubeta_nat_pdf x)%:E :> \bar R)%E. +Proof. by []. Qed. + +Lemma ubeta_nat_lty U : (ubeta_nat U < +oo)%E. +Proof. +Admitted. + Let ubeta_nat0 : ubeta_nat set0 = 0%:E. Proof. by rewrite /ubeta_nat set0I integral_set0. Qed. @@ -687,7 +717,6 @@ rewrite /ubeta_nat integral_ge0//= => x [Ux]. by rewrite in_itv/= => x01; rewrite lee_fin ubeta_nat_pdf_ge0. Qed. -(* TODO: should be shorter *) Let ubeta_nat_sigma_additive : semi_sigma_additive ubeta_nat. Proof. move=> /= F mF tF mUF; rewrite /ubeta_nat setI_bigcupl; apply: cvg_toP. @@ -697,20 +726,7 @@ move=> /= F mF tF mUF; rewrite /ubeta_nat setI_bigcupl; apply: cvg_toP. by rewrite lee_fin; exact: ubeta_nat_pdf_ge0. rewrite ge0_integral_bigcup//=. - by move=> k; exact: measurableI. -- apply/integrableP; split. - by apply/EFin_measurable_fun; exact: measurable_funTS measurable_ubeta_nat_pdf. - apply: le_lt_trans => /=. - apply: (@subset_integral _ _ _ mu _ `[0%R, 1%R]) => //=. - - rewrite -setI_bigcupl; apply: measurableI => //. - - apply/measurableT_comp => //; apply/measurableT_comp => //. - exact: measurable_funTS measurable_ubeta_nat_pdf. - - by apply: bigcup_sub => k _; exact: subIsetr. - rewrite /=. - under eq_integral. - move=> /= x; rewrite inE/= in_itv/= => x01. - rewrite ger0_norm//; last by rewrite ubeta_nat_pdf_ge0. - over. - by rewrite -beta_nat_normE ltry. +- by apply/measurable_EFinP; exact: measurable_funTS measurable_ubeta_nat_pdf. - move=> x [k _ [_]]; rewrite /= in_itv/= => x01. by rewrite lee_fin ubeta_nat_pdf_ge0. - exact: trivIset_setIr. @@ -735,19 +751,6 @@ Proof. move=> /= F mF tF mUF; exact: measure_semi_sigma_additive. Qed. HB.instance Definition _ := isMeasure.Build _ _ _ beta_nat beta_nat0 beta_nat_ge0 beta_nat_sigma_additive. -(* HB.instance Definition _ := Measure.on beta_nat. *) - -(*Let beta_nat0 : beta_nat set0 = 0. -Proof. exact: measure0. Qed. - -Let beta_nat_ge0 U : (0 <= beta_nat U)%E. -Proof. exact: measure_ge0. Qed. - -Let beta_nat_sigma_additive : semi_sigma_additive beta_nat. -Proof. move=> /= F mF tF mUF; exact: measure_semi_sigma_additive. Qed. - -HB.instance Definition _ := isMeasure.Build _ _ _ beta_nat beta_nat0 beta_nat_ge0 beta_nat_sigma_additive.*) - Let beta_nat_setT : beta_nat setT = 1%:E. Proof. rewrite /beta_nat /= /mscale /=. @@ -767,6 +770,8 @@ Qed. End beta_probability. +Arguments beta_nat {R}. + Section beta_probability11. Local Open Scope ring_scope. Context {R : realType}. @@ -788,6 +793,356 @@ Qed. End beta_probability11. +Lemma factD n m : (n`! * m`! <= (n + m).+1`!)%N. +Proof. +elim: n m => /= [m|n ih m]. + by rewrite fact0 mul1n add0n factS leq_pmull. +rewrite 2!factS [in X in (_ <= _ * X)%N]addSn -mulnA leq_mul//. +by rewrite ltnS addSnnS leq_addr. +Qed. + +Lemma factD' n m : (n`! * m.-1`! <= (n + m)`!)%N. +Proof. +case: m => //= [|m]. + by rewrite fact0 muln1 addn0. +by rewrite addnS factD. +Qed. + +Lemma leq_prod2 (x y n m : nat) : (n <= x)%N -> (m <= y)%N -> + (\prod_(m <= i < y) i * \prod_(n <= i < x) i <= \prod_(n + m <= i < x + y) i)%N. +Proof. +move=> nx my. +rewrite big_addn. +rewrite -addnBA//. +rewrite {3}/index_iota. +rewrite -addnBAC//. +rewrite iotaD. +rewrite big_cat/=. +rewrite mulnC. +rewrite leq_mul//. + rewrite /index_iota. + apply: leq_prod. + by move=> i _; rewrite leq_addr. +rewrite subnKC//. +rewrite -{1}(add0n m). +rewrite big_addn. +rewrite {2}(_ : (y - m) = ((y - m + x) - x))%N; last first. + by rewrite -addnBA// subnn addn0. +rewrite -{1}(add0n x). +rewrite big_addn. +rewrite -addnBA// subnn addn0. +apply: leq_prod => i _. +by rewrite leq_add2r leq_addr. +Qed. + +Lemma leq_fact2 (x y n m : nat) : + (n <= x) %N -> (m <= y)%N -> + (x`! * y`! * ((n + m).+1)`! <= n`! * m`! * ((x + y).+1)`!)%N. +Proof. +move=> nx my. +rewrite (_ : x`! = n`! * \prod_(n.+1 <= i < x.+1) i)%N; last first. + by rewrite -fact_split. +rewrite -!mulnA leq_mul2l; apply/orP; right. +rewrite (_ : y`! = m`! * \prod_(m.+1 <= i < y.+1) i)%N; last first. + by rewrite -fact_split. +rewrite mulnCA -!mulnA leq_mul2l; apply/orP; right. +rewrite (_ : (x + y).+1`! = (n + m).+1`! * \prod_((n + m).+2 <= i < (x + y).+2) i)%N; last first. + rewrite -fact_split//. + by rewrite ltnS leq_add. +rewrite mulnA mulnC leq_mul2l; apply/orP; right. +rewrite -addSn -addnS. +rewrite -addSn -addnS. +exact: leq_prod2. +Qed. + +Section integral_beta. +Context {R : realType}. +Variables a b a' b' : nat. + +Local Notation mu := lebesgue_measure. + +Lemma integralMl f g1 g2 A : +measurable A -> measurable_fun A f -> + measurable_fun A g1 -> measurable_fun A g2 -> +(ae_eq mu A g1 (EFin \o g2)) -> + \int[mu]_(x in A) (f x * g1 x) = + \int[mu]_(x in A) (f x * (g2 x)%:E) :> \bar R. +Proof. +move=> mA mf mg1 mg2 Hg. +apply: ae_eq_integral => //. + by apply: emeasurable_funM. + apply: emeasurable_funM => //. + by apply/measurable_EFinP. +by apply: ae_eq_mul2l. +Qed. + +Let beta_nat_dom : (@beta_nat R a b `<< mu). +Proof. +move=> A mA muA0. +rewrite /beta_nat /mscale/= /ubeta_nat. +have -> : \int[mu]_(x0 in A `&` `[0%R, 1%R]) (ubeta_nat_pdf a b x0)%:E = 0%:E. + apply/eqP; rewrite eq_le. + apply/andP; split; last first. + apply: integral_ge0 => x [Ax /=]. + rewrite in_itv /= => x01. + by rewrite lee_fin ubeta_nat_pdf_ge0. + apply: le_trans. + apply: (@ge0_subset_integral _ _ _ _ _ A). + by apply: measurableI. + by []. + apply/measurable_EFinP. + apply: (@measurable_funS _ _ _ _ setT) => //=. + apply: measurable_ubeta_nat_pdf. + move=> x Ax. + have : (`[0%R, 1%R]%classic x). + admit. + rewrite /= in_itv/=. + apply: ubeta_nat_pdf_ge0. + apply: subIsetl. + rewrite /=. + (* rewrite integral_abs_eq0. *) (* without abs *) + admit. +by rewrite mule0. +Admitted. + +Lemma integral_beta_nat f : + measurable_fun setT f -> + \int[beta_nat a b]_(x in `[0%R, 1%R]) `|f x| < +oo -> + \int[beta_nat a b]_(x in `[0%R, 1%R]) f x = + \int[mu]_(x in `[0%R, 1%R]) (f x * (beta_nat_pdf a b x)%:E) :> \bar R. +Proof. +move=> mf finf. +rewrite -(Radon_Nikodym_change_of_variables beta_nat_dom) //=. +apply: integralMl => //. + apply: (@measurable_funS _ _ _ _ [set: R]) => //. + apply: (@measurable_funS _ _ _ _ [set: R]) => //. + rewrite Radon_NikodymE. + by exact: beta_nat_dom. + move=> /= H. + case: cid => /= h [h1 h2 h3]. + have : (measurable_fun setT h /\ \int[mu]_x `|h x| < +oo). + apply/integrableP/h2. + move=> /= [mh _]. + apply: mh. + apply: (@measurable_funS _ _ _ _ [set: R]) => //. + apply: measurable_beta_nat_pdf. + rewrite Radon_NikodymE => /= A. + by exact: beta_nat_dom. +case: cid => /= h [h1 h2 h3]. +apply: integral_ae_eq => //. + apply: integrableS h2 => //. (* integrableST? *) + apply: (@measurable_funS _ _ _ _ [set: R]) => //. + apply: measurableT_comp => //. + apply: measurable_beta_nat_pdf. + move=> E E01 mE. + have mB : measurable_fun E (EFin \o ubeta_nat_pdf a b). + apply: measurableT_comp => //. + apply: (@measurable_funS _ _ _ _ [set: R]) => //. + apply: measurable_ubeta_nat_pdf. + rewrite -(h3 _ mE). + rewrite /beta_nat/mscale/ubeta_nat/beta_nat_pdf/=. + under eq_integral do rewrite mulrC EFinM. + rewrite (integralZl mE). + rewrite /ubeta_nat setIidl //. + rewrite /=. + apply/integrableP; split. + by apply: mB. + under eq_integral => x x01. + rewrite gee0_abs /=. + over. + apply: ubeta_nat_pdf_ge0. + have : x \in `[0%R, 1%R]. + apply: (@subset_trans _ _ `[x,x] _ _ E01). + by rewrite set_interval.set_itv1 sub1set x01. + by rewrite /= in_itv/= lexx. + by rewrite in_itv/=. + rewrite /=. + have <- := (setIidl E01). + by rewrite -ubeta_natE ubeta_nat_lty. +apply/integrableP; split. + by apply: (@measurable_funS _ _ _ _ [set: R]). +exact: finf. +Qed. + +Local Open Scope ring_scope. + +(* TODO: `[0, 1]? *) +Definition beta_nat_bern U : \bar R := + \int[beta_nat a b]_y bernoulli_trunc (ubeta_nat_pdf a'.+1 b'.+1 y) U. + +Local Notation B := beta_nat_norm. + +Definition Baa'bb'Bab : R := (beta_nat_norm (a + a') (b + b')) / beta_nat_norm a b. + +Lemma Baa'bb'Bab_ge0 : 0 <= Baa'bb'Bab. +Proof. by rewrite /Baa'bb'Bab divr_ge0// beta_nat_norm_ge0. Qed. + +Definition Baa'bb'Bab_nneg : {nonneg R} := NngNum Baa'bb'Bab_ge0. + +Lemma Baa'bb'Bab_le1 : Baa'bb'Bab_nneg%:num <= 1. +Proof. +rewrite /Baa'bb'Bab_nneg/= /Baa'bb'Bab. +rewrite ler_pdivrMr// ?mul1r ?beta_nat_norm_gt0//. +rewrite /B /beta_nat_norm. +rewrite ler_pdivrMr ?ltr0n ?fact_gt0//. +rewrite mulrAC. +rewrite ler_pdivlMr ?ltr0n ?fact_gt0//. +rewrite -!natrM ler_nat. +case: a. + rewrite /= fact0 mul1n !add0n. + case: b => /=. + case: a' => //. + case: b' => //= m. + by rewrite fact0 !mul1n muln1. + move=> n/=. + by rewrite fact0 add0n muln1 mul1n factD'. + move=> m. + rewrite mulnC leq_mul// mulnC. + by rewrite (leq_trans (factD' _ _))// addSn addnS//= addnC. +move=> n. +rewrite addSn. +case: b. + rewrite !fact0 add0n muln1 [leqRHS]mulnC addn0/= leq_mul//. + by rewrite factD'. +move=> m. +clear a b. +rewrite [(n + a').+1.-1]/=. +rewrite [n.+1.-1]/=. +rewrite [m.+1.-1]/=. +rewrite addnS. +rewrite [(_ + m).+1.-1]/=. +rewrite (addSn m b'). +rewrite [(m + _).+1.-1]/=. +rewrite (addSn (n + a')). +rewrite [_.+1.-1]/=. +rewrite addSn addnS. +by rewrite leq_fact2// leq_addr. +Qed. + +Lemma onem_Baa'bb'Bab_ge0 : 0 <= 1 - (Baa'bb'Bab_nneg%:num). +Proof. by rewrite subr_ge0 Baa'bb'Bab_le1. Qed. + +Lemma onem_Baa'bb'Bab_ge0_fix : 0 <= B a b * (1 - Baa'bb'Bab_nneg%:num). +Proof. +rewrite mulr_ge0//. + rewrite /B. + exact: beta_nat_norm_ge0. +rewrite subr_ge0. +exact: Baa'bb'Bab_le1. +Qed. + +Lemma ubeta_nat_pdf_ge0' t : 0 <= ubeta_nat_pdf a'.+1 b'.+1 t :> R. +Proof. +apply: ubeta_nat_pdf_ge0. (* TODO: needs 0 <= t <= 1 *) +Admitted. + +Lemma ubeta_nat_pdf_le1' t : (NngNum (ubeta_nat_pdf_ge0' t))%:num <= 1 :> R. +Proof. +rewrite /=. +rewrite /ubeta_nat_pdf. +rewrite /ubeta_nat_pdf'. (* TODO: needs 0 <= t <= 1 *) +Admitted. + +Lemma integral_ubeta_nat' : + (\int[ubeta_nat a b]_x (ubeta_nat_pdf a'.+1 b'.+1 x)%:E = + \int[mu]_(x in `[0%R, 1%R]) + (x ^+ a'.-1 * `1-x ^+ b'.-1 * x ^+ a * `1-x ^+ b)%:E :> \bar R)%E. +Proof. +rewrite /ubeta_nat/ubeta_nat_pdf. +Admitted. + +Lemma beta_nat_bern_bern U : + (a > 0)%N -> (b > 0)%N -> + beta_nat_bern U = + bernoulli_trunc Baa'bb'Bab U. +Proof. +rewrite /beta_nat_bern. +transitivity ((\int[beta_nat a b]_(y in `[0%R, 1%R]) + bernoulli_trunc (ubeta_nat_pdf a'.+1 b'.+1 y) U)%E : \bar R). + admit. +rewrite integral_beta_nat /=; last 2 first. + apply: (@measurableT_comp _ _ _ _ _ _ (bernoulli_trunc ^~ U)). + apply: (measurability (ErealGenInftyO.measurableE R)) => //=. + move=> /= _ [_ [x ->] <-]; apply: measurableI => //. + admit. + exact: measurable_ubeta_nat_pdf. + admit. +under eq_integral => x. + rewrite inE/= in_itv/= => x01. + rewrite bernoulli_truncE. + over. + apply/andP; split. + apply/ubeta_nat_pdf_ge0/x01. + apply/ubeta_nat_pdf_le1/x01. +rewrite /=. +rewrite bernoulli_truncE; last first. + apply/andP; split. + exact: Baa'bb'Bab_ge0. + exact: Baa'bb'Bab_le1. +under eq_integral => x _. + rewrite muleC muleDr//. + over. +rewrite integralD//=; last 2 first. + (* TODO: integrableM *) + admit. + admit. +congr (_ + _). + under eq_integral do rewrite muleA muleC. + rewrite integralZl//=; last first. + admit. + rewrite muleC. + congr (_ * _)%E. + rewrite /beta_nat_pdf. + under eq_integral do rewrite EFinM -muleA muleC -muleA. + rewrite integralZl//=; last first. + admit. + transitivity (((beta_nat_norm a b)^-1)%:E * \int[lebesgue_measure]_(x in `[0%R, 1%R]) ((ubeta_nat_pdf (a+a') (b+b') x)%:E) : \bar R)%E. + congr (_ * _)%E. + apply: eq_integral => x x01. + rewrite /ubeta_nat_pdf /ubeta_nat_pdf' muleC /onem -EFinM/=. + rewrite mulrCA -mulrA -exprD mulrA -exprD. + congr (_ ^+ _ * _ ^+ _)%:E. + rewrite -!subn1 subDnCA//. + rewrite addnC -!subn1 subDnCA//. + rewrite -beta_nat_normE. + rewrite /Baa'bb'Bab/B -!EFinM. + congr _%:E. + rewrite mulrC//. +under eq_integral do rewrite muleA muleC. +rewrite integralZl//=; last first. + admit. +rewrite muleC. +congr (_ * _)%E. +rewrite /beta_nat_pdf. +under eq_integral do rewrite EFinM -muleA muleC -muleA. +rewrite integralZl//=; last first. + admit. +transitivity (((beta_nat_norm a b)^-1)%:E * \int[lebesgue_measure]_(x in `[0%R, 1%R]) ((ubeta_nat_pdf a b x)%:E - (ubeta_nat_pdf (a+a') (b+b') x)%:E) : \bar R)%E. + congr (_ * _)%E. + apply: eq_integral => x x01. + rewrite /onem -EFinM mulrBl mul1r EFinB. + congr (_ - _)%E. + rewrite /ubeta_nat_pdf /ubeta_nat_pdf'/=. + rewrite mulrCA -mulrA -exprD mulrA -exprD. + congr (_ ^+ _ * _ ^+ _)%:E. + rewrite addnC -!subn1 subDnCA//. + rewrite -!subn1 subDnCA//. +rewrite integralB_EFin//=; last 2 first. + admit. + admit. +rewrite -!beta_nat_normE -EFinM mulrBr /onem mulVf; last first. + rewrite /B mulf_eq0 negb_or. + apply/andP; split. + rewrite mulf_eq0 negb_or. + rewrite gt_eqF ?ltr0n ?fact_gt0//=. + rewrite gt_eqF ?ltr0n ?fact_gt0//=. + rewrite invr_eq0 gt_eqF ?ltr0n ?fact_gt0//=. +congr (_ - _)%:E. +by rewrite mulrC. +Admitted. + +End integral_beta. + Section mscore. Context d (T : measurableType d) (R : realType). Variable f : T -> R. From bcc89f28b80ea659cd84197d39d1f90334718d74 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Sun, 24 Mar 2024 21:57:17 +0900 Subject: [PATCH 21/48] complete the casino example - new version of Bernoulli - beta_nat_norm_sym - generic change variables - beta_nat_normE - cleaning --- _CoqProject | 2 +- coq-mathcomp-analysis.opam | 4 +- theories/Make | 2 +- theories/lang_syntax.v | 1631 ++++++++-- theories/lang_syntax_examples.v | 616 ++-- theories/lang_syntax_table_game.v | 717 +++++ theories/normedtype_theory/ereal_normedtype.v | 6 +- theories/prob_lang.v | 2614 ++++++----------- theories/prob_lang_wip.v | 222 +- 9 files changed, 3527 insertions(+), 2287 deletions(-) create mode 100644 theories/lang_syntax_table_game.v diff --git a/_CoqProject b/_CoqProject index 0936ccbc72..bda072280d 100644 --- a/_CoqProject +++ b/_CoqProject @@ -142,4 +142,4 @@ theories/lang_syntax_util.v theories/lang_syntax_toy.v theories/lang_syntax.v theories/lang_syntax_examples.v -theories/lang_syntax_examples_wip.v +theories/lang_syntax_table_game.v diff --git a/coq-mathcomp-analysis.opam b/coq-mathcomp-analysis.opam index c7f50a6eed..ed4e7d5181 100644 --- a/coq-mathcomp-analysis.opam +++ b/coq-mathcomp-analysis.opam @@ -19,8 +19,8 @@ depends: [ "coq-mathcomp-solvable" "coq-mathcomp-field" "coq-mathcomp-bigenough" { (>= "1.0.0") } - "coq-mathcomp-algebra-tactics" { (>= "1.2.0") } - "coq-mathcomp-zify" { (>= "1.4.0") } + "coq-mathcomp-algebra-tactics" { (>= "1.2.3") } + "coq-mathcomp-zify" { (>= "1.5.0") } ] tags: [ diff --git a/theories/Make b/theories/Make index 80466f4fc1..b17786b9dc 100644 --- a/theories/Make +++ b/theories/Make @@ -103,4 +103,4 @@ lang_syntax_util.v lang_syntax_toy.v lang_syntax.v lang_syntax_examples.v -lang_syntax_examples_wip.v +lang_syntax_table_game.v diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index aaf541d2f0..089915ef6f 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -2,16 +2,25 @@ Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. From mathcomp Require Import interval_inference. -From mathcomp Require Import mathcomp_extra boolp classical_sets. +From mathcomp Require Import unstable mathcomp_extra boolp classical_sets. From mathcomp Require Import functions cardinality fsbigop interval_inference. From mathcomp Require Import reals ereal topology normedtype sequences. -From mathcomp Require Import esum measure lebesgue_measure numfun. -From mathcomp Require Import lebesgue_integral kernel prob_lang. -From mathcomp Require Import lang_syntax_util. +From mathcomp Require Import esum measure lebesgue_measure numfun derive realfun. +From mathcomp Require Import lebesgue_integral probability ftc kernel charge. +From mathcomp Require Import prob_lang lang_syntax_util. From mathcomp Require Import ring lra. -(******************************************************************************) -(* Syntax and Evaluation for a Probabilistic Programming Language *) +(**md**************************************************************************) +(* # Syntax and Evaluation for a Probabilistic Programming Language *) +(* *) +(* Reference: *) +(* - R. Saito, R. Affeldt. Experimenting with an Intrinsically-Typed *) +(* Probabilistic Programming Language in Coq using s-finite kernels in Coq. *) +(* APLAS 2023 *) +(* *) +(* beta distribution specialized to nat *) +(* beta_pdf == probability density function for beta *) +(* beta_prob == beta probability measure *) (* *) (* typ == syntax for types of data structures *) (* measurable_of_typ t == the measurable type corresponding to type t *) @@ -39,9 +48,6 @@ From mathcomp Require Import ring lra. (* function from mctx g to mtyp t *) (* pval R g t == "probabilistic value", i.e., *) (* s-finite kernel, from mctx g to mtyp t *) -(* mkswap k == given a kernel k : (Y * X) ~> Z, *) -(* returns a kernel of type (X * Y) ~> Z *) -(* letin' := mkcomp \o mkswap *) (* e -D> f ; mf == the evaluation of the deterministic expression e *) (* leads to the deterministic value f *) (* (mf is the proof that f is measurable) *) @@ -62,8 +68,6 @@ Unset Printing Implicit Defensive. Import Order.TTheory GRing.Theory Num.Def Num.Theory. Import numFieldTopology.Exports. -Reserved Notation "f .; g" (at level 60, right associativity, - format "f .; '/ ' g"). Reserved Notation "e -D> f ; mf" (at level 40). Reserved Notation "e -P> k" (at level 40). @@ -71,193 +75,1297 @@ Local Open Scope classical_set_scope. Local Open Scope ring_scope. Local Open Scope ereal_scope. -(* TODO: mv *) -Arguments measurable_fst {d1 d2 T1 T2}. -Arguments measurable_snd {d1 d2 T1 T2}. +Lemma RintegralZl d {T : measurableType d} {R : realType} + {mu : measure T R} {D : set T} : d.-measurable D -> + forall f : T -> R, + mu.-integrable D (EFin \o f) -> + forall r : R, (\int[mu]_(x in D) (r * f x) = r * \int[mu]_(x in D) f x)%R. +Proof. +move=> mD f intf r. +rewrite /Rintegral. +under eq_integral do rewrite EFinM. +rewrite integralZl// fineM//=. +by apply: integral_fune_fin_num. +Qed. -Section mswap. -Context d d' d3 (X : measurableType d) (Y : measurableType d') - (Z : measurableType d3) (R : realType). -Variable k : R.-ker Y * X ~> Z. +(* TODO: naming *) +Lemma cvg_atNP {T : topologicalType} {R : numFieldType} (f : R -> T) (a : R) (l : T) : + f x @[x --> a] --> l <-> (f \o -%R) x @[x --> (- a)%R] --> l. +Proof. +rewrite nbhsN. +have <-// : f x @[x --> a] = fmap [eta f \o -%R] ((- x)%R @[x --> a]). +by apply/seteqP; split=> A; move=> [/= e e0 H]; exists e => //= B /H/=; rewrite opprK. +Qed. -Definition mswap xy U := k (swap xy) U. +Lemma derivable_oo_bnd_id {R : numFieldType} (a b : R) : + derivable_oo_continuous_bnd (@id R^o) a b. +Proof. +by split => //; + [exact/cvg_at_right_filter/cvg_id|exact/cvg_at_left_filter/cvg_id]. +Qed. + +Lemma derivable_oo_bndN {R : realFieldType} (f : R -> R^o) a b : + derivable_oo_continuous_bnd f (- a) (- b) -> + derivable_oo_continuous_bnd (f \o -%R) b a. +Proof. +move=> [dF cFa cFb]. +have oppK : (-%R \o -%R) = @id R by apply/funext => x/=; rewrite opprK. +split. +- move=> x xba; apply/derivable1_diffP. + apply/(@differentiable_comp _ _ R^o _ -%R f x) => //. + by apply/derivable1_diffP/dF; rewrite oppr_itvoo 2!opprK. +- by apply/cvg_at_rightNP; rewrite -compA oppK. +- by apply/cvg_at_leftNP; rewrite -compA oppK. +Qed. -Let mswap0 xy : mswap xy set0 = 0. -Proof. done. Qed. +Module increasing_change_of_variables_from_decreasing. +Section lt0. +Context {R : realType}. +Notation mu := lebesgue_measure. +Implicit Types (F G f : R -> R) (a b : R). -Let mswap_ge0 x U : 0 <= mswap x U. -Proof. done. Qed. +Lemma continuous_withinN f a b : (a < b)%R -> + {within `[(- b)%R, (- a)%R], continuous (f \o -%R)} -> + {within `[a, b], continuous f}. +Proof. +move=> ab cf. +- apply/continuous_within_itvP (* TODO: us [/\ ...] *) => //. +- split; rewrite -ltrN2 in ab. + + move=> x xab. + move/continuous_within_itvP : cf => /(_ ab) [cf _ _]. + rewrite (_ : f = (f \o -%R) \o -%R); last first. + by apply/funext => y; rewrite /= opprK. + apply: continuous_comp; first exact: opp_continuous. + by apply: cf; rewrite -oppr_itvoo opprK. + + move/continuous_within_itvP : cf => /(_ ab) [_ _ cf]. + apply/cvg_at_rightNP. + by rewrite /= opprK in cf. + + move/continuous_within_itvP : cf => /(_ ab) [_ cf _]. + apply/cvg_at_leftNP. + by rewrite /= opprK in cf. +Qed. -Let mswap_sigma_additive x : semi_sigma_additive (mswap x). -Proof. exact: measure_semi_sigma_additive. Qed. +Lemma oppr_change (f : R -> R) a b : (a < b)%R -> + {within `[a, b], continuous f} -> + \int[mu]_(x in `[a, b]) (f x)%:E = + \int[mu]_(x in `[-%R b, -%R a]) ((f \o -%R) x)%:E. +Proof. +move=> ab cf. +have dN : ((-%R : R -> R^o)^`() = cst (-1) :> (R -> R))%R. (* TODO: lemma? *) + by apply/funext => x/=; rewrite derive1E deriveN// derive_id. +rewrite integration_by_substitution_decreasing//. +- by apply: eq_integral => /= x _; rewrite dN/= opprK mulr1 -compA/= opprK. +- by move=> x y _ _ yx; rewrite ltrN2. +- by move=> y yab; rewrite dN; exact: cvg_cst. +- by rewrite dN; exact: is_cvg_cst. +- by rewrite dN; exact: is_cvg_cst. +- by apply: (@derivable_oo_bndN _ id) => //; exact: derivable_oo_bnd_id. +- apply: continuous_withinN. + + by rewrite ltrN2. + + rewrite -(_ : f = (f \o -%R) \o -%R)//; last first. + by apply/funext => y; rewrite /= opprK. + by rewrite !opprK. +Qed. -HB.instance Definition _ x := isMeasure.Build _ _ R - (mswap x) (mswap0 x) (mswap_ge0 x) (@mswap_sigma_additive x). +End lt0. +End increasing_change_of_variables_from_decreasing. -Definition mkswap : _ -> {measure set Z -> \bar R} := - fun x => mswap x. +Lemma decreasing_nonincreasing {R : realType} (F : R -> R) (J : interval R) : + {in J &, {homo F : x y /~ (x < y)%R}} -> + {in J &, {homo F : x y /~ (x <= y)%R}}. +Proof. +move=> dF. +move=> x y x01 y01. +by rewrite le_eqVlt => /predU1P[->//|/dF] => /(_ x01 y01)/ltW. +Qed. -Let measurable_fun_kswap U : - measurable U -> measurable_fun setT (mkswap ^~ U). +Lemma derive1_onem {R: realType} : (fun x0 : R => (1 - x0)%R : R^o)^`() = (cst (-1)%R). Proof. -move=> mU. -rewrite [X in measurable_fun _ X](_ : _ = k ^~ U \o @swap _ _)//. -apply measurableT_comp => //=; first exact: measurable_kernel. -exact: measurable_swap. +apply/funext => x. +by rewrite derive1E deriveB// derive_id derive_cst sub0r. +Qed. + +Local Close Scope ereal_scope. +Lemma cvg_comp_filter {R : realType} (f g : R -> R) (r l : R) : + continuous f -> + (f \o g) x @[x --> r] --> l -> + f x @[x --> g r] --> l. +Proof. +move=> cf fgrl. +apply/(@cvgrPdist_le _ R^o) => /= e e0. +have e20 : 0 < e / 2 by rewrite divr_gt0. +move/(@cvgrPdist_le _ R^o) : fgrl => /(_ _ e20) fgrl. +have := cf (g r). +move=> /(@cvgrPdist_le _ R^o) => /(_ _ e20)[x x0]H. +exists (minr x (e/2)). + by rewrite lt_min x0. +move=> z. +rewrite /ball_ /= => grze. +rewrite -[X in X - _](subrK (f (g r))). +rewrite -(addrA _ _ (- f z)). +apply: (le_trans (ler_normD _ _)). +rewrite (splitr e) lerD//. + case: fgrl => d /= d0 K. + apply: K. + by rewrite /ball_/= subrr normr0. +apply: H => /=. +by rewrite (lt_le_trans grze)// ge_min lexx. Qed. +Local Open Scope ereal_scope. -HB.instance Definition _ := isKernel.Build _ _ - (X * Y)%type Z R mkswap measurable_fun_kswap. +Section change_of_variables_onem. +Context {R : realType}. +Let mu := (@lebesgue_measure R). -End mswap. +Lemma onem_change (G : R -> R) (r : R) : + (0 < r <= 1)%R -> + {within `[0%R, r], continuous G} -> + (\int[mu]_(x in `[0%R, r]) (G x)%:E = + \int[mu]_(x in `[(1 - r)%R, 1%R]) (G (1 - x))%:E). +Proof. +move=> r01 cG. +have := @integration_by_substitution_decreasing R (fun x => 1 - x)%R G (1 - r) 1%R. +rewrite opprB subrr addrCA subrr addr0. +move=> ->//. +- apply: eq_integral => x xr. + rewrite !fctE. + by rewrite derive1_onem opprK mulr1. +- rewrite ltrBlDl ltrDr. + by case/andP : r01. +- by move=> x y _ _ xy; rewrite ler_ltB. +- by rewrite derive1_onem; move=> ? ?; apply: cvg_cst. +- by rewrite derive1_onem; exact: is_cvg_cst. +- by rewrite derive1_onem; exact: is_cvg_cst. +- split => /=. + + move=> x xr1. + by apply: derivableB => //. + + apply: cvg_at_right_filter. + rewrite opprB addrCA addrA addrK. + apply: (@cvg_comp_filter _ _ (fun x => 1 - x)%R)=> //=. + move=> x. + apply: (@continuousB _ R^o) => //. + exact: cvg_cst. + under eq_fun do rewrite opprD addrA subrr add0r opprK. + apply: cvg_id. + apply: cvg_at_left_filter. + apply: (@cvgB _ R^o) => //. + exact: cvg_cst. +Qed. -Section mswap_sfinite_kernel. -Variables (d d' d3 : _) (X : measurableType d) (Y : measurableType d') - (Z : measurableType d3) (R : realType). -Variable k : R.-sfker Y * X ~> Z. +Lemma Ronem_change (G : R -> R) (r : R) : + (0 < r <= 1)%R -> + locally_integrable [set: R] G -> + {within `[0%R, r], continuous G} -> + (forall r, lebesgue_measure.-integrable `[0%R, r] (EFin \o G)) -> + (\int[mu]_(x in `[0%R, r]) (G x) = + \int[mu]_(x in `[(1 - r)%R, 1%R]) (G (1 - x)))%R. +Proof. +move=> r01 locG cG iG. +rewrite [in LHS]/Rintegral. +by rewrite onem_change. +Qed. + +End change_of_variables_onem. + +Section factD. -Let mkswap_sfinite : - exists2 k_ : (R.-ker X * Y ~> Z)^nat, - forall n, measure_fam_uub (k_ n) & - forall x U, measurable U -> mkswap k x U = kseries k_ x U. +Let factD' n m : (n`! * m`! <= (n + m).+1`!)%N. Proof. -have [k_ /= kE] := sfinite_kernel k. -exists (fun n => mkswap (k_ n)). - move=> n. - have /measure_fam_uubP[M hM] := measure_uub (k_ n). - by exists M%:num => x/=; exact: hM. -move=> xy U mU. -by rewrite /mswap/= kE. +elim: n m => /= [m|n ih m]. + by rewrite fact0 mul1n add0n factS leq_pmull. +rewrite 2!factS [in X in (_ <= _ * X)%N]addSn -mulnA leq_mul//. +by rewrite ltnS addSnnS leq_addr. Qed. -HB.instance Definition _ := - isSFiniteKernel_subdef.Build _ _ _ Z R (mkswap k) mkswap_sfinite. +Lemma factD n m : (n`! * m.-1`! <= (n + m)`!)%N. +Proof. +case: m => //= [|m]. + by rewrite fact0 muln1 addn0. +by rewrite addnS factD'. +Qed. + +End factD. -End mswap_sfinite_kernel. +Lemma leq_prod2 (x y n m : nat) : (n <= x)%N -> (m <= y)%N -> + (\prod_(m <= i < y) i * \prod_(n <= i < x) i <= \prod_(n + m <= i < x + y) i)%N. +Proof. +move=> nx my; rewrite big_addn -addnBA//. +rewrite [in leqRHS]/index_iota -addnBAC// iotaD big_cat/=. +rewrite mulnC leq_mul//. + by apply: leq_prod; move=> i _; rewrite leq_addr. +rewrite subnKC//. +rewrite -[in leqLHS](add0n m) big_addn. +rewrite [in leqRHS](_ : y - m = ((y - m + x) - x))%N; last first. + by rewrite -addnBA// subnn addn0. +rewrite -[X in iota X _](add0n x) big_addn -addnBA// subnn addn0. +by apply: leq_prod => i _; rewrite leq_add2r leq_addr. +Qed. -Section kswap_finite_kernel_finite. -Context d d' d3 (X : measurableType d) (Y : measurableType d') - (Z : measurableType d3) (R : realType) - (k : R.-fker Y * X ~> Z). +Lemma leq_fact2 (x y n m : nat) : (n <= x) %N -> (m <= y)%N -> + (x`! * y`! * ((n + m).+1)`! <= n`! * m`! * ((x + y).+1)`!)%N. +Proof. +move=> nx my. +rewrite (fact_split nx) -!mulnA leq_mul2l; apply/orP; right. +rewrite (fact_split my) mulnCA -!mulnA leq_mul2l; apply/orP; right. +rewrite [leqRHS](_ : _ = (n + m).+1`! * \prod_((n + m).+2 <= i < (x + y).+2) i)%N; last first. + by rewrite -fact_split// ltnS leq_add. +rewrite mulnA mulnC leq_mul2l; apply/orP; right. +do 2 rewrite -addSn -addnS. +exact: leq_prod2. +Qed. -Let mkswap_finite : measure_fam_uub (mkswap k). +Lemma bounded_norm_expn_onem {R : realType} (a b : nat) : + [bounded `|x ^+ a * (1 - x) ^+ b|%R : R^o | x in (`[0%R, 1%R]%classic : set R)]. Proof. -have /measure_fam_uubP[r hr] := measure_uub k. -apply/measure_fam_uubP; exists (PosNum [gt0 of r%:num%R]) => x /=. -exact: hr. +exists 1%R; split; [by rewrite num_real|move=> x x1 /= y]. +rewrite in_itv/= => /andP[y0 y1]. +rewrite ger0_norm// ger0_norm; last first. + by rewrite mulr_ge0 ?exprn_ge0// subr_ge0. +rewrite (le_trans _ (ltW x1))// mulr_ile1 ?exprn_ge0//. +- by rewrite subr_ge0. +- by rewrite exprn_ile1. +- rewrite exprn_ile1 ?subr_ge0//. + by rewrite lerBlDl addrC -lerBlDl subrr. +Qed. + +Lemma measurable_fun_expn_onem {R : realType} a b : + measurable_fun setT (fun x : R => x ^+ a * `1-x ^+ b)%R. +Proof. +apply/measurable_funM => //; apply/measurable_funX => //. +exact: measurable_funB. +Qed. + +Section ubeta_nat_pdf. +Local Open Scope ring_scope. +Context {R : realType}. +Variables a b : nat. + +(* unnormalized pdf *) +(*Definition ubeta_nat_pdf (t : R) := + if (0 <= t <= 1)%R then (t ^+ a.-1 * (`1-t) ^+ b.-1)%R else 0%R. + +Lemma ubeta_nat_pdf_ge0 t : 0 <= ubeta_nat_pdf t. +Proof. +rewrite /ubeta_nat_pdf; case: ifPn => // /andP[t0 t1]. +by rewrite mulr_ge0// exprn_ge0// onem_ge0. +Qed. + +Lemma ubeta_nat_pdf_le1 t : ubeta_nat_pdf t <= 1. +Proof. +rewrite /ubeta_nat_pdf; case: ifPn => // /andP[t0 t1]. +by rewrite mulr_ile1// ?(exprn_ge0,onem_ge0,exprn_ile1,onem_le1). +Qed. + +Lemma measurable_ubeta_nat_pdf : measurable_fun setT ubeta_nat_pdf. +Proof. +rewrite /ubeta_nat_pdf /=; apply: measurable_fun_if => //=; last first. + by rewrite setTI; apply: measurable_funTS; exact: measurable_fun_expn_onem. +by apply: measurable_and => /=; exact: measurable_fun_ler. +Qed. + +Local Notation mu := lebesgue_measure. + +Lemma integral_ubeta_nat_pdf U : + (\int[mu]_(x in U) (ubeta_nat_pdf x)%:E = + \int[mu]_(x in U `&` `[0%R, 1%R]) (ubeta_nat_pdf x)%:E)%E. +Proof. +rewrite [RHS]integral_mkcondr/=; apply: eq_integral => x xU. +rewrite patchE; case: ifPn => //. +rewrite notin_setE/= in_itv/= => /negP. +rewrite negb_and -!ltNge => /orP[x0|x1]. + by rewrite /ubeta_nat_pdf leNgt x0/=. +by rewrite /ubeta_nat_pdf !leNgt x1/= andbF. +Qed. + +Lemma integral_ubeta_nat_pdfT : + (\int[mu]_x (ubeta_nat_pdf x)%:E = + \int[mu]_(x in `[0%R, 1%R]) (ubeta_nat_pdf x)%:E)%E. +Proof. by rewrite integral_ubeta_nat_pdf/= setTI. Qed.*) + +End ubeta_nat_pdf. + +(*Lemma ubeta_nat_pdf11 {R : realType} (x : R) : (0 <= x <= 1)%R -> + ubeta_nat_pdf 1 1 x = 1%R. +Proof. +move=> x01. +by rewrite /ubeta_nat_pdf !expr0 mulr1 x01. +Qed. + +(* normalization constant *) +Definition beta_nat_norm {R : realType} (a b : nat) : R := + fine (\int[@lebesgue_measure R]_x (ubeta_nat_pdf a b x)%:E).*) + +Section beta_nat_Gamma. +Context {R : realType}. +(* +Let mu := @lebesgue_measure R. + +Let B (a b : nat) : \bar R := + \int[mu]_(x in `[0%R, 1%R]%classic) (ubeta_nat_pdf a b x)%:E. +*) +End beta_nat_Gamma. + +(* +Lemma integral_beta_nat_normTE {R : realType} (a b : nat) : + beta_nat_norm a b = + fine (\int[lebesgue_measure]_(t in `[0%R, 1%R]) (t^+a.-1 * (`1-t)^+b.-1)%:E) :> R. +Proof. +rewrite /beta_nat_norm /ubeta_nat_pdf [in RHS]integral_mkcond/=; congr fine. +by apply: eq_integral => /= x _; rewrite patchE mem_setE in_itv/=; case: ifPn. +Qed.*) + +Lemma onemXn_derivable {R : realType} n (x : R) : + derivable (fun y : R^o => `1-y ^+ n : R^o)%R x 1. +Proof. +have := @derivableX R R^o (@onem R) n x 1%R. +rewrite fctE. +apply. +exact: derivableB. +Qed. + +Lemma deriveX_idfun {R : realType} n x : + 'D_1 (@GRing.exp R^o ^~ n.+1) x = n.+1%:R *: (x ^+ n)%R. +Proof. by rewrite exp_derive /GRing.scale/= mulr1. Qed. + +Lemma derive1Mr [R : realFieldType] [f : R^o -> R^o] [x r : R^o] : + derivable f x 1 -> ((fun x => f x * r)^`() x = (r * f^`() x)%R :> R)%R. +Proof. +move=> fx1. +rewrite derive1E (deriveM fx1); last by []. +by rewrite -derive1E derive1_cst scaler0 add0r derive1E. +Qed. + +Lemma derive1Ml [R : realFieldType] [f : R^o -> R^o] [x r : R^o] : + derivable f x 1 -> ((fun x => r * f x)^`() x = (r * f^`() x)%R :> R)%R. +Proof. +under eq_fun do rewrite mulrC. +exact: derive1Mr. +Qed. + +Lemma decreasing_onem {R : numDomainType} : {homo (fun x : R => (1 - x)%R) : x y /~ (x < y)%R}. +Proof. +move=> b a ab. +by rewrite -ltrN2 !opprB ltr_leB. +Qed. + +Lemma continuous_onemXn {R : realType} (n : nat) x : + {for x, continuous (fun y : R => `1-y ^+ n)%R}. +Proof. +apply: (@continuous_comp _ _ _ (@onem R) (fun x => GRing.exp x n)). + by apply: (@cvgB _ R^o); [exact: cvg_cst|exact: cvg_id]. +exact: exprn_continuous. +Qed. + +Local Close Scope ereal_scope. + +(* we define a function to help formalizing the beta distribution *) +Section XMonemX. +Context {R : numDomainType}. + +Definition XMonemX a b := fun x : R => x ^+ a * `1-x ^+ b. + +Lemma XMonemX_ge0 a b x : x \in `[0%R, 1%R] -> 0 <= XMonemX a b x :> R. +Proof. +rewrite in_itv/= => /andP[x0 x1]. +by rewrite /XMonemX mulr_ge0// exprn_ge0// subr_ge0. +Qed. + +Lemma XMonemX0 n x : XMonemX 0 n x = `1-x ^+ n :> R. +Proof. by rewrite /XMonemX/= expr0 mul1r. Qed. + +Lemma XMonemX0' n x : XMonemX n 0 x = x ^+ n :> R. +Proof. by rewrite /XMonemX/= expr0 mulr1. Qed. + +Lemma XMonemX00 x : XMonemX 0 0 x = 1 :> R. +Proof. by rewrite XMonemX0 expr0. Qed. + +Lemma XMonemXC a b x : XMonemX a b (1 - x) = XMonemX b a x :> R. +Proof. +by rewrite /XMonemX [in LHS]/onem opprB addrCA subrr addr0 mulrC. +Qed. + +End XMonemX. + +Lemma continuous_XMonemX {R : realType} a b : + continuous (XMonemX a b : R -> R). +Proof. +by move=> x; apply: cvgM; [exact: exprn_continuous|exact: continuous_onemXn]. +Qed. + +Lemma within_continuous_XMonemX {R : realType} a b (A : set R) : + {within A, continuous (XMonemX a b : R -> R)}. +Proof. +by apply: continuous_in_subspaceT => x _; exact: continuous_XMonemX. +Qed. + +Lemma bounded_XMonemX {R : realType} (a b : nat) : + [bounded XMonemX a b x : R^o | x in (`[0, 1]%classic : set R)]. +Proof. +exists 1%R; split; [by rewrite num_real|move=> x x1 /= y y01]. +rewrite ger0_norm//; last by rewrite XMonemX_ge0. +move: y01; rewrite in_itv/= => /andP[? ?]. +rewrite (le_trans _ (ltW x1))// mulr_ile1 ?exprn_ge0//. +- by rewrite subr_ge0. +- by rewrite exprn_ile1. +- by rewrite exprn_ile1 ?subr_ge0// lerBlDl addrC -lerBlDl subrr. +Qed. + +Lemma measurable_fun_XMonemX {R : realType} a b (A : set R) : + measurable_fun A (XMonemX a b). +Proof. +apply/measurable_funM => //; apply/measurable_funX => //. +exact: measurable_funB. +Qed. + +Local Open Scope ereal_scope. + +Lemma integral_exprn {R : realType} (n : nat) : + fine (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (x ^+ n)%:E) = n.+1%:R^-1%R :> R. +Proof. +pose F (x : R) : R^o := (n.+1%:R^-1 * x ^+ n.+1)%R. +have cX m : {in `[0%R, 1%R], continuous (fun x : R => x ^+ m)%R}. + by move=> x x01; exact: exprn_continuous. +have cF0 : {for 0%R, continuous F}. + apply: continuousM; first exact: cvg_cst. + by apply: cX; rewrite /= in_itv/= lexx ler01. +have cF1 : {for 1%R, continuous F}. + apply: continuousM; first exact: cvg_cst. + by apply: cX; rewrite /= in_itv/= lexx ler01. +have dcF : derivable_oo_continuous_bnd F 0 1. + split. + - by move=> x x01; apply: derivableM => //; exact: exprn_derivable. + - apply: continuous_cvg; first exact: mulrl_continuous. + by apply/cvg_at_right_filter/cX; rewrite in_itv/= lexx ler01. + - apply: continuous_cvg; first exact: mulrl_continuous. + by apply/cvg_at_left_filter/cX; rewrite in_itv/= lexx ler01. +have dFE : {in `]0%R, 1%R[, F^`() =1 (fun x : R => x ^+ n : R)%R}. + move=> x x01. + rewrite derive1Ml; last exact: exprn_derivable. + by rewrite derive1E deriveX_idfun mulrA mulVf// mul1r. +rewrite (@continuous_FTC2 _ (fun x : R => x ^+ n)%R F)//. + by rewrite /F/= expr1n expr0n/= mulr1 mulr0 subr0. +by apply: continuous_subspaceT; exact: exprn_continuous. +Qed. + +Lemma derivable_oo_continuous_bnd_onemXnMr {R : realType} (n : nat) (r : R) : + derivable_oo_continuous_bnd (fun x : R => `1-x ^+ n.+1 * r : R^o)%R 0 1. +Proof. +split. +- by move=> x x01; apply: derivableM => //=; exact: onemXn_derivable. +- apply: cvgM; last exact: cvg_cst. + apply: cvg_at_right_filter. + apply: (@cvg_comp _ _ _ (fun x => 1 - x)%R (fun x => GRing.exp x n.+1)%R). + by apply: (@cvgB _ R^o); [exact: cvg_cst|exact: cvg_id]. + exact: exprn_continuous. +- apply: cvg_at_left_filter. + apply: cvgM; last exact: cvg_cst. + apply: (@cvg_comp _ _ _ (fun x => 1 - x)%R (fun x => GRing.exp x n.+1)%R). + by apply: (@cvgB _ R^o); [exact: cvg_cst|exact: cvg_id]. + exact: exprn_continuous. +Qed. + +Lemma derive_onemXn {R : realType} (n : nat) x : + ((fun y : R => `1-y ^+ n.+1 : R^o)^`() x = - n.+1%:R * `1-x ^+ n)%R. +Proof. +rewrite (@derive1_comp _ (@onem R) (fun x => GRing.exp x n.+1))//; last first. + exact: exprn_derivable. +rewrite derive1E deriveX_idfun derive1E deriveB//. +by rewrite -derive1E derive1_cst derive_id sub0r mulrN1 [in RHS]mulNr. +Qed. + +Lemma integral_onemXn {R : realType} (n : nat) : + fine (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (`1-x ^+ n)%:E) = n.+1%:R^-1%R :> R. +Proof. +rewrite (@continuous_FTC2 _ _ (fun x : R => ((1 - x) ^+ n.+1 / - n.+1%:R))%R)//=. +- rewrite subrr subr0 expr0n/= mul0r expr1n mul1r sub0r. + by rewrite -invrN -2!mulNrn opprK. +- apply: continuous_in_subspaceT. + by move=> x x01; exact: continuous_onemXn. +- exact: derivable_oo_continuous_bnd_onemXnMr. +- move=> x x01. + rewrite derive1Mr//; last exact: onemXn_derivable. + by rewrite derive_onemXn mulrA mulVf// mul1r. +Qed. + +Lemma integrable_XMonemX {R : realType} (a b : nat) : + lebesgue_measure.-integrable `[0%R, 1%R] (fun x : R => (XMonemX a b x)%:E). +Proof. +apply: continuous_compact_integrable => //. + exact: segment_compact. +apply: continuous_in_subspaceT => x _. +exact: continuous_XMonemX. +Qed. + +Lemma Rintegral_onemXn {R : realType} (n : nat) : + (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (`1-x ^+ n))%R = n.+1%:R^-1%R :> R. +Proof. +rewrite /Rintegral. +rewrite (@continuous_FTC2 _ _ (fun x : R => ((1 - x) ^+ n.+1 / - n.+1%:R))%R)//=. +- rewrite subrr subr0 expr0n/= mul0r expr1n mul1r sub0r. + by rewrite -invrN -2!mulNrn opprK. +- apply: continuous_in_subspaceT. +- by move=> x x01; exact: continuous_onemXn. +- exact: derivable_oo_continuous_bnd_onemXnMr. +- move=> x x01. + rewrite derive1Mr//; last exact: onemXn_derivable. + by rewrite derive_onemXn mulrA mulVf// mul1r. +Qed. + +(* TODO: move *) +Lemma normr_onem {R : realType} (x : R) : (0 <= x <= 1 -> `| `1-x | <= 1)%R. +Proof. +move=> /andP[x0 x1]; rewrite ler_norml; apply/andP; split. + by rewrite lerBrDl lerBlDr (le_trans x1)// lerDl. +by rewrite lerBlDr lerDl. +Qed. + +Local Open Scope ereal_scope. + +Local Open Scope ring_scope. + +Section XMonemX01. +Local Open Scope ring_scope. +Context {R : realType}. +Variables a b : nat. + +Definition XMonemX01 := (@XMonemX R a.-1 b.-1) \_ `[0, 1]. + +Lemma XMonemX01_ge0 t : 0 <= XMonemX01 t. +Proof. +rewrite /XMonemX01 patchE ; case: ifPn => //. +rewrite inE/= in_itv/= => /andP[t0 t1]. +by rewrite mulr_ge0// exprn_ge0// onem_ge0. +Qed. + +Lemma XMonemX01_le1 t : XMonemX01 t <= 1. +Proof. +rewrite /XMonemX01 patchE ; case: ifPn => //. +rewrite inE/= in_itv/= => /andP[t0 t1]. +by rewrite mulr_ile1// ?(exprn_ge0,onem_ge0,exprn_ile1,onem_le1). +Qed. + +Lemma measurable_XMonemX01 : measurable_fun [set: R] XMonemX01. +Proof. +rewrite /XMonemX01 /=; apply/(measurable_restrictT _ _).1 => //. +exact: measurable_fun_XMonemX. +Qed. + +Local Notation mu := lebesgue_measure. + +(* TODO: maybe not that useful *) +Lemma integral_XMonemX01 U : + (\int[mu]_(x in U) (XMonemX01 x)%:E = + \int[mu]_(x in U `&` `[0%R, 1%R]) (XMonemX01 x)%:E)%E. +Proof. +rewrite [RHS]integral_mkcondr /=; apply: eq_integral => x xU /=. +by rewrite /XMonemX01/= restrict_EFin -patch_setI setIid. +Qed. + +End XMonemX01. + +Lemma XMonemX_XMonemX01 {R : realType} a b a' b' (x : R) : (0 < a)%N -> (0 < b)%N -> + x \in `[0%R, 1%R]%classic -> + (XMonemX a' b' x * XMonemX01 a b x = XMonemX01 (a + a') (b + b') x :> R)%R. +Proof. +move=> a0 b0 x01; rewrite /XMonemX01 /= !patchE x01. +rewrite mulrCA -mulrA -exprD mulrA -exprD. +congr (_ ^+ _ * _ ^+ _)%R. + by rewrite addnC -!subn1 subDnCA. +by rewrite -!subn1 subDnCA. +Qed. + +Lemma XMonemX01_11 {R : realType} (x : R) : (0 <= x <= 1)%R -> + XMonemX01 1 1 x = 1%R. +Proof. +by move=> x01; rewrite /XMonemX01 patchE mem_setE in_itv/= x01/= XMonemX00. +Qed. + +(* normalization constant *) +Section betafun. +Context {R : realType}. +Notation mu := (@lebesgue_measure R). +Local Open Scope ring_scope. + +Definition betafun (a b : nat) : R := (\int[mu]_x (XMonemX01 a b x))%R. + +Lemma betafun0 (b : nat) : (0 < b)%N -> betafun 0 b = b%:R ^-1:> R. +Proof. +move=> b0. +rewrite -[LHS]Rintegral_mkcond. +under eq_Rintegral do rewrite XMonemX0. +by rewrite Rintegral_onemXn// prednK. +Qed. + +Lemma betafun00 : betafun 0 0 = 1%R :> R. +Proof. +rewrite -[LHS]Rintegral_mkcond. +under eq_Rintegral do rewrite XMonemX00. +rewrite /Rintegral. +rewrite integral_cst/= ?mul1e; last by exact: measurable_itv. +by rewrite lebesgue_measure_itv/= lte_fin ltr01 -EFinB subr0. +Qed. + +Lemma betafun_sym (a b : nat) : betafun a b = betafun b a :> R. +Proof. +rewrite -[LHS]Rintegral_mkcond. +rewrite Ronem_change//=; last 4 first. + by rewrite ltr01 lexx. + split. + - exact: measurable_fun_XMonemX. + - exact: openT. + - move=> K _ cK. + by have /integrableP[] := @continuous_compact_integrable R _ K cK + (@within_continuous_XMonemX R a.-1 b.-1 K). + apply: continuous_subspaceT. + by move=> x x01; exact: continuous_XMonemX. + move=> r. + exact: (@continuous_compact_integrable R _ _ + (@segment_compact _ _ _) (@within_continuous_XMonemX R a.-1 b.-1 `[0%R, r])). +rewrite subrr. +rewrite -[RHS]Rintegral_mkcond. +apply: eq_Rintegral => x x01. +by rewrite XMonemXC. +Qed. + +Lemma betafunS (a b : nat) : + (betafun a.+2 b.+1 = a.+1%:R / b.+1%:R * betafun a.+1 b.+2 :> R)%R. +Proof. +rewrite -[LHS]Rintegral_mkcond. +rewrite (@Rintegration_by_parts _ _ + (fun x => `1-x ^+ b.+1 / - b.+1%:R)%R (fun x => a.+1%:R * x ^+ a)%R); last 7 first. + exact: ltr01. + apply/continuous_subspaceT. + move=> x. + apply: cvgM; [exact: cvg_cst|]. + exact: exprn_continuous. + split. + by move=> x x01; exact: exprn_derivable. + by apply: cvg_at_right_filter; exact: exprn_continuous. + by apply: cvg_at_left_filter; exact: exprn_continuous. + by move=> x x01; rewrite derive1E deriveX_idfun. + apply/continuous_subspaceT. + by move=> x x01; exact: continuous_onemXn. + exact: derivable_oo_continuous_bnd_onemXnMr. + move=> x x01. + rewrite derive1Mr; last exact: onemXn_derivable. + by rewrite derive_onemXn mulrA mulVf// mul1r. +rewrite {1}/onem !(expr1n,mul1r,expr0n,subr0,subrr,mul0r,oppr0)/=. +rewrite sub0r. +transitivity (a.+1%:R / b.+1%:R * (\int[lebesgue_measure]_(x in `[0, 1]) (XMonemX a b.+1 x)) : R)%R. + under [in LHS]eq_Rintegral. + move=> x x01. + rewrite mulrA mulrC mulrA (mulrA _ a.+1%:R) -(mulrA (_ * _)%R). + over. + rewrite /=. + rewrite RintegralZl//=; last exact: integrable_XMonemX. + by rewrite -mulNrn -2!mulNr -invrN -mulNrn opprK (mulrC _ a.+1%:R)//=. +by rewrite Rintegral_mkcond. +Qed. + +Lemma betafunSS (a b : nat) : + (betafun a.+1 b.+1 = + a`!%:R / (\prod_(b.+1 <= i < (a + b).+1) i)%:R * betafun 1 (a + b).+1 :> R)%R. +Proof. +elim: a b => [b|a ih b]. + by rewrite fact0 mul1r add0n /index_iota subnn big_nil invr1 mul1r. +rewrite betafunS. +rewrite ih. +rewrite !mulrA. +congr *%R; last by rewrite addSnnS. +rewrite -mulrA. +rewrite mulrCA. +rewrite 2!mulrA. +rewrite -natrM. +rewrite (mulnC a`!). +rewrite -factS. +rewrite -mulrA. +rewrite -invfM. +rewrite big_add1. +rewrite [in RHS]big_nat_recl/=; last by rewrite addSn ltnS leq_addl. +by rewrite -natrM addSnnS. +Qed. + +Lemma betafun1S (n : nat) : (betafun 1 n.+1 = n.+1%:R^-1 :> R)%R. +Proof. +rewrite /betafun -Rintegral_mkcond. +under eq_Rintegral do rewrite XMonemX0. +by rewrite Rintegral_onemXn. +Qed. + +Lemma betafun_fact (a b : nat) : + (betafun a.+1 b.+1 = (a`! * b`!)%:R / (a + b).+1`!%:R :> R)%R. +Proof. +rewrite betafunSS betafun1S. +rewrite natrM -!mulrA; congr *%R. +(* (b+1 b+2 ... b+1 b+a)^-1 / (a+b+1) = b! / (a+b+1)! *) +rewrite factS. +rewrite [in RHS]mulnC. +rewrite natrM. +rewrite invfM. +rewrite mulrA; congr (_ / _). +rewrite -(@invrK _ b`!%:R) -invfM; congr (_^-1). +apply: (@mulfI _ b`!%:R). + by rewrite gt_eqF// ltr0n fact_gt0. +rewrite mulrA divff// ?gt_eqF// ?ltr0n ?fact_gt0//. +rewrite mul1r. +rewrite [in RHS]fact_prod. +rewrite -natrM; congr (_%:R). +rewrite fact_prod -big_cat/=. +by rewrite /index_iota subn1 -iotaD subn1/= subSS addnK addnC. +Qed. + +Lemma betafunE (a b : nat) : betafun a b = + if (a == 0)%N && (0 < b)%N then + b%:R^-1 + else if (b == 0)%N && (0 < a)%N then + a%:R^-1 + else + a.-1`!%:R * b.-1`!%:R / (a + b).-1`!%:R :> R. +Proof. +case: a => [|a]. + rewrite eqxx/=; case: ifPn => [|]. + by case: b => [|b _] //; rewrite betafun0. + rewrite -leqNgt leqn0 => /eqP ->. + by rewrite betafun00 eqxx ltnn/= fact0 mul1r divr1. +case: b => [|b]. + by rewrite betafun_sym betafun0// fact0 addn0/= mulr1 divff. +by rewrite betafun_fact/= natrM// addnS. +Qed. + +Lemma betafun_gt0 (a b : nat) : (0 < betafun a b :> R)%R. +Proof. +rewrite betafunE. +case: ifPn => [/andP[_ b0]|]; first by rewrite invr_gt0 ltr0n. +rewrite negb_and => /orP[a0|]. + case: ifPn => [/andP[_]|]; first by rewrite invr_gt0// ltr0n. + rewrite negb_and => /orP[b0|]. + by rewrite divr_gt0// ?mulr_gt0 ?ltr0n ?fact_gt0. + by rewrite -leqNgt leqn0 (negbTE a0). +rewrite -leqNgt leqn0 => /eqP ->. +rewrite eqxx/=. +case: ifPn; first by rewrite invr_gt0 ltr0n. +rewrite -leqNgt leqn0 => /eqP ->. +by rewrite fact0 mul1r divr1. +Qed. + +Lemma betafun_ge0 (a b : nat) : (0 <= betafun a b :> R)%R. +Proof. exact/ltW/betafun_gt0. Qed. + +Lemma betafun11 : betafun 1 1 = 1%R :> R. +Proof. by rewrite (betafun1S O) invr1. Qed. + +(* NB: this is not exactly betafun because EFin *) +Definition betafunEFin a b : \bar R := \int[mu]_x (XMonemX01 a b x)%:E. + +(* TODO: rev eq *) +Lemma betafunEFinT a b : + (betafunEFin a b = \int[mu]_(x in `[0%R, 1%R]) (XMonemX01 a b x)%:E)%E. +Proof. by rewrite /betafunEFin integral_XMonemX01/= setTI. Qed. + +Lemma betafunEFin_lty a b : (betafunEFin a b < +oo)%E. +Proof. +have := betafun_gt0 a b; rewrite /betafun /Rintegral /betafunEFin. +by case: (integral _ _ _) => [r _| |//]; rewrite ?ltxx ?ltry. +Qed. + +Lemma betafunEFin_fin_num a b : betafunEFin a b \is a fin_num. +Proof. +rewrite ge0_fin_numE ?betafunEFin_lty//. +by apply: integral_ge0 => //= x _; rewrite lee_fin XMonemX01_ge0. +Qed. + +Lemma betafunEFinE a b : (betafunEFin a b = (betafun a b)%:E :> \bar R)%E. +Proof. by rewrite -[LHS]fineK ?betafunEFin_fin_num. Qed. + +Lemma integrable_XMonemX01 a b : mu.-integrable setT (EFin \o XMonemX01 a b). +Proof. +apply/integrableP; split. + by apply/measurable_EFinP; exact: measurable_XMonemX01. +under eq_integral. + move=> /= x _. + rewrite ger0_norm//; last by rewrite XMonemX01_ge0. + over. +exact: betafunEFin_lty. +Qed. + +End betafun. + +(* normalized pdf for the beta distribution *) +Section beta_pdf. +Local Open Scope ring_scope. +Context {R : realType}. +Variables a b : nat. + +Definition beta_pdf t : R := XMonemX01 a b t / betafun a b. + +Lemma measurable_beta_pdf : measurable_fun setT beta_pdf. +Proof. by apply: measurable_funM => //; exact: measurable_XMonemX01. Qed. + +Lemma beta_pdf_ge0 t : 0 <= beta_pdf t. +Proof. +by rewrite /beta_pdf divr_ge0//; [exact: XMonemX01_ge0|exact: betafun_ge0]. +Qed. + +Lemma beta_pdf_le_betafunV x : beta_pdf x <= (betafun a b)^-1. +Proof. +rewrite /beta_pdf ler_pdivrMr ?betafun_gt0// mulVf ?gt_eqF ?betafun_gt0//. +exact: XMonemX01_le1. +Qed. + +Local Notation mu := lebesgue_measure. + +(* TODO: really useful? *) +Lemma int_beta_pdf01 : + (\int[mu]_(x in `[0%R, 1%R]) (beta_pdf x)%:E = + \int[mu]_x (beta_pdf x)%:E :> \bar R)%E. +Proof. +rewrite /beta_pdf. +under eq_integral do rewrite EFinM. +rewrite /=. +rewrite ge0_integralZr//=; last 3 first. + apply: measurable_funTS => /=; apply/measurable_EFinP => //. + exact: measurable_XMonemX01. + by move=> x _; rewrite lee_fin XMonemX01_ge0. + by rewrite lee_fin invr_ge0// betafun_ge0. +rewrite -betafunEFinT -ge0_integralZr//=. +- by apply/measurableT_comp => //; exact: measurable_XMonemX01. +- by move=> x _; rewrite lee_fin XMonemX01_ge0. +- by rewrite lee_fin invr_ge0// betafun_ge0. +Qed. + +Lemma integrable_beta_pdf : mu.-integrable setT (EFin \o beta_pdf). +Proof. +apply/integrableP; split. + by apply/measurable_EFinP; exact: measurable_beta_pdf. +under eq_integral. + move=> /= x _. + rewrite ger0_norm//; last by rewrite beta_pdf_ge0. + over. +rewrite -int_beta_pdf01. +apply: (@le_lt_trans _ _ (\int[mu]_(x in `[0%R, 1%R]) (betafun a b)^-1%:E)%E). + apply: ge0_le_integral => //=. + - by move=> x _; rewrite lee_fin beta_pdf_ge0. + - by apply/measurable_funTS/measurable_EFinP => /=; exact: measurable_beta_pdf. + - by move=> x _; rewrite lee_fin invr_ge0// betafun_ge0. + - by move=> x _; rewrite lee_fin beta_pdf_le_betafunV. +rewrite integral_cst//= lebesgue_measure_itv//=. +by rewrite lte01 oppr0 adde0 mule1 ltry. +Qed. + +Local Open Scope ring_scope. +Lemma bounded_beta_pdf_01 : + [bounded beta_pdf x : R^o | x in `[0%R, 1%R]%classic : set R]. +Proof. +exists (betafun a b)^-1; split; first by rewrite num_real. +move=> // y y1. +near=> M => /=. +rewrite (le_trans _ (ltW y1))//. +near: M. +move=> M /=. +rewrite in_itv/= => /andP[M0 M1]. +rewrite ler_norml; apply/andP; split. + rewrite lerNl (@le_trans _ _ 0%R)// ?invr_ge0 ?betafun_ge0//. + by rewrite lerNl oppr0 beta_pdf_ge0. +rewrite /beta_pdf ler_pdivrMr ?betafun_gt0//. +by rewrite mulVf ?XMonemX01_le1// gt_eqF// betafun_gt0. +Unshelve. all: by end_near. Qed. +Local Close Scope ring_scope. + +End beta_pdf. + +Section beta. +Local Open Scope ring_scope. +Context {R : realType}. +Variables a b : nat. + +Local Notation mu := (@lebesgue_measure R). + +Let beta_num (U : set (measurableTypeR R)) : \bar R := + \int[mu]_(x in U) (XMonemX01 a b x)%:E. + +Let beta_numT : beta_num setT = betafunEFin a b. +Proof. by rewrite /beta_num/= -/(betafunEFin a b) betafunEFinE. Qed. + +Let beta_num_lty U : measurable U -> (beta_num U < +oo)%E. +Proof. +move=> mU. +apply: (@le_lt_trans _ _ (\int[mu]_(x in U `&` `[0%R, 1%R]%classic) 1)%E); last first. + rewrite integral_cst//= ?mul1e//. + rewrite (le_lt_trans (measureIr _ _ _))//= lebesgue_measure_itv//= lte01//. + by rewrite EFinN sube0 ltry. + exact: measurableI. +rewrite /beta_num integral_XMonemX01 ge0_le_integral//=. +- exact: measurableI. +- by move=> x _; rewrite lee_fin XMonemX01_ge0. +- by apply/measurable_funTS/measurableT_comp => //; exact: measurable_XMonemX01. +- by move=> x _; rewrite lee_fin XMonemX01_le1. +Qed. + +Let beta_num0 : beta_num set0 = 0%:E. +Proof. by rewrite /beta_num integral_set0. Qed. + +Let beta_num_ge0 U : (0 <= beta_num U)%E. +Proof. +by rewrite /beta_num integral_ge0//= => x Ux; rewrite lee_fin XMonemX01_ge0. +Qed. + +Let beta_num_sigma_additive : semi_sigma_additive beta_num. +Proof. +move=> /= F mF tF mUF; rewrite /beta_num; apply: cvg_toP. + apply: ereal_nondecreasing_is_cvgn => m n mn. + apply: lee_sum_nneg_natr => // k _ _; apply: integral_ge0 => /= x Fkx. + by rewrite lee_fin; exact: XMonemX01_ge0. +rewrite ge0_integral_bigcup//=. +- by apply/measurable_funTS/measurableT_comp => //; exact: measurable_XMonemX01. +- by move=> x _; rewrite lee_fin XMonemX01_ge0. +Qed. + +HB.instance Definition _ := isMeasure.Build _ _ _ beta_num + beta_num0 beta_num_ge0 beta_num_sigma_additive. + +Definition beta_prob := + @mscale _ _ _ (invr_nonneg (NngNum (betafun_ge0 a b))) beta_num. + +HB.instance Definition _ := Measure.on beta_prob. + +Let beta_prob_setT : beta_prob setT = 1%:E. +Proof. +rewrite /beta_prob /= /mscale /= beta_numT betafunEFinE//. +by rewrite -EFinM mulVf// gt_eqF// betafun_gt0. Qed. HB.instance Definition _ := - Kernel_isFinite.Build _ _ _ Z R (mkswap k) mkswap_finite. + @Measure_isProbability.Build _ _ _ beta_prob beta_prob_setT. + +Lemma beta_prob01 : beta_prob `[0, 1] = 1%:E. +Proof. +rewrite /beta_prob /= /mscale/= /beta_num -betafunEFinT betafunEFinE//. +by rewrite -EFinM mulVf// gt_eqF// betafun_gt0. +Qed. -End kswap_finite_kernel_finite. +Lemma beta_prob_fin_num U : measurable U -> beta_prob U \is a fin_num. +Proof. +move=> mU; rewrite ge0_fin_numE//. +rewrite /beta_prob/= /mscale/= /beta_num lte_mul_pinfty//. + by rewrite lee_fin// invr_ge0 betafun_ge0. +apply: (@le_lt_trans _ _ (betafunEFin a b)). + apply: ge0_subset_integral => //=. + by apply/measurable_EFinP; exact: measurable_XMonemX01. + by move=> x _; rewrite lee_fin XMonemX01_ge0. +by rewrite betafunEFin_lty. +Qed. -Notation "l .; k" := (mkcomp l (mkswap k)) : ereal_scope. +Lemma integral_beta_pdf U : measurable U -> + (\int[mu]_(x in U) (beta_pdf a b x)%:E = beta_prob U :> \bar R)%E. +Proof. +move=> mU. +rewrite /beta_pdf. +under eq_integral do rewrite EFinM/=. +rewrite ge0_integralZr//=. +- by rewrite /beta_prob/= /mscale/= muleC. +- by apply/measurable_funTS/measurableT_comp => //; exact: measurable_XMonemX01. +- by move=> x _; rewrite lee_fin XMonemX01_ge0. +- by rewrite lee_fin invr_ge0// betafun_ge0. +Qed. -Section letin'. -Variables (d d' d3 : _) (X : measurableType d) (Y : measurableType d') - (Z : measurableType d3) (R : realType). +End beta. +Arguments beta_prob {R}. -Definition letin' (l : R.-sfker X ~> Y) (k : R.-sfker Y * X ~> Z) := - locked [the R.-sfker X ~> Z of l .; k]. +Lemma integral_beta_prob_bernoulli_lty {R : realType} a b (f : R -> R) U : + measurable_fun setT f -> + (forall x, x \in `[0%R, 1%R] -> 0 <= f x <= 1)%R -> + (\int[beta_prob a b]_x `|bernoulli (f x) U| < +oo :> \bar R)%E. +Proof. +move=> mf /= f01. +apply: (@le_lt_trans _ _ (\int[beta_prob a b]_x cst 1 x))%E. + apply: ge0_le_integral => //=. + apply: measurableT_comp => //=. + by apply: (measurableT_comp (measurable_bernoulli2 _)). + by move=> x _; rewrite gee0_abs// probability_le1. +by rewrite integral_cst//= mul1e -ge0_fin_numE// beta_prob_fin_num. +Qed. -Lemma letin'E (l : R.-sfker X ~> Y) (k : R.-sfker Y * X ~> Z) x U : - letin' l k x U = \int[l x]_y k (y, x) U. -Proof. by rewrite /letin'; unlock. Qed. +Lemma integral_beta_prob_bernoulli_onemX_lty {R : realType} n a b U : + (\int[beta_prob a b]_x `|bernoulli (`1-x ^+ n) U| < +oo :> \bar R)%E. +Proof. +apply: integral_beta_prob_bernoulli_lty => //=. + by apply: measurable_funX => //; exact: measurable_funB. +move=> x; rewrite in_itv/= => /andP[x0 x1]. +rewrite exprn_ge0 ?subr_ge0//= exprn_ile1// ?subr_ge0//. +by rewrite lerBlDl -lerBlDr subrr. +Qed. -Lemma letin'_letin (l : R.-sfker X ~> Y) (k : R.-sfker Y * X ~> Z) : - letin' l k = letin l (mkswap k). -Proof. by rewrite /letin'; unlock. Qed. +Lemma integral_beta_prob_bernoulli_onem_lty {R : realType} n a b U : + (\int[beta_prob a b]_x `|bernoulli (1 - `1-x ^+ n) U| < +oo :> \bar R)%E. +Proof. +apply: integral_beta_prob_bernoulli_lty => //=. + apply: measurable_funB => //. + by apply: measurable_funX => //; exact: measurable_funB. +move=> x; rewrite in_itv/= => /andP[x0 x1]. +rewrite -lerBlDr opprK add0r. +rewrite andbC lerBlDl -lerBlDr subrr. +rewrite exprn_ge0 ?subr_ge0//= exprn_ile1// ?subr_ge0//. +by rewrite lerBlDl -lerBlDr subrr. +Qed. -End letin'. +Local Open Scope ring_scope. -Section letin'C. -Import Notations. -Context d d1 d' (X : measurableType d) (Y : measurableType d1) - (Z : measurableType d') (R : realType). -Variables (t : R.-sfker Z ~> X) - (u' : R.-sfker X * Z ~> Y) - (u : R.-sfker Z ~> Y) - (t' : R.-sfker Y * Z ~> X) - (tt' : forall y, t =1 fun z => t' (y, z)) - (uu' : forall x, u =1 fun z => u' (x, z)). - -Definition T' z : set X -> \bar R := t z. -Let T0 z : (T' z) set0 = 0. Proof. by []. Qed. -Let T_ge0 z x : 0 <= (T' z) x. Proof. by []. Qed. -Let T_semi_sigma_additive z : semi_sigma_additive (T' z). -Proof. exact: measure_semi_sigma_additive. Qed. -HB.instance Definition _ z := @isMeasure.Build _ X R (T' z) (T0 z) (T_ge0 z) - (@T_semi_sigma_additive z). - -Let sfinT z : sfinite_measure (T' z). Proof. exact: sfinite_kernel_measure. Qed. -HB.instance Definition _ z := @isSFinite.Build _ X R (T' z) (sfinT z). - -Definition U' z : set Y -> \bar R := u z. -Let U0 z : (U' z) set0 = 0. Proof. by []. Qed. -Let U_ge0 z x : 0 <= (U' z) x. Proof. by []. Qed. -Let U_semi_sigma_additive z : semi_sigma_additive (U' z). -Proof. exact: measure_semi_sigma_additive. Qed. -HB.instance Definition _ z := @isMeasure.Build _ Y R (U' z) (U0 z) (U_ge0 z) - (@U_semi_sigma_additive z). - -Let sfinU z : sfinite_measure (U' z). Proof. exact: sfinite_kernel_measure. Qed. -HB.instance Definition _ z := @isSFinite.Build _ Y R (U' z) (sfinU z). - -Lemma letin'C z A : measurable A -> - letin' t - (letin' u' - (ret (measurable_fun_prod macc1of3' macc0of3'))) z A = - letin' u - (letin' t' - (ret (measurable_fun_prod macc0of3' macc1of3'))) z A. -Proof. -move=> mA. -rewrite !letin'E. +Lemma beta_prob_uniform {R : realType} : beta_prob 1 1 = uniform_prob (@ltr01 R). +Proof. +apply/funext => U. +rewrite /beta_prob /uniform_prob. +rewrite /mscale/= betafun11 invr1 !mul1e. +rewrite integral_XMonemX01 integral_uniform_pdf. under eq_integral. - move=> x _. - rewrite letin'E -uu'. - under eq_integral do rewrite retE /=. + move=> /= x. + rewrite inE => -[Ux/=]; rewrite in_itv/= => x10. + rewrite XMonemX01_11//=. over. -rewrite (sfinite_Fubini (T' z) (U' z) (fun x => \d_(x.1, x.2) A ))//; last first. - apply/measurable_EFinP => /=; rewrite (_ : (fun x => _) = mindic R mA)//. - by apply/funext => -[]. rewrite /=. -apply: eq_integral => y _. -by rewrite letin'E/= -tt'; apply: eq_integral => // x _; rewrite retE. +under [RHS]eq_integral. + move=> /= x. + rewrite inE => -[Ux/=]; rewrite in_itv/= => x10. + rewrite /uniform_pdf x10 subr0 invr1. + over. +by []. Qed. -End letin'C. -Arguments letin'C {d d1 d' X Y Z R} _ _ _ _. +Lemma beta_prob_integrable {R :realType} a b a' b' : + (beta_prob a b).-integrable `[0, 1] + (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (XMonemX a' b' x)%:E). +Proof. +apply/integrableP; split. + by apply/measurableT_comp => //; exact: measurable_fun_XMonemX. +apply: (@le_lt_trans _ _ (\int[beta_prob a b]_(x in `[0%R, 1%R]) 1)%E). + apply: ge0_le_integral => //=. + do 2 apply/measurableT_comp => //. + exact: measurable_fun_XMonemX. + move=> x; rewrite in_itv/= => /andP[x0 x1]. + rewrite lee_fin ger0_norm; last first. + by rewrite !mulr_ge0// exprn_ge0// onem_ge0. + by rewrite mulr_ile1// ?exprn_ge0 ?onem_ge0// exprn_ile1// ?onem_ge0// onem_le1. +rewrite integral_cst//= mul1e. +by rewrite -ge0_fin_numE// beta_prob_fin_num. +Qed. -Section letin'A. -Context d d' d1 d2 d3 (X : measurableType d) (Y : measurableType d') - (T1 : measurableType d1) (T2 : measurableType d2) (T3 : measurableType d3) - (R : realType). -Import Notations. -Variables (t : R.-sfker X ~> T1) - (u : R.-sfker T1 * X ~> T2) - (v : R.-sfker T2 * X ~> Y) - (v' : R.-sfker T2 * (T1 * X) ~> Y) - (vv' : forall y, v =1 fun xz => v' (xz.1, (y, xz.2))). - -Lemma letin'A x A : measurable A -> - letin' t (letin' u v') x A - = - (letin' (letin' t u) v) x A. -Proof. -move=> mA. -rewrite !letin'E. -under eq_integral do rewrite letin'E. -rewrite letin'_letin/=. -rewrite integral_kcomp; [|by []|]. - apply: eq_integral => z _. - apply: eq_integral => y _. - by rewrite (vv' z). -exact: measurableT_comp (@measurable_kernel _ _ _ _ _ v _ mA) _. -Qed. - -End letin'A. +Lemma beta_prob_integrable_onem {R : realType} a b a' b' : + (beta_prob a b).-integrable `[0, 1] + (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (`1-(XMonemX a' b' x))%:E). +Proof. +apply: (eq_integrable _ (cst 1 \- (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => + (XMonemX a' b' x)%:E))%E) => //. +apply: (@integrableB _ (g_sigma_algebraType R.-ocitv.-measurable)) => //=. + (* TODO: lemma? *) + apply/integrableP; split => //. + rewrite (eq_integral (fun x => (\1_setT x)%:E))/=; last first. + by move=> x _; rewrite /= indicT normr1. + rewrite integral_indic//= setTI /beta_prob /mscale/= lte_mul_pinfty//. + by rewrite lee_fin invr_ge0 betafun_ge0. + have /integrableP[_] := @integrable_XMonemX01 R a b. + under eq_integral. + move=> x _. + rewrite gee0_abs//; last first. + by rewrite lee_fin XMonemX01_ge0. + over. + by rewrite /= -/(betafunEFin a b) /= betafunEFinT. +exact: beta_prob_integrable. +Qed. + +Lemma beta_prob_integrable_dirac {R : realType} a b a' b' (c : bool) U : + (beta_prob a b).-integrable `[0, 1] + (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (XMonemX a' b' x)%:E * \d_c U)%E. +Proof. +apply: integrableMl => //=; last first. + exists 1; split => // x x1/= _ _; rewrite (le_trans _ (ltW x1))//. + by rewrite ger0_norm// indicE; case: (_ \in _). +exact: beta_prob_integrable. +Qed. + +Lemma beta_prob_integrable_onem_dirac {R : realType} a b a' b' (c : bool) U : + (beta_prob a b).-integrable `[0, 1] + (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (`1-(XMonemX a' b' x))%:E * \d_c U)%E. +Proof. +apply: integrableMl => //=; last first. + exists 1; split => // x x1/= _ _; rewrite (le_trans _ (ltW x1))//. + by rewrite ger0_norm// indicE; case: (_ \in _). +exact: beta_prob_integrable_onem. +Qed. + +Local Close Scope ring_scope. + +Section integral_beta. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +Lemma beta_prob_dom a b : @beta_prob R a b `<< mu. +Proof. +move=> A mA muA0; rewrite /beta_prob /mscale/=. +apply/eqP; rewrite mule_eq0 eqe invr_eq0 gt_eqF/= ?betafun_gt0//; apply/eqP. +rewrite integral_XMonemX01; apply/eqP; rewrite eq_le; apply/andP; split; last first. + by apply: integral_ge0 => x _; rewrite lee_fin XMonemX01_ge0. +apply: (@le_trans _ _ (\int[mu]_(x in A `&` `[0%R, 1%R]%classic) 1)); last first. + rewrite integral_cst ?mul1e//=; last exact: measurableI. + by rewrite -[leRHS]muA0 measureIl. +apply: ge0_le_integral => //=; first exact: measurableI. +- by move=> x _; rewrite lee_fin XMonemX01_ge0. +- apply/measurable_funTS/measurableT_comp => //. + exact: measurable_XMonemX01. +- by move=> x _; rewrite lee_fin XMonemX01_le1. +Qed. + +Lemma integral_beta_prob a b f U : measurable U -> measurable_fun U f -> + \int[beta_prob a b]_(x in U) `|f x| < +oo -> + \int[beta_prob a b]_(x in U) f x = \int[mu]_(x in U) (f x * (beta_pdf a b x)%:E) :> \bar R. +Proof. +move=> mU mf finf. +rewrite -(Radon_Nikodym_change_of_variables (beta_prob_dom a b)) //=; last first. + by apply/integrableP; split. +apply: ae_eq_integral => //. +- apply: emeasurable_funM => //; apply: measurable_int. + apply: integrableS (Radon_Nikodym_integrable _) => //=. + exact: beta_prob_dom. +- apply: emeasurable_funM => //=; apply/measurableT_comp => //=. + by apply/measurable_funTS; exact: measurable_beta_pdf. +- apply: ae_eqe_mul2l => /=. + rewrite Radon_NikodymE//=; first exact: beta_prob_dom. + move=> ?. + case: cid => /= h [h1 h2 h3]. + apply: integral_ae_eq => //. + + apply: integrableS h2 => //. (* integrableST? *) + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_beta_pdf. + + by move=> E E01 mE; rewrite -h3//= integral_beta_pdf. +Qed. + +End integral_beta. + +Section beta_prob_bernoulliE. +Context {R : realType}. +Local Notation mu := lebesgue_measure. +Local Open Scope ring_scope. + +Definition div_betafun a b c d : R := betafun (a + c) (b + d) / betafun a b. + +Lemma div_betafun_ge0 a b c d : 0 <= div_betafun a b c d. +Proof. by rewrite /div_betafun divr_ge0// betafun_ge0. Qed. + +Lemma div_betafun_le1 a b c d : (0 < a)%N -> (0 < b)%N -> + div_betafun a b c d <= 1. +Proof. +move=> a0 b0. +rewrite /div_betafun ler_pdivrMr// ?mul1r ?betafun_gt0//. +rewrite !betafunE. +rewrite addn_eq0 (gtn_eqF a0)/=. +rewrite addn_eq0 (gtn_eqF b0)/=. +rewrite ler_pdivrMr ?ltr0n ?fact_gt0//. +rewrite mulrAC. +rewrite ler_pdivlMr ?ltr0n ?fact_gt0//. +rewrite -!natrM ler_nat. +case: a a0 => //. +move=> n. +rewrite addSn. +case: b b0 => //. +move=> m. +rewrite [(n + c).+1.-1]/=. +rewrite [n.+1.-1]/=. +rewrite [m.+1.-1]/=. +rewrite addnS. +rewrite [(_ + m).+1.-1]/=. +rewrite (addSn m d). +rewrite [(m + _).+1.-1]/=. +rewrite (addSn (n + c)). +rewrite [_.+1.-1]/=. +rewrite addSn addnS. +by rewrite leq_fact2// leq_addr. +Qed. + +Definition beta_prob_bernoulli a b c d U : \bar R := + \int[beta_prob a b]_(y in `[0, 1]) bernoulli (XMonemX01 c.+1 d.+1 y) U. + +Lemma beta_prob_bernoulliE a b c d U : (a > 0)%N -> (b > 0)%N -> + beta_prob_bernoulli a b c d U = bernoulli (div_betafun a b c d) U. +Proof. +move=> a0 b0. +rewrite /beta_prob_bernoulli. +under eq_integral => x. + rewrite inE/= in_itv/= => x01. + rewrite bernoulliE/= ?XMonemX01_ge0 ?XMonemX01_le1//. + over. +rewrite /=. +rewrite [in RHS]bernoulliE/= ?div_betafun_ge0 ?div_betafun_le1//=. +under eq_integral => x x01. + rewrite /XMonemX01 patchE x01/=. + over. +rewrite /=. +rewrite integralD//=; last 2 first. + exact: beta_prob_integrable_dirac. + exact: beta_prob_integrable_onem_dirac. +congr (_ + _). + rewrite integralZr//=; last exact: beta_prob_integrable. + congr (_ * _)%E. + rewrite integral_beta_prob//; last 2 first. + by apply/measurableT_comp => //; exact: measurable_fun_XMonemX. + by have /integrableP[_] := @beta_prob_integrable R a b c d. + rewrite /beta_pdf. + under eq_integral do rewrite EFinM -muleA muleC -muleA. + rewrite /=. + transitivity ((betafun a b)^-1%:E * + \int[mu]_(x in `[0%R, 1%R]) (XMonemX01 (a + c) (b + d) x)%:E : \bar R)%E. + rewrite -integralZl//=; last first. + apply/integrableP; split. + apply/measurable_EFinP/measurable_funTS. + exact: measurable_XMonemX01. + under eq_integral. + move=> x x01. + rewrite gee0_abs; last by rewrite lee_fin XMonemX01_ge0. + over. + by rewrite /= -betafunEFinT betafunEFin_lty. + apply: eq_integral => x x01. + (* TODO: lemma? property of XMonemX? *) + rewrite muleA muleC muleA -(EFinM (x ^+ c)). + rewrite -/(XMonemX c d x) -EFinM mulrA XMonemX_XMonemX01//. + by rewrite -EFinM mulrC. + by rewrite -betafunEFinT betafunEFinE -EFinM mulrC. +under eq_integral do rewrite muleC. +rewrite /=. +rewrite integralZl//=; last exact: beta_prob_integrable_onem. +rewrite muleC; congr (_ * _)%E. +rewrite integral_beta_prob//=; last 2 first. + apply/measurableT_comp => //=. + by apply/measurable_funB => //; exact: measurable_fun_XMonemX. + by have /integrableP[] := @beta_prob_integrable_onem R a b c d. +rewrite /beta_pdf. +under eq_integral do rewrite EFinM muleA. +rewrite integralZr//=; last first. + apply: integrableMr => //=. + - by apply/measurable_funB => //=; exact: measurable_fun_XMonemX. + - apply/ex_bound => //. + + apply: (@globally_properfilter _ _ 0%R) => //=. + by apply: inferP; rewrite in_itv/= lexx ler01. + + exists 1 => t. + rewrite /= in_itv/= => t01. + apply: normr_onem; apply/andP; split. + by rewrite mulr_ge0// exprn_ge0// ?onem_ge0//; case/andP: t01. + by rewrite mulr_ile1// ?exprn_ge0 ?exprn_ile1// ?onem_ge0 ?onem_le1//; case/andP: t01. + - exact: integrableS (integrable_XMonemX01 _ _). +transitivity ( + (\int[mu]_(x in `[0%R, 1%R]) + ((XMonemX01 a b x)%:E - (XMonemX01 (a + c) (b + d) x)%:E) : \bar R) + * (betafun a b)^-1%:E)%E. + congr (_ * _)%E. + apply: eq_integral => x x01. + rewrite /onem -EFinM mulrBl mul1r EFinB; congr (_ - _)%E. + by rewrite XMonemX_XMonemX01. +rewrite integralB_EFin//=; last 2 first. + exact: integrableS (integrable_XMonemX01 _ _). + exact: integrableS (integrable_XMonemX01 _ _). +rewrite -!betafunEFinT !betafunEFinE. +rewrite -EFinM. +rewrite mulrBl /onem mulfV; last by rewrite gt_eqF// betafun_gt0. +by []. +Qed. + +End beta_prob_bernoulliE. Declare Scope lang_scope. Delimit Scope lang_scope with P. @@ -295,7 +1403,7 @@ Definition mtyp t : measurableType (mtyp_disp t) := projT2 (measurable_of_typ t). Definition measurable_of_seq (l : seq typ) : {d & measurableType d} := - iter_mprod (map measurable_of_typ l). + iter_mprod (List.map measurable_of_typ l). End syntax_of_types. Arguments measurable_of_typ {R}. @@ -379,8 +1487,6 @@ match b with | binop_mult => Real end. -(* Import Notations. *) - Definition fun_of_binop g (b : binop) : (mctx g -> mtyp (type_of_binop b)) -> (mctx g -> mtyp (type_of_binop b)) -> @mctx R g -> @mtyp R (type_of_binop b) := match b with @@ -443,16 +1549,10 @@ Inductive exp : flag -> ctx -> typ -> Type := | exp_proj1 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t1 | exp_proj2 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t2 | exp_var g str t : t = lookup Unit g str -> exp D g t -| exp_bernoulli g (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : - exp D g (Prob Bool) -| exp_bernoulli_trunc g : - exp D g Real -> exp D g (Prob Bool) -| exp_binomial g (n : nat) (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : - exp D g (Prob Nat) -| exp_binomial_trunc g (n : nat) : - exp D g Real -> exp D g (Prob Nat) -| exp_uniform g (a b : R) (ab0 : (0 < b - a)%R) : exp D g (Prob Real) -| exp_beta g (a b : nat) (* NB: shound be R *) : exp D g (Prob Real) +| exp_bernoulli g : exp D g Real -> exp D g (Prob Bool) +| exp_binomial g (n : nat) : exp D g Real -> exp D g (Prob Nat) +| exp_uniform g (a b : R) (ab : (a < b)%R) : exp D g (Prob Real) +| exp_beta g (a b : nat) (* NB: should be R *) : exp D g (Prob Real) | exp_poisson g : nat -> exp D g Real -> exp D g Real | exp_normalize g t : exp P g t -> exp D g (Prob t) | exp_letin g t1 t2 str : exp P g t1 -> exp P ((str, t1) :: g) t2 -> @@ -484,12 +1584,10 @@ Arguments exp_bin {R} b {g} &. Arguments exp_rel {R} r {g} &. Arguments exp_pair {R g} & {t1 t2}. Arguments exp_var {R g} _ {t} & H. -Arguments exp_bernoulli {R g}. -Arguments exp_bernoulli_trunc {R g} &. -Arguments exp_binomial {R g}. +Arguments exp_bernoulli {R g} &. +Arguments exp_binomial {R g} &. Arguments exp_uniform {R g} &. Arguments exp_beta {R g} &. -Arguments exp_binomial_trunc {R g} &. Arguments exp_poisson {R g}. Arguments exp_normalize {R g _}. Arguments exp_letin {R g} & {_ _}. @@ -574,12 +1672,10 @@ Fixpoint free_vars k g t (e : @exp R k g t) : seq string := | exp_proj1 _ _ _ e => free_vars e | exp_proj2 _ _ _ e => free_vars e | exp_var _ x _ _ => [:: x] - | exp_bernoulli _ _ _ => [::] - | exp_bernoulli_trunc _ e => free_vars e - | exp_binomial _ _ _ _ => [::] + | exp_bernoulli _ e => free_vars e + | exp_binomial _ _ e => free_vars e | exp_uniform _ _ _ _ => [::] | exp_beta _ _ _ => [::] - | exp_binomial_trunc _ _ e => free_vars e | exp_poisson _ _ e => free_vars e | exp_normalize _ _ e => free_vars e | exp_letin _ _ _ x e1 e2 => free_vars e1 ++ rem x (free_vars e2) @@ -595,9 +1691,11 @@ End free_vars. Definition dval R g t := @mctx R g -> @mtyp R t. Definition pval R g t := R.-sfker @mctx R g ~> @mtyp R t. + Section weak. Context {R : realType}. Implicit Types (g h : ctx) (x : string * typ). +Local Open Scope ring_scope. Fixpoint mctx_strong g h x (f : @mctx R (g ++ x :: h)) : @mctx R (g ++ h) := match g as g0 return mctx (g0 ++ x :: h) -> mctx (g0 ++ h) with @@ -612,7 +1710,7 @@ Lemma measurable_fun_mctx_strong g h x : measurable_fun setT (@mctx_strong g h x). Proof. elim: g h x => [h x|x g ih h x0]; first exact: measurable_snd. -apply/prod_measurable_funP; split. +apply/measurable_fun_pairP; split. - rewrite [X in measurable_fun _ X](_ : _ = fst)//. by apply/funext => -[]. - rewrite [X in measurable_fun _ X](_ : _ = @mctx_strong g h x0 \o snd). @@ -685,18 +1783,6 @@ Implicit Type (g : ctx) (str : string). Local Open Scope lang_scope. Local Open Scope ring_scope. -Definition bernoulli0 := @bernoulli R 0%R%:nng ler01. - -HB.instance Definition _ := Probability.on bernoulli0. - -Lemma __ : Measurable.sort - (pprobability - [the measurableType (R.-ocitv.-measurable).-sigma of - g_sigma_algebraType (R.-ocitv.-measurable)] R) = -Measurable.sort (@mtyp R (Prob Real)). -rewrite /=. -(* done. *) -Abort. Inductive evalD : forall g t, exp D g t -> forall f : dval R g t, measurable_fun setT f -> Prop := @@ -708,7 +1794,7 @@ Inductive evalD : forall g t, exp D g t -> | eval_real g r : ([r:R] : exp D g _) -D> cst r ; kr r -| eval_pow g n (e : exp D g _) f mf : e -D> f ; mf -> +| eval_pow g n (e : exp D g _) f mf : e -D> f ; mf -> [e ^+ {n}] -D> (fun x => f x ^+ n) ; (measurable_funX n mf) | eval_bin g bop (e1 : exp D g _) f1 mf1 e2 f2 mf2 : @@ -721,7 +1807,7 @@ Inductive evalD : forall g t, exp D g t -> | eval_pair g t1 (e1 : exp D g t1) f1 mf1 t2 (e2 : exp D g t2) f2 mf2 : e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> - [(e1, e2)] -D> fun x => (f1 x, f2 x) ; measurable_fun_prod mf1 mf2 + [(e1, e2)] -D> fun x => (f1 x, f2 x) ; measurable_fun_pair mf1 mf2 | eval_proj1 g t1 t2 (e : exp D g (Pair t1 t2)) f mf : e -D> f ; mf -> @@ -737,35 +1823,25 @@ Inductive evalD : forall g t, exp D g t -> | eval_var g x H : let i := index x (dom g) in exp_var x H -D> acc_typ (map snd g) i ; measurable_acc_typ (map snd g) i -| eval_bernoulli g r r1 : - (exp_bernoulli r r1 : exp D g _) -D> cst (bernoulli r1) ; - measurable_cst _ - -| eval_bernoulli_trunc g e r mr : - e -D> r ; mr -> - (exp_bernoulli_trunc e : exp D g _) -D> bernoulli_trunc \o r ; - measurableT_comp measurable_bernoulli_trunc mr - -| eval_binomial g n (p : {nonneg R}) (p1 : (p%:num <= 1)%R) : - (exp_binomial n p p1 : exp D g _) -D> cst (binomial_probability n p1) ; - measurable_cst _ +| eval_bernoulli g e r mr : + e -D> r ; mr -> (exp_bernoulli e : exp D g _) -D> bernoulli \o r ; + measurableT_comp measurable_bernoulli mr -| eval_binomial_trunc g n e r mr : - e -D> r ; mr -> - (exp_binomial_trunc n e : exp D g _) -D> binomial_probability_trunc n \o r ; - measurableT_comp (measurable_binomial_probability_trunc n) mr +| eval_binomial g n e r mr : + e -D> r ; mr -> (exp_binomial n e : exp D g _) -D> binomial_prob n \o r ; + measurableT_comp (measurable_binomial_prob n) mr -| eval_uniform g (a b : R) (ab0 : (0 < b - a)%R) : - (exp_uniform a b ab0 : exp D g _) -D> cst (uniform_probability ab0) ; - measurable_cst _ +| eval_uniform g (a b : R) (ab : (a < b)%R) : + (exp_uniform a b ab : exp D g _) -D> cst (uniform_prob ab) ; + measurable_cst _ | eval_beta g (a b : nat) : - (exp_beta a b : exp D g _) -D> cst (beta_nat a b) ; measurable_cst _ + (exp_beta a b : exp D g _) -D> cst (beta_prob a b) ; measurable_cst _ | eval_poisson g n (e : exp D g _) f mf : e -D> f ; mf -> - exp_poisson n e -D> poisson n \o f ; - measurableT_comp (measurable_poisson n) mf + exp_poisson n e -D> poisson_pdf n \o f ; + measurableT_comp (measurable_poisson_pdf n) mf | eval_normalize g t (e : exp P g t) k : e -P> k -> @@ -886,29 +1962,20 @@ all: (rewrite {g t e u v mu mv hu}). inversion 1; subst g0. inj_ex H9; rewrite -H9. by inj_ex H10. -- move=> g r r1 {}v {}mv. - inversion 1; subst g0 r0. - inj_ex H3; subst v. - by have -> : r1 = r3 by []. - move=> g e r mr ev IH {}v {}mv. inversion 1; subst g0. inj_ex H0; subst e0. - inj_ex H4; subst v. - by rewrite (IH _ _ H2). -- move=> g n p p1 {}v {}mv. - inversion 1; subst g0 n0 p0. - inj_ex H4; subst v. - by have -> : p1 = p3 by []. + inj_ex H3; subst v. + by rewrite (IH _ _ H4). - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. inj_ex H2; subst e0. inj_ex H4; subst v. - inj_ex H5; subst mv. - by rewrite (IH _ _ H3). -- move=> g a b ab0 {}v {}mv. + by rewrite (IH _ _ H5). +- move=> g a b ab {}v {}mv. inversion 1; subst g0 a0 b0. inj_ex H4; subst v. - by have -> : ab0 = ab2. + by have -> : ab = ab1. - (* TODO: beta *) move=> g a b {}v {}mv. inversion 1; subst g0 a0 b0. by inj_ex H4; subst v. @@ -1044,29 +2111,20 @@ all: rewrite {g t e u v eu}. inversion 1; subst g0. inj_ex H9; rewrite -H9. by inj_ex H10. -- move=> g r r1 {}v {}mv. - inversion 1; subst g0 r0. - inj_ex H3; subst v. - by have -> : r1 = r3 by []. - move=> g e r mr ev IH {}v {}mv. inversion 1; subst g0. inj_ex H0; subst e0. - inj_ex H4; subst v. - by rewrite (IH _ _ H2). -- move=> g n p p1 {}v {}mv. - inversion 1; subst g0 n0 p0. - inj_ex H4; subst v. - by have -> : p1 = p3 by []. + inj_ex H3; subst v. + by rewrite (IH _ _ H4). - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. inj_ex H2; subst e0. inj_ex H4; subst v. - inj_ex H5; subst mv. - by rewrite (IH _ _ H3). -- move=> g a b ab0 {}v {}mv. + by rewrite (IH _ _ H5). +- move=> g a b ab {}v {}mv. inversion 1; subst g0 a0 b0. inj_ex H4; subst v. - by have -> : ab0 = ab2. + by have -> : ab = ab1. - (* TODO: beta case*) move=> g a b {}v {}mv. inversion 1; subst g0 a0 b0. by inj_ex H4; subst v. @@ -1163,17 +2221,16 @@ all: rewrite {z g t}. - move=> g t1 t2 e [f [mf H]]. by exists (snd \o f); eexists; exact: eval_proj2. - by move=> g x t tE; subst t; eexists; eexists; exact: eval_var. -- by move=> r r1; eexists; eexists; exact: eval_bernoulli. - move=> g e [p [mp H]]. - by exists (bernoulli_trunc \o p); eexists; exact: eval_bernoulli_trunc. -- by move=> p p1; eexists; eexists; exact: eval_binomial. + exists ((bernoulli : R -> pprobability bool R) \o p). + by eexists; exact: eval_bernoulli. - move=> g n e [p [mp H]]. - exists (binomial_probability_trunc n \o p). - eexists; exact: (eval_binomial_trunc n). + exists ((binomial_prob n : R -> pprobability nat R) \o p). + by eexists; exact: (eval_binomial n). - by eexists; eexists; exact: eval_uniform. - by eexists; eexists; exact: eval_beta. - move=> g h e [f [mf H]]. - by exists (poisson h \o f); eexists; exact: eval_poisson. + by exists (poisson_pdf h \o f); eexists; exact: eval_poisson. - move=> g t e [k ek]. by exists (normalize_pt k); eexists; exact: eval_normalize. - move=> g t1 t2 x e1 [k1 ev1] e2 [k2 ev2]. @@ -1307,7 +2364,7 @@ Lemma execD_pair g t1 t2 (e1 : exp D g t1) (e2 : exp D g t2) : let mf1 := projT2 (execD e1) in let mf2 := projT2 (execD e2) in execD [(e1, e2)] = @existT _ _ (fun z => (f1 z, f2 z)) - (@measurable_fun_prod _ _ _ (mctx g) (mtyp t1) (mtyp t2) + (@measurable_fun_pair _ _ _ (mctx g) (mtyp t1) (mtyp t2) f1 f2 mf1 mf2). Proof. by move=> f1 f2 mf1 mf2; apply/execD_evalD/eval_pair; exact: evalD_execD. @@ -1341,35 +2398,26 @@ Lemma execD_var g x (H : nth Unit (map snd g) (index x (dom g)) = lookup Unit g (measurable_acc_typ (map snd g) i). Proof. by move=> i; apply/execD_evalD; exact: eval_var. Qed. -Lemma execD_bernoulli g r (r1 : (r%:num <= 1)%R) : - @execD g _ (exp_bernoulli r r1) = - existT _ (cst [the probability _ _ of bernoulli r1]) (measurable_cst _). -Proof. exact/execD_evalD/eval_bernoulli. Qed. - -Lemma execD_bernoulli_trunc g e : - @execD g _ (exp_bernoulli_trunc e) = - existT _ (bernoulli_trunc \o projT1 (execD e)) (measurableT_comp measurable_bernoulli_trunc (projT2 (execD e))). -Proof. exact/execD_evalD/eval_bernoulli_trunc/evalD_execD. Qed. - -Lemma execD_binomial g n p (p1 : (p%:num <= 1)%R) : - @execD g _ (exp_binomial n p p1) = - existT _ (cst [the probability _ _ of binomial_probability n p1]) (measurable_cst _). -Proof. exact/execD_evalD/eval_binomial. Qed. +Lemma execD_bernoulli g e : + @execD g _ (exp_bernoulli e) = + existT _ ((bernoulli : R -> pprobability bool R) \o projT1 (execD e)) + (measurableT_comp measurable_bernoulli (projT2 (execD e))). +Proof. exact/execD_evalD/eval_bernoulli/evalD_execD. Qed. -Lemma execD_binomial_trunc g n e : - @execD g _ (exp_binomial_trunc n e) = - existT _ (binomial_probability_trunc n \o projT1 (execD e)) - (measurableT_comp (measurable_binomial_probability_trunc n) (projT2 (execD e))). -Proof. exact/execD_evalD/eval_binomial_trunc/evalD_execD. Qed. +Lemma execD_binomial g n e : + @execD g _ (exp_binomial n e) = + existT _ ((binomial_prob n : R -> pprobability nat R) \o projT1 (execD e)) + (measurableT_comp (measurable_binomial_prob n) (projT2 (execD e))). +Proof. exact/execD_evalD/eval_binomial/evalD_execD. Qed. Lemma execD_uniform g a b ab0 : @execD g _ (exp_uniform a b ab0) = - existT _ (cst [the probability _ _ of uniform_probability ab0]) (measurable_cst _). + existT _ (cst [the probability _ _ of uniform_prob ab0]) (measurable_cst _). Proof. exact/execD_evalD/eval_uniform. Qed. -Lemma execD_beta_nat g a b : +Lemma execD_beta g a b : @execD g _ (exp_beta a b) = - existT _ (cst [the probability _ _ of beta_nat a b]) (measurable_cst _). + existT _ (cst [the probability _ _ of beta_prob a b]) (measurable_cst _). Proof. exact/execD_evalD/eval_beta. Qed. Lemma execD_normalize_pt g t (e : exp P g t) : @@ -1380,8 +2428,8 @@ Proof. exact/execD_evalD/eval_normalize/evalP_execP. Qed. Lemma execD_poisson g n (e : exp D g Real) : execD (exp_poisson n e) = - existT _ (poisson n \o (projT1 (execD e))) - (measurableT_comp (measurable_poisson n) (projT2 (execD e))). + existT _ (poisson_pdf n \o projT1 (execD e)) + (measurableT_comp (measurable_poisson_pdf n) (projT2 (execD e))). Proof. exact/execD_evalD/eval_poisson/evalD_execD. Qed. Lemma execP_if g st e1 e2 e3 : @@ -1417,3 +2465,32 @@ End execution_functions. Arguments execD_var_erefl {R g} str. Arguments execP_weak {R} g h x {t} e. Arguments exp_var'E {R} str. + +Local Open Scope lang_scope. +Lemma congr_letinl {R : realType} g t1 t2 str (e1 e2 : @exp _ _ g t1) + (e : @exp _ _ (_ :: g) t2) x U : + (forall y V, execP e1 y V = execP e2 y V) -> + measurable U -> + @execP R g t2 [let str := e1 in e] x U = + @execP R g t2 [let str := e2 in e] x U. +Proof. by move=> + mU; move/eq_sfkernel => He; rewrite !execP_letin He. Qed. + +Lemma congr_letinr {R : realType} g t1 t2 str (e : @exp _ _ _ t1) + (e1 e2 : @exp _ _ (_ :: g) t2) x U : + (forall y V, execP e1 (y, x) V = execP e2 (y, x) V) -> + @execP R g t2 [let str := e in e1] x U = @execP R g t2 [let str := e in e2] x U. +Proof. +by move=> He; rewrite !execP_letin !letin'E; apply: eq_integral => ? _; exact: He. +Qed. + +Lemma congr_normalize {R : realType} g t (e1 e2 : @exp R _ g t) : + (forall x U, execP e1 x U = execP e2 x U) -> + execD [Normalize e1] = execD [Normalize e2]. +Proof. +move=> He; apply: eq_execD. +rewrite !execD_normalize_pt /=. +f_equal. +apply: eq_kernel => y V. +exact: He. +Qed. +Local Close Scope lang_scope. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index 77e7a78dda..57335e3176 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -2,33 +2,45 @@ Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. From mathcomp Require Import interval_inference. -From mathcomp Require Import mathcomp_extra boolp classical_sets. +From mathcomp Require Import unstable mathcomp_extra boolp classical_sets. From mathcomp Require Import functions cardinality fsbigop. From mathcomp Require Import reals ereal topology normedtype sequences esum. From mathcomp Require Import measure lebesgue_measure numfun lebesgue_integral. From mathcomp Require Import kernel prob_lang lang_syntax_util lang_syntax. +From mathcomp Require Import probability. From mathcomp Require Import ring lra. -(******************************************************************************) -(* Examples using the Probabilistic Programming Language of lang_syntax.v *) +(**md**************************************************************************) +(* # Examples using the Probabilistic Programming Language of lang_syntax.v *) (* *) -(* sample_pair_syntax := normalize ( *) +(* sample_pair1213 := normalize ( *) (* let x := sample (bernoulli 1/2) in *) (* let y := sample (bernoulli 1/3) in *) (* return (x, y)) *) (* *) +(* sample_and1213 := normalize ( *) +(* let x := sample (bernoulli 1/2) in *) +(* let y := sample (bernoulli 1/3) in *) +(* return (x && y)) *) +(* *) (* bernoulli13_score := normalize ( *) (* let x := sample (bernoulli 1/3) in *) (* let _ := if x then score (1/3) else score (2/3) in *) (* return x) *) (* *) -(* bernoulli12_score := normalize ( *) -(* let x := sample (bernoulli 1/2) in *) -(* let _ := if x then score (1/3) else score (2/3) in *) -(* return x) *) +(* sample_binomial3 := *) +(* let x := sample (binomial 3 1/2) in *) +(* return x *) (* *) (* hard_constraint := let x := Score {0}:R in return TT *) (* *) +(* guard := *) +(* let p := sample (bernoulli (1 / 3)) in *) +(* let _ := if p then return TT else score 0 in *) +(* return p *) +(* *) +(* more examples about uniform, beta, and bernoulli distributions *) +(* *) (* associativity of let-in expressions *) (* *) (* staton_bus_syntax == example from [Staton, ESOP 2017] *) @@ -52,90 +64,8 @@ Local Open Scope ring_scope. Local Open Scope ereal_scope. Local Open Scope string_scope. -(* letin' versions of rewriting laws *) -Lemma letin'_sample_bernoulli d d' (T : measurableType d) - (T' : measurableType d') (R : realType)(r : {nonneg R}) (r1 : (r%:num <= 1)%R) - (u : R.-sfker bool * T ~> T') x y : - letin' (sample_cst (bernoulli r1)) u x y = - r%:num%:E * u (true, x) y + (`1- (r%:num))%:E * u (false, x) y. -Proof. -rewrite letin'E/=. -rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. -by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT 2!mul1e. -Qed. - -Section letin'_return. -Context d d' d3 (X : measurableType d) (Y : measurableType d') - (Z : measurableType d3) (R : realType). - -Lemma letin'_kret (k : R.-sfker X ~> Y) - (f : Y * X -> Z) (mf : measurable_fun setT f) x U : - measurable U -> - letin' k (ret mf) x U = k x (curry f ^~ x @^-1` U). -Proof. -move=> mU; rewrite letin'E. -under eq_integral do rewrite retE. -rewrite integral_indic ?setIT// -[X in measurable X]setTI. -exact: (measurableT_comp mf). -Qed. - -Lemma letin'_retk (f : X -> Y) (mf : measurable_fun setT f) - (k : R.-sfker Y * X ~> Z) x U : - measurable U -> letin' (ret mf) k x U = k (f x, x) U. -Proof. -move=> mU; rewrite letin'E retE integral_dirac ?diracT ?mul1e//. -exact: (measurableT_comp (measurable_kernel k _ mU)). -Qed. - -End letin'_return. - -Section letin'_ite. -Context d d2 d3 (T : measurableType d) (T2 : measurableType d2) - (Z : measurableType d3) (R : realType). -Variables (k1 k2 : R.-sfker T ~> Z) - (u : R.-sfker Z * T ~> T2) - (f : T -> bool) (mf : measurable_fun setT f) - (t : T) (U : set T2). - -Lemma letin'_iteT : f t -> letin' (ite mf k1 k2) u t U = letin' k1 u t U. -Proof. -move=> ftT. -rewrite !letin'E/=. -apply: eq_measure_integral => V mV _. -by rewrite iteE ftT. -Qed. - -Lemma letin'_iteF : ~~ f t -> letin' (ite mf k1 k2) u t U = letin' k2 u t U. -Proof. -move=> ftF. -rewrite !letin'E/=. -apply: eq_measure_integral => V mV _. -by rewrite iteE (negbTE ftF). -Qed. - -End letin'_ite. -(* /letin' versions of rewriting laws *) - Local Open Scope lang_scope. -Lemma execP_letinL {R : realType} g t1 t2 x (e1 : @exp R P g t1) (e1' : exp P g t1) - (e2 : exp P ((x, t1) :: g) t2) : - forall U, measurable U -> - execP e1 = execP e1' -> - execP [let x := e1 in e2] ^~ U = execP [let x := e1' in e2] ^~ U. -Proof. -by move=> U mU e1e1'; rewrite !execP_letin e1e1'. -Qed. - -Lemma execP_letinR {R : realType} g t1 t2 x (e1 : @exp R P g t1) - (e2 : exp P _ t2) (e2' : exp P ((x, t1) :: g) t2) : - forall U, measurable U -> - execP e2 = execP e2' -> - execP [let x := e1 in e2] ^~ U = execP [let x := e1 in e2'] ^~ U. -Proof. -by move=> U mU e1e1'; rewrite !execP_letin e1e1'. -Qed. - Local Close Scope lang_scope. (* simple tests to check bidirectional hints *) @@ -190,7 +120,7 @@ Context {R : realType}. Lemma exec_normalize_return g x r : projT1 (@execD _ g _ [Normalize return r:R]) x = @dirac _ (measurableTypeR R) r _ :> probability _ R. - (* TODO: try to use the notation \d_r *) + (* NB: \d_r notation? *) Proof. by rewrite execD_normalize_pt execP_return execD_real//=; exact: normalize_kdirac. Qed. @@ -203,16 +133,15 @@ Local Open Scope ring_scope. Import Notations. Context {R : realType}. -Definition sample_pair_syntax0 : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in - let "y" := Sample {exp_bernoulli (1 / 3%:R)%:nng (p1S 2)} in +Definition sample_pair1213' : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli [{1 / 2}:R]} in + let "y" := Sample {exp_bernoulli [{1 / 3}:R]} in return (#{"x"}, #{"y"})]. -Definition sample_pair_syntax : exp _ [::] _ := - [Normalize {sample_pair_syntax0}]. +Definition sample_pair1213 : exp _ [::] _ := [Normalize {sample_pair1213'}]. -Lemma exec_sample_pair0 (A : set (bool * bool)) : - @execP R [::] _ sample_pair_syntax0 tt A = +Lemma exec_sample_pair1213' (A : set (bool * bool)) : + @execP R [::] _ sample_pair1213' tt A = ((1 / 2)%:E * ((1 / 3)%:E * \d_(true, true) A + (1 - 1 / 3)%:E * \d_(true, false) A) + @@ -220,110 +149,116 @@ Lemma exec_sample_pair0 (A : set (bool * bool)) : ((1 / 3)%:E * \d_(false, true) A + (1 - 1 / 3)%:E * \d_(false, false) A))%E. Proof. -rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. +rewrite !execP_letin !execP_sample !execD_bernoulli !execP_return /=. rewrite execD_pair !exp_var'E (execD_var_erefl "x") (execD_var_erefl "y") /=. -rewrite letin'E ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !diracT !mul1e. -rewrite !letin'E !ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. -by rewrite !integral_dirac//= !diracT !mul1e. +rewrite !execD_real//=. +do 2 (rewrite letin'E/= integral_bernoulli//=; last lra). +by rewrite letin'E/= integral_bernoulli//=; lra. Qed. -Lemma exec_sample_pair0_TandT : - @execP R [::] _ sample_pair_syntax0 tt [set (true, true)] = (1 / 6)%:E. +Lemma exec_sample_pair1213'_TandT : + @execP R [::] _ sample_pair1213' tt [set (true, true)] = (1 / 6)%:E. Proof. -rewrite exec_sample_pair0 !diracE mem_set//; do 3 rewrite memNset//=. +rewrite exec_sample_pair1213' !diracE mem_set//; do 3 rewrite memNset//=. by rewrite /= !mule0 mule1 !add0e mule0 adde0; congr (_%:E); lra. Qed. -Lemma exec_sample_pair0_TandF : - @execP R [::] _ sample_pair_syntax0 tt [set (true, false)] = (1 / 3)%:E. +Lemma exec_sample_pair1213'_TandT' : + @execP R [::] _ sample_pair1213' tt [set p | p.1 && p.2] = (1 / 6)%:E. Proof. -rewrite exec_sample_pair0 !diracE memNset// mem_set//; do 2 rewrite memNset//. +rewrite exec_sample_pair1213' !diracE mem_set//; do 3 rewrite memNset//=. by rewrite /= !mule0 mule1 !add0e mule0 adde0; congr (_%:E); lra. Qed. -Lemma exec_sample_pair0_TandT' : - @execP R [::] _ sample_pair_syntax0 tt [set p | p.1 && p.2] = (1 / 6)%:E. +Lemma exec_sample_pair1213'_TandF : + @execP R [::] _ sample_pair1213' tt [set (true, false)] = (1 / 3)%:E. Proof. -rewrite exec_sample_pair0 !diracE mem_set//; do 3 rewrite memNset//=. +rewrite exec_sample_pair1213' !diracE memNset// mem_set//; do 2 rewrite memNset//. by rewrite /= !mule0 mule1 !add0e mule0 adde0; congr (_%:E); lra. Qed. -Lemma exec_sample_pair_TorT : - (projT1 (execD sample_pair_syntax)) tt [set p | p.1 || p.2] = (2 / 3)%:E. +Lemma exec_sample_pair1213_TorT : + (projT1 (execD sample_pair1213)) tt [set p | p.1 || p.2] = (2 / 3)%:E. Proof. -rewrite execD_normalize_pt normalizeE/= exec_sample_pair0. +rewrite execD_normalize_pt normalizeE/= exec_sample_pair1213'. rewrite !diracE; do 4 rewrite mem_set//=. rewrite eqe ifF; last by apply/negbTE/negP => /orP[/eqP|//]; lra. -rewrite exec_sample_pair0 !diracE; do 3 rewrite mem_set//; rewrite memNset//=. +rewrite exec_sample_pair1213' !diracE; do 3 rewrite mem_set//; rewrite memNset//=. by rewrite !mule1; congr (_%:E); field. Qed. -Definition sample_and_syntax0 : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in - let "y" := Sample {exp_bernoulli (1 / 3%:R)%:nng (p1S 2)} in +End sample_pair. + +Section sample_and. +Local Open Scope lang_scope. +Local Open Scope ring_scope. +Import Notations. +Context {R : realType}. + +Definition sample_and1213' : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli [{1 / 2}:R]} in + let "y" := Sample {exp_bernoulli [{1 / 3}:R]} in return #{"x"} && #{"y"}]. -Lemma exec_sample_and0 (A : set bool) : - @execP R [::] _ sample_and_syntax0 tt A = ((1 / 6)%:E * \d_true A + - (1 - 1 / 6)%:E * \d_false A)%E. -Proof. -rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. -rewrite (@execD_bin _ _ binop_and) !exp_var'E (execD_var_erefl "x") (execD_var_erefl "y") /=. -rewrite letin'E ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !diracT !mul1e. -rewrite !letin'E !ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !diracT !mul1e. +Lemma exec_sample_and1213' (A : set bool) : + @execP R [::] _ sample_and1213' tt A = ((1 / 6)%:E * \d_true A + + (1 - 1 / 6)%:E * \d_false A)%E. +Proof. +rewrite !execP_letin !execP_sample/= !execD_bernoulli execP_return /=. +rewrite !(@execD_bin _ _ binop_and) !exp_var'E (execD_var_erefl "x"). +rewrite (execD_var_erefl "y") /= !letin'E/= !execD_real/=. +rewrite integral_bernoulli//=; last lra. +rewrite !letin'E/= integral_bernoulli//=; last lra. +rewrite integral_bernoulli//=; last lra. +rewrite /onem. rewrite muleDr// -addeA; congr (_ + _)%E. by rewrite !muleA; congr (_%:E); congr (_ * _); field. rewrite -muleDl// !muleA -muleDl//. by congr (_%:E); congr (_ * _); field. Qed. -Definition sample_bernoulli_and3 : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in - let "y" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in - let "z" := Sample {exp_bernoulli (1 / 2)%:nng (p1S 1)} in +Definition sample_and121212 : @exp R _ [::] _ := + [let "x" := Sample {exp_bernoulli [{1 / 2}:R]} in + let "y" := Sample {exp_bernoulli [{1 / 2}:R]} in + let "z" := Sample {exp_bernoulli [{1 / 2}:R]} in return #{"x"} && #{"y"} && #{"z"}]. -Lemma exec_sample_bernoulli_and3 t U : - execP sample_bernoulli_and3 t U = ((1 / 8)%:E * \d_true U + - (1 - 1 / 8)%:E * \d_false U)%E. -Proof. -rewrite !execP_letin !execP_sample !execD_bernoulli execP_return /=. -rewrite !(@execD_bin _ _ binop_and) !exp_var'E. -rewrite (execD_var_erefl "x") (execD_var_erefl "y") (execD_var_erefl "z") /=. -rewrite letin'E ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !diracT !mul1e. -rewrite !letin'E !ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !diracT !mul1e. -rewrite !letin'E !ge0_integral_measure_add//= !ge0_integral_mscale//= /onem. -rewrite !integral_dirac//= !diracT !mul1e. -rewrite !muleDr// -!addeA. -by congr (_ + _)%E; rewrite ?addeA !muleA -?muleDl//; -congr (_ * _)%E; congr (_%:E); field. +Lemma exec_sample_and121212 t U : + execP sample_and121212 t U = ((1 / 8)%:E * \d_true U + + (1 - 1 / 8)%:E * \d_false U)%E. +Proof. +rewrite !execP_letin !execP_sample !execD_bernoulli !execP_return /=. +rewrite !(@execD_bin _ _ binop_and) !exp_var'E (execD_var_erefl "x"). +rewrite (execD_var_erefl "y") (execD_var_erefl "z") /= !execD_real/=. +do 3 (rewrite !letin'E/= integral_bernoulli//=; last lra). +do 2 (rewrite integral_bernoulli//=; last lra). +rewrite !letin'E/= integral_bernoulli//=; last lra. +rewrite !muleDr// -!addeA; congr (_ + _)%E. + by rewrite !muleA; congr *%E; congr EFin; field. +rewrite !muleA -!muleDl//; congr *%E; congr EFin. +by rewrite /onem; field. Qed. -End sample_pair. +End sample_and. -Section bernoulli_examples. +Section sample_score. Local Open Scope ring_scope. Local Open Scope lang_scope. Import Notations. Context {R : realType}. Definition bernoulli13_score := [Normalize - let "x" := Sample {@exp_bernoulli R [::] (1 / 3%:R)%:nng (p1S 2)} in - let "_" := if #{"x"} then Score {(1 / 3)}:R else Score {(2 / 3)}:R in + let "x" := Sample {@exp_bernoulli R [::] [{1 / 3}:R]} in + let "_" := if #{"x"} then Score {1 / 3}:R else Score {2 / 3}:R in return #{"x"}]. Lemma exec_bernoulli13_score : - execD bernoulli13_score = execD (exp_bernoulli (1 / 5%:R)%:nng (p1S 4)). + execD bernoulli13_score = execD (exp_bernoulli [{1 / 5}:R]). Proof. apply: eq_execD. rewrite execD_bernoulli/= /bernoulli13_score execD_normalize_pt 2!execP_letin. rewrite execP_sample/= execD_bernoulli/= execP_if /= exp_var'E. -rewrite (execD_var_erefl "x")/= !execP_return/= 2!execP_score 2!execD_real/=. +rewrite (execD_var_erefl "x")/= !execP_return/= 2!execP_score !execD_real/=. apply: funext=> g; apply: eq_probability => U. rewrite normalizeE !letin'E/=. under eq_integral. @@ -331,41 +266,40 @@ under eq_integral. rewrite !letin'E. under eq_integral do rewrite retE /=. over. -rewrite !ge0_integral_measure_add //=; last by move=> b _; rewrite integral_ge0. -rewrite !ge0_integral_mscale //=; last 2 first. - by move=> b _; rewrite integral_ge0. - by move=> b _; rewrite integral_ge0. -rewrite !integral_dirac// !diracT !mul1e. +rewrite /=. +rewrite integral_bernoulli//=; [|lra|by move=> b; rewrite integral_ge0]. rewrite iteE/= !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_indic//= !iteE/= /mscale/=. rewrite setTI !diracT !mule1. rewrite ger0_norm//. rewrite -EFinD/= eqe ifF; last first. - apply/negbTE/negP => /orP[/eqP|//]; rewrite /onem. lra. + by apply/negbTE/negP => /orP[/eqP|//]; rewrite /onem; lra. +rewrite integral_bernoulli//=; last lra. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_dirac//= !diracT !mul1e ger0_norm//. rewrite exp_var'E (execD_var_erefl "x")/=. -rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. +rewrite !indicT/= !mulr1. +rewrite bernoulliE//=; last lra. by rewrite muleDl//; congr (_ + _)%E; rewrite -!EFinM; congr (_%:E); - rewrite !indicT !indicE /onem /=; case: (_ \in _); field. + rewrite !indicE /onem /=; case: (_ \in _); field. Qed. Definition bernoulli12_score := [Normalize - let "x" := Sample {@exp_bernoulli R [::] (1 / 2)%:nng (p1S 1)} in - let "r" := if #{"x"} then Score {(1 / 3)}:R else Score {(2 / 3)}:R in + let "x" := Sample {@exp_bernoulli R [::] [{1 / 2}:R]} in + let "r" := if #{"x"} then Score {1 / 3}:R else Score {2 / 3}:R in return #{"x"}]. Lemma exec_bernoulli12_score : - execD bernoulli12_score = execD (exp_bernoulli (1 / 3%:R)%:nng (p1S 2)). + execD bernoulli12_score = execD (exp_bernoulli [{1 / 3}:R]). Proof. apply: eq_execD. rewrite execD_bernoulli/= /bernoulli12_score execD_normalize_pt 2!execP_letin. rewrite execP_sample/= execD_bernoulli/= execP_if /= exp_var'E. -rewrite (execD_var_erefl "x")/= !execP_return/= 2!execP_score 2!execD_real/=. +rewrite (execD_var_erefl "x")/= !execP_return/= 2!execP_score !execD_real/=. apply: funext=> g; apply: eq_probability => U. rewrite normalizeE !letin'E/=. under eq_integral. @@ -373,11 +307,7 @@ under eq_integral. rewrite !letin'E. under eq_integral do rewrite retE /=. over. -rewrite !ge0_integral_measure_add //=; last by move=> b _; rewrite integral_ge0. -rewrite !ge0_integral_mscale //=; last 2 first. - by move=> b _; rewrite integral_ge0. - by move=> b _; rewrite integral_ge0. -rewrite !integral_dirac// !diracT !mul1e. +rewrite /= integral_bernoulli//=; [|lra|by move=> b; rewrite integral_ge0]. rewrite iteE/= !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_indic//= !iteE/= /mscale/=. @@ -386,13 +316,15 @@ rewrite ger0_norm//. rewrite -EFinD/= eqe ifF; last first. apply/negbTE/negP => /orP[/eqP|//]. by rewrite /onem; lra. +rewrite integral_bernoulli//=; last lra. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_dirac//= !diracT !mul1e ger0_norm//. rewrite exp_var'E (execD_var_erefl "x")/=. -rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. -by rewrite muleDl//; congr (_ + _)%E; +rewrite bernoulliE//=; last lra. +rewrite !mul1r. +rewrite muleDl//; congr (_ + _)%E; rewrite -!EFinM; congr (_%:E); by rewrite !indicT !indicE /onem /=; case: (_ \in _); field. @@ -400,15 +332,12 @@ Qed. (* https://dl.acm.org/doi/pdf/10.1145/2933575.2935313 (Sect. 4) *) Definition bernoulli14_score := [Normalize - let "x" := Sample {@exp_bernoulli R [::] (1 / 4%:R)%:nng (p1S 3)} in + let "x" := Sample {@exp_bernoulli R [::] [{1 / 4}:R]} in let "r" := if #{"x"} then Score {5}:R else Score {2}:R in return #{"x"}]. -Let p511 : ((5%:R / 11%:R)%:nng%:num <= (1 : R)). -Proof. by rewrite /=; lra. Qed. - Lemma exec_bernoulli14_score : - execD bernoulli14_score = execD (exp_bernoulli (5%:R / 11%:R)%:nng p511). + execD bernoulli14_score = execD (exp_bernoulli [{5%:R / 11%:R}:R]). Proof. apply: eq_execD. rewrite execD_bernoulli/= execD_normalize_pt 2!execP_letin. @@ -422,41 +351,39 @@ under eq_integral. rewrite !letin'E. under eq_integral do rewrite retE /=. over. -rewrite !ge0_integral_measure_add //=; last by move=> b _; rewrite integral_ge0. -rewrite !ge0_integral_mscale //=; last 2 first. - by move=> b _; rewrite integral_ge0. - by move=> b _; rewrite integral_ge0. -rewrite !integral_dirac// !diracT !mul1e. +rewrite /= integral_bernoulli//=; [|lra|by move=> b; rewrite integral_ge0]. rewrite iteE/= !ge0_integral_mscale//=. rewrite ger0_norm//. -rewrite !integral_dirac//= !diracT !mule1. -rewrite iteE/= !ge0_integral_mscale//=. +rewrite !integral_cst//= !diracT !(mule1,mul1e). +rewrite !indicT/= !mule1. +rewrite !iteE/= /mscale/=. rewrite ger0_norm//. -rewrite !integral_dirac//= !diracT !mule1. +rewrite !diracT/= !mul1r. rewrite -EFinD/= eqe ifF; last first. apply/negbTE/negP => /orP[/eqP|//]. by rewrite /onem; lra. +rewrite integral_bernoulli//=; last lra. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. -rewrite !integral_dirac//= !diracT /= !mul1e ger0_norm//. -rewrite /bernoulli/= measure_addE/= /mscale/= !mul1r. -rewrite !indicT !mulr1. -by rewrite muleDl//; congr (_ + _)%E; +rewrite !integral_dirac//= !diracT !mul1e ger0_norm//. +rewrite bernoulliE//=; last lra. +rewrite !indicT. +rewrite muleDl//; congr (_ + _)%E; rewrite -!EFinM; congr (_%:E); by rewrite !indicE /onem /=; case: (_ \in _); field. Qed. -End bernoulli_examples. +End sample_score. -Section binomial_examples. +Section sample_binomial. Context {R : realType}. Open Scope lang_scope. Open Scope ring_scope. Definition sample_binomial3 : @exp R _ [::] _ := - [let "x" := Sample {exp_binomial 3 (1 / 2)%:nng (p1S 1)} in + [let "x" := Sample {exp_binomial 3 [{1 / 2}:R]} in return #{"x"}]. Lemma exec_sample_binomial3 t U : measurable U -> @@ -466,46 +393,19 @@ Lemma exec_sample_binomial3 t U : measurable U -> (1 / 8)%:E * \d_3%N U)%E. Proof. move=> mU; rewrite /sample_binomial3 execP_letin execP_sample execP_return. -rewrite exp_var'E (execD_var_erefl "x") !execD_binomial/=. -rewrite letin'E ge0_integral_measure_sum//=. -rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//=; [|exact: measurable_fun_dirac..]. -rewrite !integral_dirac// /bump; [|exact: measurable_fun_dirac..]. +rewrite exp_var'E (execD_var_erefl "x") !execD_binomial/= execD_real//=. +rewrite letin'E/= /= integral_binomial//=; [lra|move=> _]. +rewrite !big_ord_recl big_ord0/=. +rewrite /bump. rewrite !binS/= !bin0 bin1 bin2 bin_small// addn0. rewrite expr0 mulr1 mul1r subn0. rewrite -2!addeA !mul1r. congr _%:E. -rewrite indicT !mul1r /onem !addrA addr0 expr1/=. +rewrite !indicE /onem !addrA addr0 expr1/=. by congr (_ + _ + _ + _); congr (_ * _); field. Qed. -End binomial_examples. - -Section hard_constraint'. -Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). - -Definition fail' := - letin' (score (@measurable_cst _ _ X _ setT (0%R : R))) - (ret (@measurable_cst _ _ _ Y setT point)). - -Lemma fail'E x U : fail' x U = 0. -Proof. by rewrite /fail' letin'E ge0_integral_mscale//= normr0 mul0e. Qed. - -End hard_constraint'. -Arguments fail' {d d' X Y R}. - -(* hard constraints to express score below 1 *) -Lemma score_fail' d (X : measurableType d) {R : realType} - (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : - score (kr r%:num) = - letin' (sample_cst (bernoulli r1) : R.-pker X ~> _) - (ite macc0of2 (ret ktt) fail'). -Proof. -apply/eq_sfkernel => x U. -rewrite letin'E/= /sample; unlock. -rewrite ge0_integral_measure_add//= ge0_integral_mscale//= ge0_integral_mscale//=. -rewrite !integral_dirac//= !diracT/= !mul1e. -by rewrite /mscale/= iteE//= iteE//= fail'E mule0 adde0 ger0_norm. -Qed. +End sample_binomial. Section hard_constraint. Local Open Scope ring_scope. @@ -524,18 +424,20 @@ rewrite letin'E integral_indic//= /mscale/= normr0 mul0e. by rewrite /fail' letin'E/= ge0_integral_mscale//= normr0 mul0e. Qed. -Lemma exec_score_fail (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : - execP (g := [::]) [Score {r%:num}:R] = - execP [let str := Sample {exp_bernoulli r r1} in +Lemma exec_score_fail (r : R) (r01 : (0 <= r <= 1)%R) : + execP (g := [::]) [Score {r}:R] = + execP [let str := Sample {exp_bernoulli [{r}:R]} in if #str then return TT else {hard_constraint _}]. Proof. -rewrite execP_score execD_real /= score_fail'. +move: r01 => /andP[r0 r1]//. +rewrite execP_score execD_real /= score_fail' ?r0 ?r1//. rewrite execP_letin execP_sample/= execD_bernoulli execP_if execP_return. rewrite execD_unit/= exp_var'E /=. - apply/ctx_prf_head. + exact/ctx_prf_head (* TODO *). move=> h. apply: eq_sfkernel=> /= -[] U. -rewrite 2!letin'E/=. +rewrite [LHS]letin'E/= [RHS]letin'E/=. +rewrite execD_real/=. apply: eq_integral => b _. rewrite 2!iteE//=. case: b => //=. @@ -552,6 +454,206 @@ Qed. End hard_constraint. +Section test_uniform. +Local Open Scope ring_scope. +Local Open Scope lang_scope. +Context (R : realType). + +Definition uniform_syntax : @exp R _ [::] _ := + [let "p" := Sample {exp_uniform 0 1 (@ltr01 R)} in + return #{"p"}]. + +Lemma exec_uniform_syntax t U : measurable U -> + execP uniform_syntax t U = uniform_prob (@ltr01 R) U. +Proof. +move=> mU. +rewrite /uniform_syntax execP_letin execP_sample execP_return !execD_uniform. +rewrite exp_var'E (execD_var_erefl "p")/=. +rewrite letin'E /=. +rewrite integral_uniform//=; last exact: measurable_fun_dirac. +rewrite subr0 invr1 mul1e. +rewrite {1}/uniform_prob. +rewrite integral_mkcond//=. +rewrite [in RHS]integral_mkcond//=. +apply: eq_integral => x _. +rewrite !patchE. +case: ifPn => //; case: ifPn => //. +- move=> xU. + rewrite inE/= in_itv/= => x01. + by rewrite /uniform_pdf x01 diracE xU subr0 invr1. +- by rewrite diracE => /negbTE ->. +- move=> xU. + rewrite notin_setE/= in_itv/= => /negP/negbTE x01. + by rewrite /uniform_pdf x01. +Qed. + +End test_uniform. + +Section guard. +Local Open Scope ring_scope. +Local Open Scope lang_scope. +Context (R : realType). + +Definition guard : @exp R _ [::] _ := [ + let "p" := Sample {exp_bernoulli [{1 / 3}:R]} in + let "_" := if #{"p"} then return TT else Score {0}:R in + return #{"p"} +]. + +Lemma exec_guard t U : execP guard t U = ((1 / 3)%:E * \d_true U)%E. +Proof. +rewrite /guard 2!execP_letin execP_sample execD_bernoulli execD_real. +rewrite execP_if/= !execP_return !exp_var'E !(execD_var_erefl "p") execD_unit. +rewrite execP_score execD_real/=. +rewrite letin'E/= integral_bernoulli//=; last lra. +rewrite !letin'E !iteE/= integral_dirac// ge0_integral_mscale//=. +by rewrite normr0 mul0e !mule0 !adde0 !diracT !mul1e. +Qed. + +End guard. + +Section test_binomial. +Local Open Scope ring_scope. +Local Open Scope lang_scope. +Context (R : realType). + +Definition binomial_le : @exp R _ [::] Bool := + [let "a2" := Sample {exp_binomial 3 [{1 / 2}:R]} in + return {1}:N <= #{"a2"}]. + +Lemma exec_binomial_le t U : + execP binomial_le t U = ((7 / 8)%:E * \d_true U + + (1 / 8)%:E * \d_false U)%E. +Proof. +rewrite /binomial_le execP_letin execP_sample execP_return execD_rel execD_nat. +rewrite exp_var'E (execD_var_erefl "a2") execD_binomial/= !execD_real/=. +rewrite letin'E//= integral_binomial//=; [lra|move=> _]. +rewrite !big_ord_recl big_ord0//=. +rewrite /bump. +rewrite !binS/= !bin0 bin1 bin2 bin_small// addn0. +rewrite addeC adde0. +congr (_ + _)%:E. + rewrite !indicE !(mul0n,add0n,lt0n,mul1r)/=. + rewrite -!mulrDl; congr (_ * _). + rewrite /onem. + lra. +rewrite !expr0 ltnn indicE/= !(mul1r,mul1e) /onem. +lra. +Qed. + +Definition binomial_guard : @exp R _ [::] Nat := + [let "a1" := Sample {exp_binomial 3 [{1 / 2}:R]} in + let "_" := if #{"a1"} == {1}:N then return TT else Score {0}:R in + return #{"a1"}]. + +Lemma exec_binomial_guard t U : + execP binomial_guard t U = ((3 / 8)%:E * \d_1%N U)%E. +Proof. +rewrite /binomial_guard !execP_letin execP_sample execP_return execP_if. +rewrite !exp_var'E execD_rel !(execD_var_erefl "a1") execP_return. +rewrite execD_unit execD_binomial execD_nat execP_score !execD_real. +rewrite !letin'E//=. +rewrite integral_binomial//=; [lra|move=> _]. +rewrite !big_ord_recl big_ord0. +rewrite /bump/=. +rewrite !binS/= !bin0 bin1 bin2 bin_small//. +rewrite !letin'E//= !iteE/=. +rewrite !ge0_integral_mscale//=. +rewrite !integral_dirac//= !diracE/=. +rewrite /bump/=. +rewrite !(normr0,mul0e,mule0,add0e,add0n,mul1e,adde0). +rewrite mem_set//=. +rewrite /onem mul1e. +congr (_%:E * _)%E. +lra. +Qed. + +End test_binomial. + +Section beta_bernoulli_bernoulli. +Local Open Scope ring_scope. +Local Open Scope lang_scope. +Context (R : realType). +Local Notation mu := (@lebesgue_measure R). + +(* TODO: move? *) +Lemma integrable_bernoulli_XMonemX01 a b U + (mu : {measure set (g_sigma_algebraType R.-ocitv.-measurable) -> \bar R}) : + measurable U -> (mu `[0%R, 1%R]%classic < +oo)%E -> + mu.-integrable `[0, 1] (fun x => bernoulli (XMonemX01 a b x) U). +Proof. +move=> mU mu01oo. +apply/integrableP; split. + apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. + by apply: measurable_funTS; exact: measurable_XMonemX01. +apply: (@le_lt_trans _ _ (\int[mu]_(x in `[0%R, 1%R]) cst 1 x)%E). + apply: ge0_le_integral => //=. + apply/measurable_funTS/measurableT_comp => //=. + apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. + exact: measurable_XMonemX01. + by move=> x _; rewrite gee0_abs// probability_le1. +by rewrite integral_cst//= mul1e. +Qed. + +Let measurable_bernoulli_XMonemX01 U : + measurable_fun setT (fun x : R => bernoulli (XMonemX01 2 1 x) U). +Proof. +apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. +exact: measurable_XMonemX01. +Qed. + +Lemma beta_bernoulli_bernoulli U : measurable U -> + @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in + Sample {exp_bernoulli [#{"p"}]}] tt U = + @execP R [::] _ [Sample {exp_bernoulli [{3 / 5}:R]}] tt U. +Proof. +move=> mU. +rewrite execP_letin !execP_sample execD_beta !execD_bernoulli/=. +rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. +transitivity (beta_prob_bernoulli 6 4 1 0 U : \bar R). + rewrite /beta_prob_bernoulli !letin'E/=. + rewrite integral_beta_prob//=; last 2 first. + exact: measurable_bernoulli2. + exact: integral_beta_prob_bernoulli_lty. + rewrite integral_beta_prob//=; last 2 first. + by apply: measurable_funTS => /=; exact: measurable_bernoulli_XMonemX01. + rewrite integral_beta_prob//=. + + suff: mu.-integrable `[0%R, 1%R] + (fun x => bernoulli (XMonemX01 2 1 x) U * (beta_pdf 6 4 x)%:E)%E. + move=> /integrableP[_]. + under eq_integral. + move=> x _. + rewrite gee0_abs//; last first. + by rewrite mule_ge0// lee_fin beta_pdf_ge0. + over. + move=> ?. + by under eq_integral do rewrite gee0_abs//. + + apply: integrableMl => //=. + * apply: integrable_bernoulli_XMonemX01 => //=. + by rewrite lebesgue_measure_itv//= lte01 EFinN sube0 ltry. + * by apply: measurable_funTS; exact: measurable_beta_pdf. + * exact: bounded_beta_pdf_01. + + apply/measurableT_comp => //; apply: measurable_funTS => /=. + exact: measurable_bernoulli_XMonemX01. + + under eq_integral do rewrite gee0_abs//=. + have : (beta_prob 6 4 `[0%R, 1%R] < +oo :> \bar R)%E. + by rewrite -ge0_fin_numE// beta_prob_fin_num. + by move=> /(@integrable_bernoulli_XMonemX01 2 1 _ (beta_prob 6 4) mU) /integrableP[]. + rewrite [RHS]integral_mkcond. + apply: eq_integral => x _ /=. + rewrite patchE. + case: ifPn => x01. + by rewrite /XMonemX01 patchE x01 XMonemX0' expr1. + by rewrite /beta_pdf /XMonemX01 patchE (negbTE x01) mul0r mule0. +rewrite beta_prob_bernoulliE// !bernoulliE//=; last 2 first. + lra. + by rewrite div_betafun_ge0 div_betafun_le1. +by congr (_ * _ + _ * _)%:E; + rewrite /div_betafun/= /onem !betafunE/=; repeat rewrite !factE/=; field. +Qed. + +End beta_bernoulli_bernoulli. + Section letinA. Local Open Scope lang_scope. Variable R : realType. @@ -600,20 +702,21 @@ Import Notations. Context {R : realType}. Definition staton_bus_syntax0 : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in + [let "x" := Sample {exp_bernoulli [{2 / 7}:R]} in let "r" := if #{"x"} then return {3}:R else return {10}:R in let "_" := Score {exp_poisson 4 [#{"r"}]} in return #{"x"}]. Definition staton_bus_syntax := [Normalize {staton_bus_syntax0}]. -Let sample_bern : R.-sfker munit ~> mbool := sample_cst (bernoulli p27). +Let sample_bern : R.-sfker munit ~> mbool := + sample _ (measurableT_comp measurable_bernoulli (measurable_cst (2 / 7 : R)%R)). Let ite_3_10 : R.-sfker mbool * munit ~> measurableTypeR R := ite macc0of2 (@ret _ _ _ (measurableTypeR R) R _ (kr 3)) (@ret _ _ _ (measurableTypeR R) R _ (kr 10)). Let score_poisson4 : R.-sfker measurableTypeR R * (mbool * munit) ~> munit := - score (measurableT_comp (measurable_poisson 4) (@macc0of2 _ _ (measurableTypeR R) _)). + score (measurableT_comp (measurable_poisson_pdf 4) (@macc0of2 _ _ (measurableTypeR R) _)). Let kstaton_bus' := letin' sample_bern @@ -622,7 +725,8 @@ Let kstaton_bus' := Lemma eval_staton_bus0 : staton_bus_syntax0 -P> kstaton_bus'. Proof. -apply: eval_letin; first by apply: eval_sample; exact: eval_bernoulli. +apply: eval_letin. + by apply: eval_sample; apply: eval_bernoulli; exact: eval_real. apply: eval_letin. apply/evalP_if; [|exact/eval_return/eval_real..]. rewrite exp_var'E. @@ -637,7 +741,7 @@ Qed. Lemma exec_staton_bus0' : execP staton_bus_syntax0 = kstaton_bus'. Proof. -rewrite 3!execP_letin execP_sample/= execD_bernoulli. +rewrite 3!execP_letin execP_sample/= execD_bernoulli/= !execD_real. rewrite /kstaton_bus'; congr letin'. rewrite !execP_if !execP_return !execD_real/=. rewrite exp_var'E (execD_var_erefl "x")/=. @@ -654,7 +758,7 @@ Lemma exec_staton_bus : execD staton_bus_syntax = existT _ (normalize_pt kstaton_bus') (measurable_normalize_pt _). Proof. by rewrite execD_normalize_pt exec_staton_bus0'. Qed. -Let poisson4 := @poisson R 4%N. +Let poisson4 := @poisson_pdf R 4%N. Let staton_bus_probability U := ((2 / 7)%:E * (poisson4 3)%:E * \d_true U + @@ -664,19 +768,21 @@ Lemma exec_staton_bus0 (U : set bool) : execP staton_bus_syntax0 tt U = staton_bus_probability U. Proof. rewrite exec_staton_bus0' /staton_bus_probability /kstaton_bus'. -rewrite letin'_sample_bernoulli. +rewrite /sample_bern. +rewrite letin'E/=. +rewrite integral_bernoulli//=; last lra. rewrite -!muleA; congr (_ * _ + _ * _)%E. - rewrite letin'_iteT//. rewrite letin'_retk//. rewrite letin'_kret//. rewrite /score_poisson4. - by rewrite /score/= /mscale/= ger0_norm//= poisson_ge0. + by rewrite /score/= /mscale/= ger0_norm//= poisson_pdf_ge0. - by rewrite onem27. - rewrite letin'_iteF//. rewrite letin'_retk//. rewrite letin'_kret//. rewrite /score_poisson4. - by rewrite /score/= /mscale/= ger0_norm//= poisson_ge0. + by rewrite /score/= /mscale/= ger0_norm//= poisson_pdf_ge0. Qed. End staton_bus. @@ -689,7 +795,7 @@ Import Notations. Context {R : realType}. Definition staton_busA_syntax0 : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli (2 / 7%:R)%:nng p27} in + [let "x" := Sample {exp_bernoulli [{2 / 7}:R]} in let "_" := let "r" := if #{"x"} then return {3}:R else return {10}:R in Score {exp_poisson 4 [#{"r"}]} in @@ -698,13 +804,14 @@ Definition staton_busA_syntax0 : @exp R _ [::] _ := Definition staton_busA_syntax : exp _ [::] _ := [Normalize {staton_busA_syntax0}]. -Let sample_bern : R.-sfker munit ~> mbool := sample_cst (bernoulli p27). +Let sample_bern : R.-sfker munit ~> mbool := + sample _ (measurableT_comp measurable_bernoulli (measurable_cst (2 / 7 : R)%R)). Let ite_3_10 : R.-sfker mbool * munit ~> measurableTypeR R := ite macc0of2 (@ret _ _ _ (measurableTypeR R) R _ (kr 3)) (@ret _ _ _ (measurableTypeR R) R _ (kr 10)). Let score_poisson4 : R.-sfker measurableTypeR R * (mbool * munit) ~> munit := - score (measurableT_comp (measurable_poisson 4) (@macc0of3' _ _ _ (measurableTypeR R) _ _)). + score (measurableT_comp (measurable_poisson_pdf 4) (@macc0of3' _ _ _ (measurableTypeR R) _ _)). (* same as kstaton_bus _ (measurable_poisson 4) but expressed with letin' instead of letin *) @@ -715,19 +822,10 @@ Let kstaton_busA' := score_poisson4) (ret macc1of3')). -(*Lemma kstaton_busA'E : kstaton_busA' = kstaton_bus _ (measurable_poisson 4). -Proof. -apply/eq_sfkernel => -[] U. -rewrite /kstaton_busA' /kstaton_bus. -rewrite letin'_letin. -rewrite /sample_bern. -congr (letin _ _ tt U). -rewrite 2!letin'_letin/=. -Abort.*) - Lemma eval_staton_busA0 : staton_busA_syntax0 -P> kstaton_busA'. Proof. -apply: eval_letin; first by apply: eval_sample; exact: eval_bernoulli. +apply: eval_letin. + by apply: eval_sample; apply: eval_bernoulli; exact: eval_real. apply: eval_letin. apply: eval_letin. apply/evalP_if; [|exact/eval_return/eval_real..]. @@ -742,7 +840,7 @@ Qed. Lemma exec_staton_busA0' : execP staton_busA_syntax0 = kstaton_busA'. Proof. -rewrite 3!execP_letin execP_sample/= execD_bernoulli. +rewrite 3!execP_letin execP_sample/= execD_bernoulli execD_real. rewrite /kstaton_busA'; congr letin'. rewrite !execP_if !execP_return !execD_real/=. rewrite exp_var'E (execD_var_erefl "x")/=. @@ -789,7 +887,7 @@ rewrite execP_return exp_var'E/= (execD_var_erefl "x") //=. by apply/eq_sfkernel => /= -[[] [a [b []]]] U0. Qed. -Let poisson4 := @poisson R 4%N. +Let poisson4 := @poisson_pdf R 4%N. Lemma exec_staton_busA0 U : execP staton_busA_syntax0 tt U = ((2 / 7%:R)%:E * (poisson4 3%:R)%:E * \d_true U + @@ -802,8 +900,6 @@ Section letinC. Local Open Scope lang_scope. Variable (R : realType). -Require Import Classical_Prop. (* TODO: mv *) - Let weak_head g {t1 t2} x (e : @exp R P g t2) (xg : x \notin dom g) := exp_weak P [::] _ (x, t1) e xg. diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v new file mode 100644 index 0000000000..cef5fe1cae --- /dev/null +++ b/theories/lang_syntax_table_game.v @@ -0,0 +1,717 @@ +Require Import String. +From HB Require Import structures. +From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. +From mathcomp Require Import unstable mathcomp_extra boolp classical_sets functions. +From mathcomp Require Import cardinality fsbigop interval_inference. +From mathcomp Require Import reals ereal topology normedtype sequences. +From mathcomp Require Import esum measure charge lebesgue_measure numfun. +From mathcomp Require Import lebesgue_integral probability kernel prob_lang. +From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. +From mathcomp Require Import ring lra. + +(**md**************************************************************************) +(* # Eddy's table game example *) +(* *) +(* ref: *) +(* - Chung-chieh Shan, Equational reasoning for probabilistic programming, *) +(* POPL TutorialFest 2018 *) +(* https://homes.luddy.indiana.edu/ccshan/rational/equational-handout.pdf *) +(* - Sean R Eddy, What is Bayesian statistics?, Nature Biotechnology 22(9), *) +(* 1177--1178 (2004) *) +(* *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import Order.TTheory GRing.Theory Num.Def Num.Theory. +Import numFieldTopology.Exports. + +Local Open Scope classical_set_scope. +Local Open Scope ring_scope. +Local Open Scope string_scope. + +Local Open Scope ereal_scope. +Lemma letin'_sample_uniform {R : realType} d d' (T : measurableType d) + (T' : measurableType d') (a b : R) (ab : (a < b)%R) + (u : R.-sfker [the measurableType _ of (_ * T)%type] ~> T') x y : + measurable y -> + letin' (sample_cst (uniform_prob ab)) u x y = + (b - a)^-1%:E * \int[lebesgue_measure]_(x0 in `[a, b]) u (x0, x) y. +Proof. +move=> my; rewrite letin'E/=. +rewrite integral_uniform//=. +move => _ /= Y mY /=. +have /= := measurable_kernel u _ my measurableT _ mY. +move/measurable_ysection => /(_ x) /=. +set A := (X in measurable X). +set B := (X in _ -> measurable X). +suff : A = B by move=> ->. +by rewrite {}/A {}/B !setTI ysectionE. +Qed. + +Local Open Scope lang_scope. +Lemma execP_letin_uniform {R : realType} + g t str (s0 s1 : exp P ((str, Real) :: g) t) : + (forall (p : R) x U, (0 <= p <= 1)%R -> + execP s0 (p, x) U = execP s1 (p, x) U) -> + forall x U, measurable U -> + execP [let str := Sample {@exp_uniform _ g 0 1 (@ltr01 R)} in {s0}] x U = + execP [let str := Sample {@exp_uniform _ g 0 1 (@ltr01 R)} in {s1}] x U. +Proof. +move=> s01 x U mU. +rewrite !execP_letin execP_sample execD_uniform/=. +rewrite !letin'_sample_uniform//. +congr *%E. +apply: eq_integral => p p01. +apply: s01. +by rewrite inE in p01. +Qed. +Local Close Scope lang_scope. +Local Close Scope ereal_scope. + +Section bounded. +Local Open Scope ring_scope. +Local Open Scope lang_scope. +Local Open Scope ereal_scope. +Context {R : realType}. + +Lemma bounded_id_01 : [bounded x0 : R^o | x0 in `[0%R, 1%R]%classic : set R]. +Proof. +exists 1%R; split => // y y1. +near=> M => /=. +rewrite (le_trans _ (ltW y1))//. +near: M. +move=> M /=. +rewrite in_itv/= => /andP[M0 M1]. +by rewrite ler_norml M1 andbT (le_trans _ M0). +Unshelve. all: by end_near. Qed. + +Lemma bounded_onem_01 : [bounded (`1- x : R^o) | x in `[0%R, 1%R]%classic : set R]. +Proof. +exists 1%R; split => // y y1. +near=> M => /=. +rewrite (le_trans _ (ltW y1))//. +near: M. +move=> M /=. +rewrite in_itv/= => /andP[M0 M1]. +rewrite ler_norml (@le_trans _ _ 0%R)//=. + by rewrite lerBlDr addrC -lerBlDr subrr. +by rewrite onem_ge0. +Unshelve. all: by end_near. Qed. + +Lemma bounded_cst_01 (x : R) : [bounded x : R^o | _ in `[0%R, 1%R]%classic : set R]. +Proof. +exists `|x|%R; split. + by rewrite num_real. +move=> y y1/= z. +rewrite in_itv/= => /andP[z0 z1]. +by rewrite (le_trans _ (ltW y1)). +Qed. + +Lemma bounded_norm (f : R -> R) : + [bounded f x : R^o | x in (`[0%R, 1%R]%classic : set R)] <-> + [bounded `|f x|%R : R^o | x in (`[0%R, 1%R]%classic : set R)]. +Proof. +split. + move=> [M [Mreal HM]]. + exists `|M|%R; split; first by rewrite normr_real. + move=> r Mr x/= x01. + by rewrite ger0_norm// HM// (le_lt_trans _ Mr)// ler_norm. +move=> [M [Mreal HM]]. +exists `|M|%R; split; first by rewrite normr_real. +move=> r Mr x/= x01. +rewrite -[leLHS]ger0_norm// HM//. +by rewrite (le_lt_trans _ Mr)// ler_norm. +Qed. + +Lemma boundedMl k (f : R -> R) : + [bounded f x : R^o | x in (`[0%R, 1%R]%classic : set R)] -> + [bounded (k * f x)%R : R^o | x in (`[0%R, 1%R]%classic : set R)]. +Proof. +move=> [M [Mreal HM]]. +exists `|k * M|%R; split; first by rewrite normr_real. +move=> r kMr x/= x01. +rewrite normrM. +have [->|k0] := eqVneq k 0%R. + by rewrite normr0 mul0r (le_trans _ (ltW kMr)). +rewrite -ler_pdivlMl ?normr_gt0//. +apply: HM => //. +rewrite ltr_pdivlMl ?normr_gt0//. +rewrite (le_lt_trans _ kMr)//. +by rewrite normrM ler_pM2l ?normr_gt0// ler_norm. +Qed. + +End bounded. + +Lemma measurable_bernoulli_expn {R : realType} U n : + measurable_fun [set: g_sigma_algebraType R.-ocitv.-measurable] + (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => bernoulli (`1-x ^+ n) U). +Proof. +apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. +by apply: measurable_funX => //=; exact: measurable_funB. +Qed. + +Lemma integrable_bernoulli_beta_pdf {R : realType} U : measurable U -> + (@lebesgue_measure R).-integrable [set: g_sigma_algebraType R.-ocitv.-measurable] + (fun x => bernoulli (1 - `1-x ^+ 3) U * (beta_pdf 6 4 x)%:E)%E. +Proof. +move=> mU. +have ? : measurable_fun [set: g_sigma_algebraType R.-ocitv.-measurable] + (fun x => bernoulli (1 - `1-x ^+ 3) U). + apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. + apply: measurable_funB => //; apply: measurable_funX => //. + exact: measurable_funB. +apply/integrableP; split => /=. + apply: emeasurable_funM => //=. + by apply/measurable_EFinP; exact: measurable_beta_pdf. +apply: (@le_lt_trans _ _ (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (beta_pdf 6 4 x)%:E))%E. + rewrite [leRHS]integral_mkcond /=. + apply: ge0_le_integral => //=. + - apply: measurableT_comp => //; apply: emeasurable_funM => //. + by apply/measurable_EFinP; exact: measurable_beta_pdf. + - move=> x _ /=; rewrite patchE; case: ifPn => // _. + by rewrite lee_fin beta_pdf_ge0. + - apply/(measurable_restrict _ _ _) => //. + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_beta_pdf. + - move=> x _. + rewrite patchE; case: ifPn => x01. + rewrite gee0_abs//. + rewrite gee_pMl// ?probability_le1//. + by rewrite ge0_fin_numE// (le_lt_trans (probability_le1 _ _))// ltry. + by rewrite lee_fin beta_pdf_ge0. + by rewrite mule_ge0// lee_fin beta_pdf_ge0. + by rewrite /beta_pdf /XMonemX01 patchE (negbTE x01) mul0r mule0 abse0. +apply: (@le_lt_trans _ _ + (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (betafun 6 4)^-1%:E)%E); last first. + by rewrite integral_cst//= lebesgue_measure_itv/= lte01 EFinN sube0 mule1 ltry. +apply: ge0_le_integral => //=. +- by move=> ? _; rewrite lee_fin beta_pdf_ge0. +- by apply/measurable_funTS/measurableT_comp => //; exact: measurable_beta_pdf. +- by move=> ? _; rewrite lee_fin invr_ge0// betafun_ge0. +- by move=> x _; rewrite lee_fin beta_pdf_le_betafunV. +Qed. + +Section game_programs. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. +Local Open Scope lang_scope. +Context (R : realType). +Local Notation mu := lebesgue_measure. + +Definition guard {g} str n : @exp R P [:: (str, _) ; g] _ := + [if #{str} == {n}:N then return TT else Score {0}:R]. + +Definition prog0 : @exp R _ [::] _ := + [Normalize + let "p" := Sample {exp_uniform 0 1 (@ltr01 R)} in + let "x" := Sample {exp_binomial 8 [#{"p"}]} in + let "_" := {guard "x" 5} in + let "y" := Sample {exp_binomial 3 [#{"p"}]} in + return {1}:N <= #{"y"}]. + +Definition tail1 : @exp R _ [:: ("_", Unit); ("x", Nat) ; ("p", Real)] _ := + [Sample {exp_bernoulli [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%R}]}]. + +Definition tail2 : @exp R _ [:: ("_", Unit); ("p", Real)] _ := + [Sample {exp_bernoulli [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%R}]}]. + +Definition tail3 : @exp R _ [:: ("p", Real); ("_", Unit)] _ := + [Sample {exp_bernoulli [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%R}]}]. + +Definition prog1 : @exp R _ [::] _ := + [Normalize + let "p" := Sample {exp_uniform 0 1 (@ltr01 R)} in + let "x" := Sample {exp_binomial 8 [#{"p"}]} in + let "_" := {guard "x" 5} in + {tail1}]. + +Definition prog2 : @exp R _ [::] _ := + [Normalize + let "p" := Sample {exp_uniform 0 1 (@ltr01 R)} in + let "_" := + Score {[{56}:R * #{"p"} ^+ {5%R} * {[{1}:R - #{"p"}]} ^+ {3%R}]} in + {tail2}]. + +Definition prog2' : @exp R _ [::] _ := + [Normalize + let "p" := Sample {exp_beta 1 1} in + let "_" := Score + {[{56}:R * #{"p"} ^+ {5%R} * {[{1}:R - #{"p"}]} ^+ {3%R}]} in + {tail2}]. + +Definition prog3 : @exp R _ [::] _ := + [Normalize + let "_" := Score {1 / 9}:R in + let "p" := Sample {exp_beta 6 4} in + {tail3}]. + +Definition prog4 : @exp R _ [::] _ := + [Normalize + let "_" := Score {1 / 9}:R in + Sample {exp_bernoulli [{10 / 11}:R]}]. + +Definition prog5 : @exp R _ [::] _ := + [Normalize Sample {exp_bernoulli [{10 / 11}:R]}]. + +End game_programs. +Arguments tail1 {R}. +Arguments tail2 {R}. +Arguments guard {R g}. + +Section from_prog0_to_prog1. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. +Local Open Scope lang_scope. +Context (R : realType). +Local Notation mu := lebesgue_measure. + +Let prog01_subproof + (x : mctx (untag (ctx_of (recurse Unit (recurse Nat (found "p" Real [::])))))) + U : (0 <= x.2.2.1 <= 1)%R -> + execP [let "y" := Sample {exp_binomial 3 [#{"p"}]} in + return {1}:N <= #{"y"}] x U = + execP (@tail1 R) x U. +Proof. +move=> x01. +rewrite /tail1. +(* reduce lhs *) +rewrite execP_letin execP_sample execD_binomial/= execP_return/= execD_rel/=. +rewrite exp_var'E (execD_var_erefl "p")/=. +rewrite exp_var'E (execD_var_erefl "y")/=. +rewrite execD_nat/=. +rewrite [LHS]letin'E/=. +(* reduce rhs *) +rewrite execP_sample/= execD_bernoulli/= (@execD_bin _ _ binop_minus)/=. +rewrite execD_real/= execD_pow/= (@execD_bin _ _ binop_minus)/= execD_real/=. +rewrite (execD_var_erefl "p")/=. +exact/integral_binomial_prob. +Qed. + +Lemma prog01 : execD (@prog0 R) = execD (@prog1 R). +Proof. +rewrite /prog0 /prog1. +apply: congr_normalize => y A. +apply: execP_letin_uniform => // p [] B p01. +apply: congr_letinr => a1 V0. +apply: congr_letinr => -[] V1. +exact: prog01_subproof. +Qed. + +End from_prog0_to_prog1. + +Section from_prog1_to_prog2. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. +Local Open Scope lang_scope. +Context (R : realType). +Local Notation mu := lebesgue_measure. + +Let prog12_subproof (y : @mctx R [::]) (V : set (@mtyp R Bool)) + (p : R) + (x : projT2 (existT measurableType default_measure_display unit)) + (U : set (mtyp Bool)) + (p0 : (0 <= p)%R) + (p1 : (p <= 1)%R) : + \int[binomial_prob 8 p]_y0 + execP [let "_" := {guard "x" 5} in {tail1}] + (y0, (p, x)) U = + \int[mscale (NngNum (normr_ge0 (56 * XMonemX 5 3 p)%R)) \d_tt]_y0 + execP tail2 (y0, (p, x)) U. +Proof. +rewrite integral_binomial//=. +rewrite (bigD1 (inord 5))//=. +rewrite big1 ?adde0; last first. + move=> i i5. + rewrite execP_letin/= execP_if/= execD_rel/=. + rewrite exp_var'E/= (execD_var_erefl "x")/=. + rewrite execD_nat/= execP_score/= execD_real/= execP_return/=. + rewrite letin'E iteE/=. + move: i => [[|[|[|[|[|[|[|[|[|//]]]]]]]]]]//= Hi in i5 *; + rewrite ?ge0_integral_mscale//= ?execD_real/= ?normr0 ?(mul0e,mule0)//. + by rewrite -val_eqE/= inordK in i5. +(* reduce lhs *) +rewrite -[(p ^+ _ * _ ^+ _)%R]/(XMonemX _ _ p). +rewrite execP_letin/= execP_if/= execD_rel/=. +rewrite exp_var'E/= (execD_var_erefl "x")/=. +rewrite execD_nat/= execP_score/= execD_real/= execP_return/=. +rewrite letin'E iteE/=. +rewrite inordK// eqxx. +rewrite integral_dirac//= execD_unit/= diracE mem_set// mul1e. +(* reduce rhs *) +rewrite ge0_integral_mscale//=. +rewrite integral_dirac//= diracE mem_set// mul1e. +rewrite ger0_norm ?mulr_ge0 ?subr_ge0//. +rewrite mulr_natl. +(* same score *) +congr *%E. +(* the tails are the same module the shape of the environment *) +rewrite /tail1 /tail2 !execP_sample/=. +rewrite !execD_bernoulli/=. +rewrite !(@execD_bin _ _ binop_minus)/=. +rewrite !execD_pow/=. +rewrite !execD_real/=. +rewrite !(@execD_bin _ _ binop_minus)/=. +by rewrite !execD_real/= !exp_var'E/= !(execD_var_erefl "p")/=. +Qed. + +Lemma prog12 : execD (@prog1 R) = execD (@prog2 R). +Proof. +apply: congr_normalize => y V. +apply: execP_letin_uniform => // p x U /andP[p0 p1]. +(* reduce the lhs *) +rewrite execP_letin execP_sample execD_binomial/=. +rewrite letin'E/=. +rewrite [in LHS]exp_var'E/= (execD_var_erefl "p")/=. +(* reduce the rhs *) +rewrite [in RHS]execP_letin execP_score/=. +rewrite letin'E/=. +do 2 rewrite (@execD_bin _ _ binop_mult)/=/=. +rewrite [in RHS]exp_var'E/=. +rewrite execD_pow/=. +rewrite (execD_var_erefl "p")/=. +rewrite execD_pow/=. +rewrite (@execD_bin _ _ binop_minus)/=/=. +rewrite 2!execD_real/=. +rewrite (execD_var_erefl "p")/=. +rewrite -(mulrA 56%R). +exact: prog12_subproof. +Qed. + +End from_prog1_to_prog2. + +Local Open Scope ereal_scope. + +Lemma measurable_bernoulli_onemXn {R : realType} U : + measurable_fun [set: g_sigma_algebraType R.-ocitv.-measurable] + (fun x => bernoulli (1 - `1-x ^+ 3) U). +Proof. +apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. +apply: measurable_funB => //. +by apply: measurable_funX; exact: measurable_funB. +Qed. + +Lemma bounded_norm_XnonemXn {R : realType} : + [bounded normr (56 * XMonemX 5 3 x)%R : R^o | x in `[0%R, 1%R] : set R]. +Proof. exact/(bounded_norm _).1/boundedMl/bounded_XMonemX. Qed. + +Lemma integrable_bernoulli_XMonemX {R : realType} U : + (beta_prob 1 1).-integrable [set: R] + (fun x => bernoulli (1 - `1-x ^+ 3) U * (normr (56 * XMonemX 5 3 x))%:E). +Proof. +apply/integrableP; split. + apply: emeasurable_funM; first exact: measurable_bernoulli_onemXn. + apply/measurable_EFinP => //; apply: measurableT_comp => //. + by apply: measurable_funM => //; exact: measurable_fun_XMonemX. +rewrite beta_prob_uniform integral_uniform//=. + rewrite subr0 invr1 mul1e. + suff : lebesgue_measure.-integrable `[0%R, 1%R] + (fun y : R => bernoulli (1 - `1-y ^+ 3) U * (normr (56 * XMonemX 5 3 y))%:E). + by move=> /integrableP[]. + apply: integrableMl => //=. + - apply/integrableP; split. + by apply: measurable_funTS; exact: measurable_bernoulli_onemXn. + have := @integral_beta_prob_bernoulli_onem_lty R 3 1%N 1%N U. + rewrite beta_prob_uniform integral_uniform//=; last first. + by apply: measurableT_comp => //=; exact: measurable_bernoulli_onemXn. + by rewrite subr0 invr1 mul1e. + - apply: @measurableT_comp => //=; apply: measurable_funM => //. + exact: measurable_fun_XMonemX. + exact: bounded_norm_XnonemXn. +apply: @measurableT_comp => //; apply: emeasurable_funM => //=. + exact: measurable_bernoulli_onemXn. +do 2 apply: @measurableT_comp => //=. +by apply: measurable_funM => //; exact: measurable_fun_XMonemX. +Qed. + +Section from_prog2_to_prog3. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. +Local Open Scope lang_scope. +Context (R : realType). +Local Notation mu := lebesgue_measure. + +Lemma prog22' : execD (@prog2 R) = execD (@prog2' R). +Proof. +apply: congr_normalize => // x U. +apply: congr_letinl => // y V. +rewrite !execP_sample execD_uniform/= execD_beta/=. +by rewrite beta_prob_uniform. +Qed. + +Lemma prog23 : execD (@prog2' R) = execD (@prog3 R). +Proof. +apply: congr_normalize => x U. +(* reduce the LHS *) +rewrite 2![in LHS]execP_letin. +rewrite ![in LHS]execP_sample. +rewrite [in LHS]execP_score. +rewrite [in LHS]execD_beta/=. +rewrite [in LHS]execD_bernoulli. +rewrite 2![in LHS](@execD_bin _ _ binop_mult)/=. +rewrite 2![in LHS]execD_pow/=. +rewrite 2![in LHS](@execD_bin _ _ binop_minus)/=. +rewrite 3![in LHS]execD_real. +rewrite [in LHS]exp_var'E [in LHS](execD_var_erefl "p")/=. +rewrite [in LHS]execD_pow/=. +rewrite [in LHS](@execD_bin _ _ binop_minus)/=. +rewrite [in LHS]execD_real. +rewrite [in LHS]exp_var'E [in LHS](execD_var_erefl "p")/=. +(* reduce the RHS *) +rewrite [in RHS]execP_letin. +rewrite [in RHS]execP_score. +rewrite [in RHS]execP_letin/=. +rewrite [in RHS]execP_sample/=. +rewrite [in RHS]execD_beta/=. +rewrite [in RHS]execP_sample/=. +rewrite [in RHS]execD_bernoulli/=. +rewrite [in RHS]execD_real/=. +rewrite [in RHS](@execD_bin _ _ binop_minus)/=. +rewrite [in RHS]execD_real/=. +rewrite [in RHS]execD_pow/=. +rewrite [in RHS](@execD_bin _ _ binop_minus)/=. +rewrite [in RHS]exp_var'E [in RHS](execD_var_erefl "p")/=. +rewrite [in RHS]execD_real/=. +rewrite [LHS]letin'E/=. +under eq_integral => y _. + rewrite letin'E/=. + rewrite integral_cst//= /mscale/= diracT mule1 -mulrA -/(XMonemX _ _ _). + over. +rewrite [RHS]letin'E/=. +under [in RHS]eq_integral => y _. + rewrite letin'E/=. + over. +rewrite /=. +rewrite [RHS]ge0_integral_mscale//=; last first. + by move=> _ _; rewrite integral_ge0. +rewrite integral_beta_prob//=; last 2 first. + - apply: emeasurable_funM => //=. + exact: measurable_bernoulli_onemXn. + apply/measurable_EFinP; apply: measurableT_comp => //. + by apply: measurable_funM => //; exact: measurable_fun_XMonemX. + - by have /integrableP[] := @integrable_bernoulli_XMonemX R U. +rewrite ger0_norm// integral_dirac// diracT mul1e. +rewrite integral_beta_prob/=; [|by []|exact: measurable_bernoulli_onemXn + |exact: integral_beta_prob_bernoulli_onem_lty]. +rewrite -integralZl//=; last exact: integrable_bernoulli_beta_pdf. +apply: eq_integral => y _. +rewrite [in RHS]muleCA -[in LHS]muleA; congr *%E. +rewrite /beta_pdf /XMonemX01 2!patchE; case: ifPn => [y01|_]; last first. + by rewrite !mul0r 2!mule0. +rewrite ger0_norm; last first. + by rewrite mulr_ge0// XMonemX_ge0//; rewrite inE in y01. +rewrite [X in _ = _ * X]EFinM [in RHS]muleCA. +rewrite /= XMonemX00 mul1r [in LHS](mulrC 56%R) [in LHS]EFinM -[in LHS]muleA; congr *%E. +by rewrite !betafunE/=; repeat rewrite !factE/=; rewrite -EFinM; congr EFin; lra. +Qed. + +End from_prog2_to_prog3. + +Local Open Scope ereal_scope. +(* TODO: move? *) +Lemma int_beta_prob01 {R : realType} (f : R -> R) a b U : + measurable_fun [set: R] f -> + (forall x, x \in `[0%R, 1%R] -> 0 <= f x <= 1)%R -> + \int[beta_prob a b]_y bernoulli (f y) U = + \int[beta_prob a b]_(y in `[0%R, 1%R] : set R) bernoulli (f y) U. +Proof. +move=> mf f01. +rewrite [LHS]integral_beta_prob//=; last 2 first. + apply: measurable_funTS. + by apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. + exact: integral_beta_prob_bernoulli_lty. +rewrite [RHS]integral_beta_prob//; last 2 first. + apply/measurable_funTS => //=. + by apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. + apply: (le_lt_trans _ (lang_syntax.integral_beta_prob_bernoulli_lty a b U mf f01)). + apply: ge0_subset_integral => //=; apply: measurableT_comp => //=. + by apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. +rewrite [RHS]integral_mkcond/=; apply: eq_integral => x _ /=. +rewrite !patchE; case: ifPn => // x01. +by rewrite /beta_pdf /XMonemX01 patchE (negbTE x01) mul0r mule0. +Qed. + +Lemma expr_onem_01 {R : realType} (x : R) : x \in `[0%R, 1%R] -> + (0 <= `1-x ^+ 3 <= 1)%R. +Proof. +rewrite in_itv/= => /andP[x0 x1]. +rewrite exprn_ge0 ?subr_ge0//= exprn_ile1// ?subr_ge0//. +by rewrite lerBlDl -lerBlDr subrr. +Qed. + +Lemma int_beta_prob_bernoulli {R : realType} (U : set (@mtyp R Bool)) : + \int[beta_prob 6 4]_y bernoulli (`1-y ^+ 3) U = bernoulli (1 / 11) U :> \bar R. +Proof. +rewrite int_beta_prob01//; last 2 first. + by apply: measurable_funX => //; exact: measurable_funB. + exact: expr_onem_01. +have := @beta_prob_bernoulliE R 6 4 0 3 U isT isT. +rewrite /beta_prob_bernoulli. +under eq_integral. + move=> x x0. + rewrite /XMonemX01 patchE x0 XMonemX0. + over. +rewrite /= => ->; congr bernoulli. +by rewrite /div_betafun addn0 !betafunE/=; repeat rewrite !factE/=; field. +Qed. + +Lemma dirac_bool {R : realType} (U : set bool) : + \d_false U + \d_true U = (\sum_(x \in U) (1%E : \bar R))%R. +Proof. +have [| | |] := set_bool U => /eqP ->; rewrite !diracE. +- by rewrite memNset// mem_set//= fsbig_set1 add0e. +- by rewrite mem_set// memNset//= fsbig_set1 adde0. +- by rewrite !in_set0 fsbig_set0 adde0. +- rewrite !in_setT setT_bool fsbigU0//=; last by move=> x [->]. + by rewrite !fsbig_set1. +Qed. + +Lemma int_beta_prob_bernoulli_onem {R : realType} (U : set (@mtyp R Bool)) : + \int[beta_prob 6 4]_y bernoulli (`1-(`1-y ^+ 3)) U = bernoulli (10 / 11) U :> \bar R. +Proof. +transitivity (\d_false U + \d_true U - bernoulli (1 / 11) U : \bar R)%E; last first. + rewrite /bernoulli ifT; last lra. + rewrite ifT; last lra. + apply/eqP; rewrite sube_eq//; last first. + rewrite ge0_adde_def// inE. + by apply/sume_ge0 => //= b _; rewrite lee_fin bernoulli_pmf_ge0//; lra. + by apply/sume_ge0 => //= b _; rewrite lee_fin bernoulli_pmf_ge0//; lra. + rewrite -fsbig_split//=. + under eq_fsbigr. + move=> /= x _. + rewrite -EFinD /bernoulli_pmf [X in X%:E](_ : _ = 1%R); last first. + case: x => //; lra. + over. + by rewrite /= dirac_bool. +rewrite -int_beta_prob_bernoulli. +apply/esym/eqP; rewrite sube_eq//; last first. + by rewrite ge0_adde_def// inE; exact: integral_ge0. +rewrite int_beta_prob01; last 2 first. + apply: measurable_funB => //; apply: measurable_funX => //. + exact: measurable_funB. + move=> x x01. + by rewrite subr_ge0 andbC lerBlDr -lerBlDl subrr expr_onem_01. +rewrite [X in _ == _ + X]int_beta_prob01; last 2 first. + by apply: measurable_funX => //; exact: measurable_funB. + exact: expr_onem_01. +rewrite -ge0_integralD//=; last 2 first. + apply: (@measurableT_comp _ _ _ _ _ _ (bernoulli ^~ U)) => /=. + exact: measurable_bernoulli2. + apply: measurable_funB => //=; apply: measurable_funX => //=. + exact: measurable_funB. + apply: (@measurableT_comp _ _ _ _ _ _ (bernoulli ^~ U)) => /=. + exact: measurable_bernoulli2. + by apply: measurable_funX => //=; exact: measurable_funB. +apply/eqP; transitivity + (\int[beta_prob 6 4]_(x in `[0%R, 1%R]) (\d_false U + \d_true U) : \bar R). + by rewrite integral_cst//= beta_prob01 mule1 EFinD. +apply: eq_integral => /= x x01. +rewrite /bernoulli subr_ge0 lerBlDr -lerBlDl subrr andbC. +rewrite (_ : (_ <= _ <= _)%R); last first. + by apply: expr_onem_01; rewrite inE in x01. +rewrite -fsbig_split//=. +under eq_fsbigr. + move=> /= y yU. + rewrite -EFinD /bernoulli_pmf. + rewrite [X in X%:E](_ : _ = 1%R); last first. + by case: ifPn => _; rewrite subrK. + over. +by rewrite /= dirac_bool. +Qed. + +Local Close Scope ereal_scope. + +Section from_prog3_to_prog4. +Local Open Scope ereal_scope. +Local Open Scope lang_scope. +Context (R : realType). + +(* NB: not used *) +Lemma prog34' U : + @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in + Sample {exp_bernoulli [{[{1}:R - #{"p"}]} ^+ {3%N}]}] tt U = + @execP R [::] _ [Sample {exp_bernoulli [{1 / 11}:R]}] tt U. +Proof. +(* reduce the lhs *) +rewrite execP_letin. +rewrite execP_sample execD_beta/=. +rewrite execP_sample/= execD_bernoulli/=. +rewrite execD_pow/= (@execD_bin _ _ binop_minus) execD_real/=. +rewrite exp_var'E (execD_var_erefl "p")/=. +(* reduce the rhs *) +rewrite execP_sample execD_bernoulli/= execD_real. +(* semantics of lhs *) +rewrite letin'E/=. +exact: int_beta_prob_bernoulli. +Qed. + +Lemma prog34 l u U : + @execP R l _ [let "p" := Sample {exp_beta 6 4} in + Sample {exp_bernoulli [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%N}]}] u U = + @execP R l _ [Sample {exp_bernoulli [{10 / 11}:R]}] u U. +Proof. +(* reduce the lhs *) +rewrite execP_letin. +rewrite execP_sample execD_beta/=. +rewrite execP_sample/= execD_bernoulli/=. +rewrite (@execD_bin _ _ binop_minus)/=. +rewrite execD_pow/= (@execD_bin _ _ binop_minus) execD_real/=. +rewrite exp_var'E (execD_var_erefl "p")/=. +(* reduce the rhs *) +rewrite execP_sample execD_bernoulli/= execD_real. +(* semantics of lhs *) +rewrite letin'E/=. +exact: int_beta_prob_bernoulli_onem. +Qed. + +End from_prog3_to_prog4. + +Section from_prog4_to_prog5. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. +Local Open Scope lang_scope. +Context (R : realType). +Local Notation mu := lebesgue_measure. + +Lemma normalize_score_bernoulli g p q (p0 : (0 < p)%R) (q01 : (0 <= q <= 1)%R) : + @execD R g _ [Normalize let "_" := Score {p}:R in + Sample {exp_bernoulli [{q}:R]}] = + execD [Normalize Sample {exp_bernoulli [{q}:R]}]. +Proof. +apply: eq_execD. +rewrite !execD_normalize_pt/= !execP_letin !execP_score. +rewrite !execP_sample !execD_bernoulli !execD_real/=. +apply: funext=> x. +apply: eq_probability=> /= U. +rewrite !normalizeE/=. +rewrite !bernoulliE//=; [|lra..]. +rewrite !diracT !mule1 -EFinD add_onemK onee_eq0/=. +rewrite !letin'E. +under eq_integral. + move=> A _ /=. + rewrite !bernoulliE//=; [|lra..]. + rewrite !diracT !mule1 -EFinD add_onemK. + over. +rewrite !ge0_integral_mscale//= (ger0_norm (ltW p0))//. +rewrite integral_dirac// !diracT !indicT /= !mule1 !mulr1. +rewrite add_onemK invr1 mule1. +rewrite gt_eqF ?lte_fin//=. +rewrite integral_dirac//= diracT mul1e. +by rewrite muleAC -EFinM divff ?gt_eqF// mul1e bernoulliE. +Qed. + +Lemma prog45 : execD (@prog4 R) = execD (@prog5 R). +Proof. by rewrite normalize_score_bernoulli//; lra. Qed. + +End from_prog4_to_prog5. + +Lemma from_prog0_to_prog5 {R : realType} : execD (@prog0 R) = execD (@prog5 R). +Proof. +rewrite prog01 prog12 prog22' prog23. +rewrite -prog45. +apply: congr_normalize => y V. +apply: congr_letinr => x U. +by rewrite -prog34. +Qed. diff --git a/theories/normedtype_theory/ereal_normedtype.v b/theories/normedtype_theory/ereal_normedtype.v index 33a6aa13b2..16bca15ce0 100644 --- a/theories/normedtype_theory/ereal_normedtype.v +++ b/theories/normedtype_theory/ereal_normedtype.v @@ -5,9 +5,9 @@ From mathcomp Require Import rat interval zmodp vector fieldext falgebra. From mathcomp Require Import archimedean. From mathcomp Require Import mathcomp_extra unstable boolp classical_sets. From mathcomp Require Import functions cardinality set_interval. -From mathcomp Require Import interval_inference ereal reals topology. -From mathcomp Require Import separation_axioms function_spaces real_interval. -From mathcomp Require Import prodnormedzmodule tvs. +From mathcomp Require Import interval_inference ereal reals constructive_ereal. +From mathcomp Require Import topology separation_axioms function_spaces. +From mathcomp Require Import real_interval prodnormedzmodule tvs. From mathcomp Require Import num_normedtype. (**md**************************************************************************) diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 54aa27cf6f..07bc09b58a 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -1,43 +1,44 @@ (* mathcomp analysis (c) 2022 Inria and AIST. License: CeCILL-C. *) From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. -From mathcomp Require Import interval_inference archimedean rat. -From mathcomp Require Import mathcomp_extra boolp classical_sets functions. -From mathcomp Require Import cardinality fsbigop reals ereal. -From mathcomp Require Import topology normedtype sequences. +From mathcomp Require Import rat archimedean. +From mathcomp Require Import unstable mathcomp_extra boolp classical_sets. +From mathcomp Require Import functions cardinality fsbigop interval_inference. +From mathcomp Require Import reals ereal topology normedtype sequences. From mathcomp Require Import esum measure lebesgue_measure numfun. -From mathcomp Require Import lebesgue_integral exp kernel charge. +From mathcomp Require Import lebesgue_integral probability exp kernel charge. From mathcomp Require Import ring lra. (**md**************************************************************************) (* # Semantics of a probabilistic programming language using s-finite kernels *) -(* bernoulli p1 == Bernoulli probability with p1 a proof that *) -(* p : {nonneg R} is smaller than 1 *) -(* bernoulli_trunc r == Bernoulli probability with real number r *) -(* bino_term n k p == $\binom{n}{k}p^k (1-p)^(n-k)$ *) -(* Computes a binomial distribution term for *) -(* k successes in n trials with success rate p *) -(* binomial_probability n p1 == binomial probability with n and p1 a proof *) -(* that p : {nonneg R} is smaller than 1 *) -(* binomial_probability_trunc n r == binomial probability with n and real *) -(* number r *) -(* uniform_probability a b ab0 == uniform probability over the interval [a,b] *) -(* sample mP == sample according to the probability P where mP is a *) -(* proof that P is a measurable function *) -(* letin l k == execute l, augment the context, and execute k *) -(* ret mf == access the context with f and return the result *) -(* score mf == observe t from d, where f is the density of d and *) -(* t occurs in f *) -(* e.g., score (r e^(-r * t)) = observe t from exp(r) *) -(* normalize k P == normalize the kernel k into a probability kernel, *) -(* P is a default probability in case normalization is *) -(* not possible *) -(* ite mf k1 k2 == access the context with the boolean function f and *) -(* behaves as k1 or k2 according to the result *) (* *) -(* poisson == Poisson distribution function *) -(* exp_density == density function for exponential distribution *) +(* Reference: *) +(* - R. Affeldt, C. Cohen, A. Saito. Semantics of probabilistic programs *) +(* using s-finite kernels in Coq. CPP 2023 *) (* *) +(* ``` *) +(* poisson_pdf == Poisson pdf *) +(* exponential_pdf == exponential distribution pdf *) +(* *) +(* sample mP == sample according to the probability P where mP is *) +(* a proof that P is a measurable function *) +(* sample_cst P == sample according to the probability P *) +(* letin l k == execute l, augment the context, and execute k *) +(* ret mf == access the context with f and return the result *) +(* score mf == observe t from d, where f is the density of d and *) +(* t occurs in f *) +(* e.g., score (r e^(-r * t)) = observe t from exp(r) *) +(* normalize k P == normalize the kernel k into a probability kernel, *) +(* P is a default probability in case normalization *) +(* is not possible *) +(* ite mf k1 k2 == access the context with the boolean function f and *) +(* behaves as k1 or k2 according to the result *) +(* case_nat == case analysis on the nat datatype *) +(* *) +(* mkswap k == given a kernel k : (Y * X) ~> Z, *) +(* returns a kernel of type (X * Y) ~> Z *) +(* letin' := mkcomp \o mkswap *) +(* ``` *) (******************************************************************************) Set Implicit Arguments. @@ -50,45 +51,6 @@ Local Open Scope classical_set_scope. Local Open Scope ring_scope. Local Open Scope ereal_scope. -(* Definition mR (R : realType) : Type := R. -HB.instance Definition _ (R : realType) := Measurable.on (mR R). -(* [the measurableType (R.-ocitv.-measurable).-sigma of - salgebraType (R.-ocitv.-measurable)]. *) *) - -Module Notations. -(*Notation var1of2 := (@measurable_fst _ _ _ _). -Notation var2of2 := (@measurable_snd _ _ _ _). -Notation var1of3 := (measurableT_comp (@measurable_fst _ _ _ _) - (@measurable_fst _ _ _ _)). -Notation var2of3 := (measurableT_comp (@measurable_snd _ _ _ _) - (@measurable_fst _ _ _ _)). -Notation var3of3 := (@measurable_snd _ _ _ _).*) - -(* Definition mR R := [the measurableType (R.-ocitv.-measurable).-sigma of - salgebraType (R.-ocitv.-measurable)]. *) -Notation munit := (unit : measurableType _). -Notation mbool := (bool : measurableType _). -Notation mnat := (nat : measurableType _). -End Notations. - -(* TODO: PR *) -Lemma onem_nonneg_proof (R : numDomainType) (p : {nonneg R}) : - (p%:num <= 1 -> 0 <= `1-(p%:num))%R. -Proof. by rewrite /onem/= subr_ge0. Qed. - -Definition onem_nonneg (R : numDomainType) (p : {nonneg R}) - (p1 : (p%:num <= 1)%R) := - NngNum (onem_nonneg_proof p1). -(* /TODO: PR *) - -Lemma invr_nonneg_proof (R : numDomainType) (p : {nonneg R}) : - (0 <= (p%:num)^-1)%R. -Proof. by rewrite invr_ge0. Qed. - -Definition invr_nonneg (R : numDomainType) (p : {nonneg R}) := - NngNum (invr_nonneg_proof p). - -(* TODO: move *) Lemma eq_probability R d (Y : measurableType d) (m1 m2 : probability Y R) : (m1 =1 m2 :> (set Y -> \bar R)) -> m1 = m2. Proof. @@ -114,249 +76,87 @@ subst p2. by f_equal. Qed. -Section constants. -Variable R : realType. -Local Open Scope ring_scope. - -Lemma onem1S n : `1- (1 / n.+1%:R) = (n%:R / n.+1%:R)%:nng%:num :> R. -Proof. -by rewrite /onem/= -{1}(@divrr _ n.+1%:R) ?unitfE// -mulrBl -natr1 addrK. -Qed. - -Lemma p1S n : (1 / n.+1%:R)%:nng%:num <= 1 :> R. -Proof. by rewrite ler_pdivrMr//= mul1r ler1n. Qed. - -Lemma p12 : (1 / 2%:R)%:nng%:num <= 1 :> R. Proof. by rewrite p1S. Qed. - -Lemma p14 : (1 / 4%:R)%:nng%:num <= 1 :> R. Proof. by rewrite p1S. Qed. - -Lemma onem27 : `1- (2 / 7%:R) = (5%:R / 7%:R)%:nng%:num :> R. -Proof. by apply/eqP; rewrite subr_eq/= -mulrDl -natrD divrr// unitfE. Qed. - -Lemma p27 : (2 / 7%:R)%:nng%:num <= 1 :> R. -Proof. by rewrite /= lter_pdivrMr// mul1r ler_nat. Qed. - -End constants. -Arguments p12 {R}. -Arguments p14 {R}. -Arguments p27 {R}. -Arguments p1S {R}. - -Section bernoulli. -Variables (R : realType) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). - -Definition bernoulli : set bool -> \bar R := - measure_add - [the measure _ _ of mscale p [the measure _ _ of dirac true]] - [the measure _ _ of mscale (onem_nonneg p1) [the measure _ _ of dirac false]]. - -HB.instance Definition _ := Measure.on bernoulli. - -Let bernoulli_setT : bernoulli [set: _] = 1. -Proof. -rewrite /bernoulli/= /measure_add/= /msum 2!big_ord_recr/= big_ord0 add0e/=. -by rewrite /mscale/= !diracT !mule1 -EFinD add_onemK. -Qed. - -HB.instance Definition _ := - @Measure_isProbability.Build _ _ R bernoulli bernoulli_setT. - -End bernoulli. - -Lemma integral_bernoulli {R : realType} (p : {nonneg R}) (p1 : (p%:num <= 1)%R) - (f : bool -> \bar R) : (forall x, 0 <= f x) -> - \int[bernoulli p1]_y (f y) = - p%:num%:E * f true + (`1-(p%:num))%:E * f false. -Proof. -move=> f0. -rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. -by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT 2!mul1e. -Qed. +Definition dep_uncurry (A : Type) (B : A -> Type) (C : Type) : + (forall a : A, B a -> C) -> {a : A & B a} -> C := + fun f p => let (a, Ba) := p in f a Ba. -Section bernoulli_trunc. -Variables (R : realType). +Section binomial_total. Local Open Scope ring_scope. +Variables (R : realType) (n : nat). +Implicit Type p : R. -Lemma sumbool_ler (x y : R) : (x <= y)%R + (x > y)%R. -Proof. -have [_|_] := leP x y. -by apply (*left*) inl. -by apply (*right*) inr. -Qed. - -(* TODO: move? *) -Definition dep_uncurry := -(fun (A : Type) (B : A -> Type) (C : Type) (f : forall a : A, B a -> C) (p : {a : A & B a}) => let (a, Ba) := p in f a Ba) : -forall [A : Type] [B : A -> Type] [C : Type], -(forall a : A, B a -> C) -> {a : A & B a} -> C. - -Definition bernoulli0 := @bernoulli R 0%R%:nng ler01. - -HB.instance Definition _ := Probability.on bernoulli0. - -Definition bernoulli_trunc (p : R) := match sumbool_ler 0%R p with -| inl l0p => match sumbool_ler (NngNum l0p)%:num 1%R with - | inl lp1 => [the probability _ _ of @bernoulli R (NngNum l0p) lp1] - | inr _ => bernoulli0 - end -| inr _ => bernoulli0 -end. - -Lemma bernoulli_truncE (p : R) U : - (0 <= p <= 1)%R -> - (bernoulli_trunc p U = - p%:E * \d_true U + (`1-p)%:E * \d_false U)%E. -Proof. -move=> /andP[p0 p1]. -rewrite /bernoulli_trunc. -case: (sumbool_ler 0 p) => [{}p0/=|]. - case: (sumbool_ler p 1) => [{}p1/=|]. - by rewrite /bernoulli/= measure_addE. - by rewrite ltNge p1. -by rewrite ltNge p0. -Qed. - -(* HB.instance Definition _ (p : R) := Probability.on (bernoulli_trunc p). *) - -Let simpe := (@mule0 R, @adde0 R, @mule1 R, @add0e R). - -Lemma measurable_bernoulli_trunc : - measurable_fun setT (bernoulli_trunc : _ -> pprobability _ _). +Lemma measurable_binomial_prob : + measurable_fun setT (binomial_prob n : R -> pprobability _ _). Proof. apply: (@measurability _ _ _ _ _ _ (@pset _ _ _ : set (set (pprobability _ R)))) => //. move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. -rewrite /bernoulli_trunc/=. -have := @subsetT _ Ys; rewrite setT_bool => UT. -have [->|->|->|->] /= := subset_set2 UT. -- rewrite [X in measurable_fun _ X](_ : _ = cst 0%E)//. - apply/funext => x/=. - by case: sumbool_ler. -- rewrite [X in measurable_fun _ X](_ : _ = (fun x => if 0 <= x <= 1 then x%:E else 0%E))//. - by apply: measurable_fun_ifT => //=; apply: measurable_and => //; exact: measurable_fun_ler. - apply/funext => x/=; case: sumbool_ler => /= x0. - case: sumbool_ler => /= x1. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. - by rewrite /mscale/= !diracE mem_set//= memNset//= ?simpe x0 x1. - rewrite /bernoulli0 /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. - by rewrite /mscale/= !diracE mem_set//= memNset//= ?simpe x0/= leNgt x1. - rewrite /bernoulli0 /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. - by rewrite /mscale/= !diracE mem_set//= memNset//= ?simpe leNgt x0. -- rewrite [X in measurable_fun _ X](_ : _ = - (fun x => if 0 <= x <= 1 then (`1-x)%:E else 1%E))//. - apply: measurable_fun_ifT => //=. - by apply: measurable_and => //; exact: measurable_fun_ler. - by apply/measurable_EFinP; apply/measurable_funB. - apply/funext => x/=; case: sumbool_ler => /= x0. - case: sumbool_ler => /= x1. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. - by rewrite /mscale/= !diracE memNset//= mem_set//= ?simpe x0 x1/=. - rewrite /bernoulli0 /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. - by rewrite /mscale/= !diracE memNset//= mem_set//= ?simpe x0/= leNgt x1/= onem0. - rewrite /bernoulli0. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. - by rewrite /mscale/= !diracE memNset//= mem_set//= leNgt x0/= ?simpe onem0. -- rewrite [X in measurable_fun _ X](_ : _ = cst 1%E)//; apply/funext => x/=. - case: sumbool_ler => /= x0. - case: sumbool_ler => /= x1. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. - rewrite /mscale/= !diracE mem_set//=; last by left. - rewrite mem_set//=; last by right. - by rewrite ?simpe -EFinD add_onemK. - rewrite /bernoulli0. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. - rewrite /mscale/= !diracE mem_set//=; last by left. - rewrite mem_set//=; last by right. - by rewrite ?simpe onem0. - rewrite /bernoulli0. - rewrite /bernoulli/= /measure_add/= /msum/= !big_ord_recl//= big_ord0//=. - rewrite /mscale/= !diracE mem_set//=; last by left. - rewrite mem_set//=; last by right. - by rewrite ?simpe onem0. -Qed. - -End bernoulli_trunc. - -Arguments bernoulli_trunc {R}. -Arguments measurable_bernoulli_trunc {R}. - -Lemma integral_bernoulli_trunc {R : realType} (p : R) (f : bool -> \bar R) : - (0 <= p <= 1)%R -> (forall x, 0 <= f x) -> - \int[bernoulli_trunc p]_y (f y) = p%:E * f true + (`1-p)%:E * f false. -Proof. -move=> /andP[p0 p1] f0; rewrite /bernoulli_trunc. -case: sumbool_ler => [? /=|]. - case: (sumbool_ler p 1) => [? /=|]. - by rewrite integral_bernoulli. - by rewrite ltNge p1. -by rewrite ltNge p0. -Qed. - -Section binomial_probability. -Context {R : realType} (n : nat) (p : {nonneg R}) (p1 : (p%:num <= 1)%R). +rewrite /binomial_prob/=. +set f := (X in measurable_fun _ X). +apply: measurable_fun_if => //=. + by apply: measurable_and => //; exact: measurable_fun_ler. +apply: (eq_measurable_fun (fun t => + \sum_(k x /set_mem[_/= x01]. + rewrite nneseries_esum// -1?[in RHS](set_mem_set Ys)// => k kYs. + by rewrite lee_fin binomial_pmf_ge0. +apply: ge0_emeasurable_sum. + by move=> k x/= [_ x01] _; rewrite lee_fin binomial_pmf_ge0. +move=> k Ysk; apply/measurableT_comp => //. +exact: measurable_binomial_pmf. +Qed. + +End binomial_total. +Arguments measurable_binomial_prob {R}. + +Section poisson_pdf. +Variable R : realType. Local Open Scope ring_scope. -(* C(n, k) * p^k * (1-p)^(n-k) *) -Definition bino_term (k : nat) : {nonneg R} := - (p%:num^+k * (NngNum (onem_ge0 p1))%:num^+(n-k)%N *+ 'C(n, k))%:nng. +Definition poisson_pdf k r : R := + if r > 0 then r ^+ k / k`!%:R^-1 * expR (- r) else 1%:R. -Lemma bino_term0 : - bino_term 0 = ((NngNum (onem_ge0 p1))%:num^+n)%:nng. +Lemma poisson_pdf_ge0 k r : 0 <= poisson_pdf k r. Proof. -rewrite /bino_term bin0 subn0/=. -apply/val_inj => /=. -by field. +rewrite /poisson_pdf; case: ifPn => r0//. +by rewrite mulr_ge0 ?expR_ge0// mulr_ge0// exprn_ge0 ?ltW. Qed. -Lemma bino_term1 : - bino_term 1 = (p%:num * (NngNum (onem_ge0 p1))%:num^+(n-1)%N *+ n)%:nng. +Lemma poisson_pdf_gt0 k r : 0 < r -> 0 < poisson_pdf k.+1 r. Proof. -rewrite /bino_term bin1/=. -apply/val_inj => /=. -by rewrite expr1. +move=> r0; rewrite /poisson_pdf r0 mulr_gt0 ?expR_gt0//. +by rewrite divr_gt0// ?exprn_gt0// invr_gt0 ltr0n fact_gt0. Qed. -Import Notations. - -(* Check \sum_(k < n.+1) (fun k => [the measure _ _ of mscale (bino_term k) - [the measure _ _ of \d_k]]). *) -(* \sum_(k < n.+1) (bino_coef p n k) * \d_k. *) -Definition binomial_probability : set nat -> \bar R := - @msum _ _ R (fun k => mscale (bino_term k) \d_k) n.+1. - -HB.instance Definition _ := Measure.on binomial_probability. - -Let binomial_setT : binomial_probability [set: _] = 1%:E. +Lemma measurable_poisson_pdf k : measurable_fun setT (poisson_pdf k). Proof. -rewrite /binomial_probability/msum/mscale/bino_term/=/mscale/=. -under eq_bigr do rewrite diracT mule1. -rewrite sumEFin. -under eq_bigr=> i _. - rewrite mulrC. - over. -rewrite -exprDn_comm; last by rewrite /GRing.comm mulrC. -by rewrite addrC add_onemK; congr _%:E; rewrite expr1n. +rewrite /poisson_pdf; apply: measurable_fun_if => //. + exact: measurable_fun_ltr. +by apply: measurable_funM => /=; + [exact: measurable_funM|exact: measurableT_comp]. Qed. -HB.instance Definition _ := - @Measure_isProbability.Build _ _ R binomial_probability binomial_setT. +End poisson_pdf. + +Section exponential_pdf. +Variable R : realType. +Local Open Scope ring_scope. + +Definition exponential_pdf x r : R := r * expR (- r * x). -End binomial_probability. +Lemma exponential_pdf_gt0 x r : 0 < r -> 0 < exponential_pdf x r. +Proof. by move=> r0; rewrite /exponential_pdf mulr_gt0// expR_gt0. Qed. -Section integral_binomial. -Variables (R : realType) (d : measure_display) (T : measurableType d). +Lemma exponential_pdf_ge0 x r : 0 <= r -> 0 <= exponential_pdf x r. +Proof. by move=> r0; rewrite /exponential_pdf mulr_ge0// expR_ge0. Qed. -Lemma integral_binomial (n : nat) (p : {nonneg R}) - (p1 : (p%:num <= 1)%R) (f : nat -> \bar R) - (mf : measurable_fun setT f) : - (forall x, 0 <= f x) -> \int[binomial_probability n p1]_y (f y) = - \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k. +Lemma measurable_exponential_pdf x : measurable_fun setT (exponential_pdf x). Proof. -move=> f0; rewrite ge0_integral_measure_sum//=; apply: eq_bigr => i _. -by rewrite ge0_integral_mscale//= integral_dirac//= diracT mul1e. +apply: measurable_funM => //=; apply: measurableT_comp => //. +exact: measurable_funM. Qed. -End integral_binomial. +End exponential_pdf. (* X + Y is a measurableType if X and Y are *) HB.instance Definition _ (X Y : pointedType) := @@ -402,117 +202,63 @@ Lemma measurable_fun_if_pair_nat d d' (X : measurableType d) Proof. move=> mx my; apply: measurable_fun_ifT => //=. - have h : measurable_fun [set: nat] (fun t => t == n) by []. - exact: (@measurableT_comp _ _ _ _ _ _ _ _ _ h). + exact: (measurableT_comp h). - exact: measurableT_comp. - exact: measurableT_comp. Qed. -Section binomial_trunc. -Variables (R : realType). -Local Open Scope ring_scope. +(* Definition mR (R : realType) : Type := R. +HB.instance Definition _ (R : realType) := Measurable.on (mR R). +(* [the measurableType (R.-ocitv.-measurable).-sigma of + salgebraType (R.-ocitv.-measurable)]. *) *) + +Module Notations. +Notation munit := (unit : measurableType _). +Notation mbool := (bool : measurableType _). +Notation mnat := (nat : measurableType _). +End Notations. -Definition binomial_probability0 := @binomial_probability R 0 0%:nng%R ler01. +Lemma invr_nonneg_proof (R : numDomainType) (p : {nonneg R}) : + (0 <= (p%:num)^-1)%R. +Proof. by rewrite invr_ge0. Qed. -Definition binomial_probability_trunc n (p : R) := - match (sumbool_ler 0%R p) with - | inl l0p => match (sumbool_ler (NngNum l0p)%:num 1%R) with - | inl lp1 => [the probability _ _ of @binomial_probability R n (NngNum l0p) lp1] - | inr _ => [the probability _ _ of binomial_probability0] - end - | inr _ => [the probability _ _ of binomial_probability0] - end. +Definition invr_nonneg (R : numDomainType) (p : {nonneg R}) := + NngNum (invr_nonneg_proof p). -Lemma measurable_binomial_probability_trunc (n : nat) - : measurable_fun setT (binomial_probability_trunc n : R -> pprobability _ _). -Proof. -apply: (@measurability _ _ _ _ _ _ - (@pset _ _ _ : set (set (pprobability _ R)))) => //. -move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. -rewrite /binomial_probability_trunc/=. -set f := (X in measurable_fun _ X). -rewrite (_ : f = fun x => if 0 <= x <= 1 then (\sum_(m < n.+1) - match sumbool_ler 0 x with - | inl l0p => - match sumbool_ler x 1 with - | inl lp1 => mscale (@bino_term _ n (NngNum l0p) lp1 m) (\d_(nat_of_ord m)) Ys - | inr _ => (x ^+ m * `1-x ^+ (n - m) *+ 'C(n, m))%:E * \d_(nat_of_ord m) Ys - end - | inr _ => (x ^+ m * `1-x ^+ (n - m) *+ 'C(n, m))%:E * \d_(nat_of_ord m) Ys - end)%E - else \d_0%N Ys)//. - move=> ?; apply: measurable_fun_ifT => //=. - by apply: measurable_and => //; exact: measurable_fun_ler. - apply: emeasurable_sum => m /=. - rewrite /mscale/= [X in measurable_fun _ X](_ : _ = (fun x => - ((x ^+ m * `1-x ^+ (n - m) *+ 'C(n, m))%:E * \d_(nat_of_ord m) Ys)%E)); last first. - by apply:funext => x; case: sumbool_ler => // x0; case: sumbool_ler. - apply: emeasurable_funM => //; apply/measurable_EFinP => //. - under eq_fun do rewrite -mulrnAr. - apply: measurable_funM => //. - under eq_fun do rewrite -mulr_natr. - apply: measurable_funM => //=. - apply: measurable_funX. - exact: measurable_funB. -rewrite {}/f. -apply/funext => x. -case: sumbool_ler => /= x0. - case: sumbool_ler => /= x1. - by rewrite /binomial_probability/= /msum /= /bino_term/= x0 x1. - rewrite /binomial_probability /= /msum big_ord_recl/= big_ord0 /mscale. - by rewrite /= expr0 mul1r subnn expr0 bin0 mul1e adde0 !leNgt x1 andbF. -rewrite /binomial_probability /= /msum big_ord_recl/= big_ord0 /mscale. -by rewrite /= expr0 mul1r subnn expr0 bin0 mul1e adde0 !leNgt x0/=. -Qed. - -End binomial_trunc. - -Arguments binomial_probability_trunc {R}. -Arguments measurable_binomial_probability_trunc {R}. - -Section integral_binomial_trunc. -Variables (R : realType) (d : measure_display) (T : measurableType d). +Section constants. +Variable R : realType. +Local Open Scope ring_scope. -Import Notations. -Lemma integral_binomial_probabilty_trunc (n : nat) (p : R) - (p0 : (0 <= p)%R) (p1 : ((NngNum p0)%:num <= 1)%R) (f : mnat -> \bar R) - (mf : measurable_fun setT f) : (forall x, 0 <= f x) -> - \int[binomial_probability_trunc n p]_y (f y) = - \sum_(k < n.+1) (bino_term n p1 k)%:num%:E * f k. +Lemma onem1S n : `1- (1 / n.+1%:R) = (n%:R / n.+1%:R)%:nng%:num :> R. Proof. -move=> f0; rewrite /binomial_probability_trunc/=. -case: sumbool_ler => [? /=|]. - case: sumbool_ler => [?/=|]. - by rewrite integral_binomial. - by rewrite ltNge p1. -by rewrite ltNge p0. +by rewrite /onem/= -{1}(@divrr _ n.+1%:R) ?unitfE// -mulrBl -natr1 addrK. Qed. -End integral_binomial_trunc. +Lemma p1S n : (1 / n.+1%:R)%:nng%:num <= 1 :> R. +Proof. by rewrite ler_pdivrMr//= mul1r ler1n. Qed. -Section binomial_example. -Context {R : realType}. -Open Scope ring_scope. +Lemma p12 : (1 / 2%:R)%:nng%:num <= 1 :> R. +Proof. by rewrite ler_pdivrMr//= mul1r ler1n. Qed. -Lemma binomial3_2 : @binomial_probability R 3 _ (p1S 1) [set 2%N] = (3 / 8)%:E. -Proof. -rewrite /binomial_probability/msum !big_ord_recl/= big_ord0 adde0 bino_term0. -rewrite /mscale/= !diracE /bump/=. -repeat rewrite ?binS ?bin0 ?bin1 ?bin_small//. -rewrite memNset//=. -rewrite memNset//=. -rewrite mem_set//=. -rewrite memNset//=. -congr _%:E. -rewrite expr0 !mulr1 !mulr0 !add0r !addn0 !add0n /onem. -by field. -Qed. +Lemma p14 : (1 / 4%:R)%:nng%:num <= 1 :> R. +Proof. by rewrite ler_pdivrMr//= mul1r ler1n. Qed. + +Lemma onem27 : `1- (2 / 7%:R) = (5%:R / 7%:R)%:nng%:num :> R. +Proof. by apply/eqP; rewrite subr_eq/= -mulrDl -natrD divrr// unitfE. Qed. -End binomial_example. +(*Lemma p27 : (2 / 7%:R)%:nng%:num <= 1 :> R. +Proof. by rewrite /= lter_pdivrMr// mul1r ler_nat. Qed.*) + +End constants. +Arguments p12 {R}. +Arguments p14 {R}. +(*Arguments p27 {R}.*) +Arguments p1S {R}. Section uniform_probability. Context (R : realType) (a b : R) (ab0 : (0 < b - a)%R). -Definition uniform_probability : set _ -> \bar R := +Definition uniform_probability (* : set _ -> \bar R *) := @mscale _ _ R (invr_nonneg (NngNum (ltW ab0))) (mrestr lebesgue_measure (measurable_itv `[a, b])). @@ -586,570 +332,14 @@ apply: eq_fsbigr => r _; rewrite ge0_integralZl//. - by rewrite lee_fin invr_ge0// ltW. Qed. -Lemma integral_uniform (f : _ -> \bar R) (a b : R) - (ab0 : (0 < b - a)%R) : measurable_fun setT f -> (forall x, 0 <= f x) -> - let m := uniform_probability ab0 in - \int[m]_x f x = - (b - a)^-1%:E * \int[lebesgue_measure]_(x in `[a, b]) f x. -Proof. -move=> mf f0 m. -pose f_ := nnsfun_approx measurableT mf. -transitivity (lim (\int[m]_(x in setT) (f_ n x)%:E @[n --> \oo])). - rewrite -monotone_convergence//=. - - apply: eq_integral => ? /[!inE] xD; apply/esym/cvg_lim => //=. - exact/cvg_nnsfun_approx. - - by move=> n; exact/measurable_EFinP/measurable_funTS. - - by move=> n ? _; rewrite lee_fin. - - by move=> ? _ ? ? mn; rewrite lee_fin; exact/lefP/nd_nnsfun_approx. -rewrite [X in _ = _ * X](_ : _ = lim - (\int[lebesgue_measure]_(x in `[a, b]) (f_ n x)%:E @[n --> \oo])); last first. - rewrite -monotone_convergence//=. - - by apply: eq_integral => ? /[!inE] xD; apply/esym/cvg_lim => //; exact: cvg_nnsfun_approx. - - by move=> n; exact/measurable_EFinP/measurable_funTS. - - by move=> n ? _; rewrite lee_fin. - - by move=> ? _ ? ? ab; rewrite lee_fin; exact/lefP/nd_nnsfun_approx. -rewrite -limeMl//. - apply: congr_lim. - by apply/funext => n /=; exact: integral_uniform_nnsfun. -apply/ereal_nondecreasing_is_cvgn => ? ? ab; apply: ge0_le_integral => //=. -- by move=> ? _; rewrite lee_fin. -- exact/measurable_EFinP/measurable_funTS. -- by move=> ? _; rewrite lee_fin. -- exact/measurable_EFinP/measurable_funTS. -- by move=> ? _; rewrite lee_fin; exact/lefP/nd_nnsfun_approx. -Qed. - End integral_uniform. -(* normalization constant *) -Section beta_nat_norm. -Context {R : realType} (a b : nat). - -Definition beta_nat_norm : R := - (a.-1`!%:R * b.-1`!%:R / (a + b).-1`!%:R)%R. - -Lemma beta_nat_norm_gt0 : (0 < beta_nat_norm :> R)%R. -Proof. -by rewrite /beta_nat_norm divr_gt0// ?mulr_gt0 ?ltr0n ?fact_gt0. -Qed. - -Lemma beta_nat_norm_ge0 : (0 <= beta_nat_norm :> R)%R. -Proof. exact/ltW/beta_nat_norm_gt0. Qed. - -End beta_nat_norm. - -Lemma beta_nat_norm00 {R : realType} : beta_nat_norm 0 0 = 1%R :> R. -Proof. by rewrite /beta_nat_norm/= fact0 mulr1/= divff. Qed. - -Lemma beta_nat_norm11 {R : realType} : beta_nat_norm 1 1 = 1%R :> R. -Proof. by rewrite /beta_nat_norm/= fact0 mulr1/= divff. Qed. - -(* definition of the beta probability specialized to natural numbers *) - -Definition ubeta_nat_pdf' {R : realType} (a b : nat) (t : R) := - (t ^+ a * (`1- t) ^+ b)%R. - -Section beta_probability. -Local Open Scope ring_scope. -Context {R : realType}. -Variables a b : nat. -(* unnormalized pdf for beta specialized to nat *) - -Definition ubeta_nat_pdf (t : R) := ubeta_nat_pdf' a.-1 b.-1 t. - -Lemma ubeta_nat_pdf_ge0 t : 0 <= t <= 1 -> 0 <= ubeta_nat_pdf t. -Proof. by move=> /andP[t0 t1]; rewrite mulr_ge0// exprn_ge0// onem_ge0. Qed. - -Lemma ubeta_nat_pdf_le1 t : 0 <= t <= 1 -> ubeta_nat_pdf t <= 1. -Proof. -move=> /andP[t0 t1]; rewrite /ubeta_nat_pdf. -by rewrite mulr_ile1// ?(exprn_ge0,onem_ge0,exprn_ile1,onem_le1). -Qed. - -Lemma measurable_ubeta_nat_pdf : measurable_fun setT ubeta_nat_pdf. -Proof. -by apply /measurable_funM => //; exact/measurable_funX/measurable_funB. -Qed. - -Lemma beta_nat_normE : - (beta_nat_norm a b)%:E = (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (ubeta_nat_pdf x)%:E)%E. -Proof. -rewrite /beta_nat_norm. -rewrite /ubeta_nat_pdf. -rewrite /ubeta_nat_pdf'. -Admitted. - -(* normalized pdf for beta specialized to nat *) -Definition beta_nat_pdf t := ubeta_nat_pdf t / (beta_nat_norm a b). - -Lemma measurable_beta_nat_pdf : measurable_fun setT beta_nat_pdf. -Proof. by apply: measurable_funM => //; exact: measurable_ubeta_nat_pdf. Qed. - -Lemma beta_nat_pdf_ge0 t : 0 <= t <= 1 -> 0 <= beta_nat_pdf t. -Proof. -move=> t01; rewrite /beta_nat_pdf divr_ge0//. - exact: ubeta_nat_pdf_ge0. -exact: beta_nat_norm_ge0. -Qed. - -Local Notation mu := lebesgue_measure. - -(* unnormalized beta specialized to nat *) -Definition ubeta_nat (U : set (measurableTypeR R)) : \bar R := - \int[mu]_(x in U `&` `[0, 1](*NB: is this correct?*)) (ubeta_nat_pdf x)%:E. -(* TODO: define as \int[uniform_probability p01]_(t in U) (ubeta_nat_pdf t)%:E ? *) - -Lemma ubeta_natE U : - (ubeta_nat U = - \int[mu]_(x in U `&` `[0%R, 1%R]) (ubeta_nat_pdf x)%:E :> \bar R)%E. -Proof. by []. Qed. - -Lemma ubeta_nat_lty U : (ubeta_nat U < +oo)%E. -Proof. -Admitted. - -Let ubeta_nat0 : ubeta_nat set0 = 0%:E. -Proof. by rewrite /ubeta_nat set0I integral_set0. Qed. - -Let ubeta_nat_ge0 U : (0 <= ubeta_nat U)%E. -Proof. -rewrite /ubeta_nat integral_ge0//= => x [Ux]. -by rewrite in_itv/= => x01; rewrite lee_fin ubeta_nat_pdf_ge0. -Qed. - -Let ubeta_nat_sigma_additive : semi_sigma_additive ubeta_nat. -Proof. -move=> /= F mF tF mUF; rewrite /ubeta_nat setI_bigcupl; apply: cvg_toP. - apply: ereal_nondecreasing_is_cvgn => m n mn. - apply: lee_sum_nneg_natr => // k _ _. - apply: integral_ge0 => /= x [_]; rewrite in_itv => x01. - by rewrite lee_fin; exact: ubeta_nat_pdf_ge0. -rewrite ge0_integral_bigcup//=. -- by move=> k; exact: measurableI. -- by apply/measurable_EFinP; exact: measurable_funTS measurable_ubeta_nat_pdf. -- move=> x [k _ [_]]; rewrite /= in_itv/= => x01. - by rewrite lee_fin ubeta_nat_pdf_ge0. -- exact: trivIset_setIr. -Qed. - -HB.instance Definition _ := isMeasure.Build _ _ _ ubeta_nat - ubeta_nat0 ubeta_nat_ge0 ubeta_nat_sigma_additive. - -Definition beta_nat (*: set [the measurableType (R.-ocitv.-measurable).-sigma of - salgebraType R.-ocitv.-measurable] -> \bar R*) := - @mscale _ _ _ (invr_nonneg (NngNum (beta_nat_norm_ge0 a b))) ubeta_nat. - -Let beta_nat0 : beta_nat set0 = 0. -Proof. exact: measure0. Qed. - -Let beta_nat_ge0 U : (0 <= beta_nat U)%E. -Proof. exact: measure_ge0. Qed. - -Let beta_nat_sigma_additive : semi_sigma_additive beta_nat. -Proof. move=> /= F mF tF mUF; exact: measure_semi_sigma_additive. Qed. - -HB.instance Definition _ := isMeasure.Build _ _ _ beta_nat - beta_nat0 beta_nat_ge0 beta_nat_sigma_additive. - -Let beta_nat_setT : beta_nat setT = 1%:E. -Proof. -rewrite /beta_nat /= /mscale /=. -rewrite /ubeta_nat/= setTI. -by rewrite -beta_nat_normE -EFinM mulVr// unitfE gt_eqF// beta_nat_norm_gt0. -Qed. - -HB.instance Definition _ := @Measure_isProbability.Build _ _ _ - beta_nat beta_nat_setT. - -Lemma beta_nat01 : beta_nat `[0, 1] = 1%:E. -Proof. -rewrite /beta_nat /= /mscale/=. -rewrite /beta_nat /ubeta_nat setIidr//. -by rewrite -beta_nat_normE -EFinM mulVr// unitfE gt_eqF// beta_nat_norm_gt0. -Qed. - -End beta_probability. - -Arguments beta_nat {R}. - -Section beta_probability11. -Local Open Scope ring_scope. -Context {R : realType}. - -Lemma ubeta_nat_pdf11 : ubeta_nat_pdf 1 1 = @cst R _ 1. -Proof. by apply/funext => r; rewrite /ubeta_nat_pdf/= /ubeta_nat_pdf' !expr0 mulr1. Qed. - -Let a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. - -Lemma beta11_uniform U : measurable U -> - beta_nat 1 1 U = uniform_probability a01 U. -Proof. -move=> mU; rewrite /beta_nat /uniform_probability. -rewrite /mscale/= beta_nat_norm11 subr0 invr1 !mul1e. -rewrite /ubeta_nat /mrestr/=. -rewrite ubeta_nat_pdf11/= integral_cst/= ?mul1e//. -exact: measurableI. -Qed. - -End beta_probability11. - -Lemma factD n m : (n`! * m`! <= (n + m).+1`!)%N. -Proof. -elim: n m => /= [m|n ih m]. - by rewrite fact0 mul1n add0n factS leq_pmull. -rewrite 2!factS [in X in (_ <= _ * X)%N]addSn -mulnA leq_mul//. -by rewrite ltnS addSnnS leq_addr. -Qed. - -Lemma factD' n m : (n`! * m.-1`! <= (n + m)`!)%N. -Proof. -case: m => //= [|m]. - by rewrite fact0 muln1 addn0. -by rewrite addnS factD. -Qed. - -Lemma leq_prod2 (x y n m : nat) : (n <= x)%N -> (m <= y)%N -> - (\prod_(m <= i < y) i * \prod_(n <= i < x) i <= \prod_(n + m <= i < x + y) i)%N. -Proof. -move=> nx my. -rewrite big_addn. -rewrite -addnBA//. -rewrite {3}/index_iota. -rewrite -addnBAC//. -rewrite iotaD. -rewrite big_cat/=. -rewrite mulnC. -rewrite leq_mul//. - rewrite /index_iota. - apply: leq_prod. - by move=> i _; rewrite leq_addr. -rewrite subnKC//. -rewrite -{1}(add0n m). -rewrite big_addn. -rewrite {2}(_ : (y - m) = ((y - m + x) - x))%N; last first. - by rewrite -addnBA// subnn addn0. -rewrite -{1}(add0n x). -rewrite big_addn. -rewrite -addnBA// subnn addn0. -apply: leq_prod => i _. -by rewrite leq_add2r leq_addr. -Qed. - -Lemma leq_fact2 (x y n m : nat) : - (n <= x) %N -> (m <= y)%N -> - (x`! * y`! * ((n + m).+1)`! <= n`! * m`! * ((x + y).+1)`!)%N. -Proof. -move=> nx my. -rewrite (_ : x`! = n`! * \prod_(n.+1 <= i < x.+1) i)%N; last first. - by rewrite -fact_split. -rewrite -!mulnA leq_mul2l; apply/orP; right. -rewrite (_ : y`! = m`! * \prod_(m.+1 <= i < y.+1) i)%N; last first. - by rewrite -fact_split. -rewrite mulnCA -!mulnA leq_mul2l; apply/orP; right. -rewrite (_ : (x + y).+1`! = (n + m).+1`! * \prod_((n + m).+2 <= i < (x + y).+2) i)%N; last first. - rewrite -fact_split//. - by rewrite ltnS leq_add. -rewrite mulnA mulnC leq_mul2l; apply/orP; right. -rewrite -addSn -addnS. -rewrite -addSn -addnS. -exact: leq_prod2. -Qed. - -Section integral_beta. -Context {R : realType}. -Variables a b a' b' : nat. - -Local Notation mu := lebesgue_measure. - -Lemma integralMl f g1 g2 A : -measurable A -> measurable_fun A f -> - measurable_fun A g1 -> measurable_fun A g2 -> -(ae_eq mu A g1 (EFin \o g2)) -> - \int[mu]_(x in A) (f x * g1 x) = - \int[mu]_(x in A) (f x * (g2 x)%:E) :> \bar R. -Proof. -move=> mA mf mg1 mg2 Hg. -apply: ae_eq_integral => //. - by apply: emeasurable_funM. - apply: emeasurable_funM => //. - by apply/measurable_EFinP. -by apply: ae_eq_mul2l. -Qed. - -Let beta_nat_dom : (@beta_nat R a b `<< mu). -Proof. -move=> A mA muA0. -rewrite /beta_nat /mscale/= /ubeta_nat. -have -> : \int[mu]_(x0 in A `&` `[0%R, 1%R]) (ubeta_nat_pdf a b x0)%:E = 0%:E. - apply/eqP; rewrite eq_le. - apply/andP; split; last first. - apply: integral_ge0 => x [Ax /=]. - rewrite in_itv /= => x01. - by rewrite lee_fin ubeta_nat_pdf_ge0. - apply: le_trans. - apply: (@ge0_subset_integral _ _ _ _ _ A). - by apply: measurableI. - by []. - apply/measurable_EFinP. - apply: (@measurable_funS _ _ _ _ setT) => //=. - apply: measurable_ubeta_nat_pdf. - move=> x Ax. - have : (`[0%R, 1%R]%classic x). - admit. - rewrite /= in_itv/=. - apply: ubeta_nat_pdf_ge0. - apply: subIsetl. - rewrite /=. - (* rewrite integral_abs_eq0. *) (* without abs *) - admit. -by rewrite mule0. -Admitted. - -Lemma integral_beta_nat f : - measurable_fun setT f -> - \int[beta_nat a b]_(x in `[0%R, 1%R]) `|f x| < +oo -> - \int[beta_nat a b]_(x in `[0%R, 1%R]) f x = - \int[mu]_(x in `[0%R, 1%R]) (f x * (beta_nat_pdf a b x)%:E) :> \bar R. -Proof. -move=> mf finf. -rewrite -(Radon_Nikodym_change_of_variables beta_nat_dom) //=. -apply: integralMl => //. - apply: (@measurable_funS _ _ _ _ [set: R]) => //. - apply: (@measurable_funS _ _ _ _ [set: R]) => //. - rewrite Radon_NikodymE. - by exact: beta_nat_dom. - move=> /= H. - case: cid => /= h [h1 h2 h3]. - have : (measurable_fun setT h /\ \int[mu]_x `|h x| < +oo). - apply/integrableP/h2. - move=> /= [mh _]. - apply: mh. - apply: (@measurable_funS _ _ _ _ [set: R]) => //. - apply: measurable_beta_nat_pdf. - rewrite Radon_NikodymE => /= A. - by exact: beta_nat_dom. -case: cid => /= h [h1 h2 h3]. -apply: integral_ae_eq => //. - apply: integrableS h2 => //. (* integrableST? *) - apply: (@measurable_funS _ _ _ _ [set: R]) => //. - apply: measurableT_comp => //. - apply: measurable_beta_nat_pdf. - move=> E E01 mE. - have mB : measurable_fun E (EFin \o ubeta_nat_pdf a b). - apply: measurableT_comp => //. - apply: (@measurable_funS _ _ _ _ [set: R]) => //. - apply: measurable_ubeta_nat_pdf. - rewrite -(h3 _ mE). - rewrite /beta_nat/mscale/ubeta_nat/beta_nat_pdf/=. - under eq_integral do rewrite mulrC EFinM. - rewrite (integralZl mE). - rewrite /ubeta_nat setIidl //. - rewrite /=. - apply/integrableP; split. - by apply: mB. - under eq_integral => x x01. - rewrite gee0_abs /=. - over. - apply: ubeta_nat_pdf_ge0. - have : x \in `[0%R, 1%R]. - apply: (@subset_trans _ _ `[x,x] _ _ E01). - by rewrite set_interval.set_itv1 sub1set x01. - by rewrite /= in_itv/= lexx. - by rewrite in_itv/=. - rewrite /=. - have <- := (setIidl E01). - by rewrite -ubeta_natE ubeta_nat_lty. -apply/integrableP; split. - by apply: (@measurable_funS _ _ _ _ [set: R]). -exact: finf. -Qed. - -Local Open Scope ring_scope. - -(* TODO: `[0, 1]? *) -Definition beta_nat_bern U : \bar R := - \int[beta_nat a b]_y bernoulli_trunc (ubeta_nat_pdf a'.+1 b'.+1 y) U. - -Local Notation B := beta_nat_norm. - -Definition Baa'bb'Bab : R := (beta_nat_norm (a + a') (b + b')) / beta_nat_norm a b. - -Lemma Baa'bb'Bab_ge0 : 0 <= Baa'bb'Bab. -Proof. by rewrite /Baa'bb'Bab divr_ge0// beta_nat_norm_ge0. Qed. - -Definition Baa'bb'Bab_nneg : {nonneg R} := NngNum Baa'bb'Bab_ge0. - -Lemma Baa'bb'Bab_le1 : Baa'bb'Bab_nneg%:num <= 1. -Proof. -rewrite /Baa'bb'Bab_nneg/= /Baa'bb'Bab. -rewrite ler_pdivrMr// ?mul1r ?beta_nat_norm_gt0//. -rewrite /B /beta_nat_norm. -rewrite ler_pdivrMr ?ltr0n ?fact_gt0//. -rewrite mulrAC. -rewrite ler_pdivlMr ?ltr0n ?fact_gt0//. -rewrite -!natrM ler_nat. -case: a. - rewrite /= fact0 mul1n !add0n. - case: b => /=. - case: a' => //. - case: b' => //= m. - by rewrite fact0 !mul1n muln1. - move=> n/=. - by rewrite fact0 add0n muln1 mul1n factD'. - move=> m. - rewrite mulnC leq_mul// mulnC. - by rewrite (leq_trans (factD' _ _))// addSn addnS//= addnC. -move=> n. -rewrite addSn. -case: b. - rewrite !fact0 add0n muln1 [leqRHS]mulnC addn0/= leq_mul//. - by rewrite factD'. -move=> m. -clear a b. -rewrite [(n + a').+1.-1]/=. -rewrite [n.+1.-1]/=. -rewrite [m.+1.-1]/=. -rewrite addnS. -rewrite [(_ + m).+1.-1]/=. -rewrite (addSn m b'). -rewrite [(m + _).+1.-1]/=. -rewrite (addSn (n + a')). -rewrite [_.+1.-1]/=. -rewrite addSn addnS. -by rewrite leq_fact2// leq_addr. -Qed. - -Lemma onem_Baa'bb'Bab_ge0 : 0 <= 1 - (Baa'bb'Bab_nneg%:num). -Proof. by rewrite subr_ge0 Baa'bb'Bab_le1. Qed. - -Lemma onem_Baa'bb'Bab_ge0_fix : 0 <= B a b * (1 - Baa'bb'Bab_nneg%:num). -Proof. -rewrite mulr_ge0//. - rewrite /B. - exact: beta_nat_norm_ge0. -rewrite subr_ge0. -exact: Baa'bb'Bab_le1. -Qed. - -Lemma ubeta_nat_pdf_ge0' t : 0 <= ubeta_nat_pdf a'.+1 b'.+1 t :> R. -Proof. -apply: ubeta_nat_pdf_ge0. (* TODO: needs 0 <= t <= 1 *) -Admitted. - -Lemma ubeta_nat_pdf_le1' t : (NngNum (ubeta_nat_pdf_ge0' t))%:num <= 1 :> R. -Proof. -rewrite /=. -rewrite /ubeta_nat_pdf. -rewrite /ubeta_nat_pdf'. (* TODO: needs 0 <= t <= 1 *) -Admitted. - -Lemma integral_ubeta_nat' : - (\int[ubeta_nat a b]_x (ubeta_nat_pdf a'.+1 b'.+1 x)%:E = - \int[mu]_(x in `[0%R, 1%R]) - (x ^+ a'.-1 * `1-x ^+ b'.-1 * x ^+ a * `1-x ^+ b)%:E :> \bar R)%E. -Proof. -rewrite /ubeta_nat/ubeta_nat_pdf. -Admitted. - -Lemma beta_nat_bern_bern U : - (a > 0)%N -> (b > 0)%N -> - beta_nat_bern U = - bernoulli_trunc Baa'bb'Bab U. -Proof. -rewrite /beta_nat_bern. -transitivity ((\int[beta_nat a b]_(y in `[0%R, 1%R]) - bernoulli_trunc (ubeta_nat_pdf a'.+1 b'.+1 y) U)%E : \bar R). - admit. -rewrite integral_beta_nat /=; last 2 first. - apply: (@measurableT_comp _ _ _ _ _ _ (bernoulli_trunc ^~ U)). - apply: (measurability (ErealGenInftyO.measurableE R)) => //=. - move=> /= _ [_ [x ->] <-]; apply: measurableI => //. - admit. - exact: measurable_ubeta_nat_pdf. - admit. -under eq_integral => x. - rewrite inE/= in_itv/= => x01. - rewrite bernoulli_truncE. - over. - apply/andP; split. - apply/ubeta_nat_pdf_ge0/x01. - apply/ubeta_nat_pdf_le1/x01. -rewrite /=. -rewrite bernoulli_truncE; last first. - apply/andP; split. - exact: Baa'bb'Bab_ge0. - exact: Baa'bb'Bab_le1. -under eq_integral => x _. - rewrite muleC muleDr//. - over. -rewrite integralD//=; last 2 first. - (* TODO: integrableM *) - admit. - admit. -congr (_ + _). - under eq_integral do rewrite muleA muleC. - rewrite integralZl//=; last first. - admit. - rewrite muleC. - congr (_ * _)%E. - rewrite /beta_nat_pdf. - under eq_integral do rewrite EFinM -muleA muleC -muleA. - rewrite integralZl//=; last first. - admit. - transitivity (((beta_nat_norm a b)^-1)%:E * \int[lebesgue_measure]_(x in `[0%R, 1%R]) ((ubeta_nat_pdf (a+a') (b+b') x)%:E) : \bar R)%E. - congr (_ * _)%E. - apply: eq_integral => x x01. - rewrite /ubeta_nat_pdf /ubeta_nat_pdf' muleC /onem -EFinM/=. - rewrite mulrCA -mulrA -exprD mulrA -exprD. - congr (_ ^+ _ * _ ^+ _)%:E. - rewrite -!subn1 subDnCA//. - rewrite addnC -!subn1 subDnCA//. - rewrite -beta_nat_normE. - rewrite /Baa'bb'Bab/B -!EFinM. - congr _%:E. - rewrite mulrC//. -under eq_integral do rewrite muleA muleC. -rewrite integralZl//=; last first. - admit. -rewrite muleC. -congr (_ * _)%E. -rewrite /beta_nat_pdf. -under eq_integral do rewrite EFinM -muleA muleC -muleA. -rewrite integralZl//=; last first. - admit. -transitivity (((beta_nat_norm a b)^-1)%:E * \int[lebesgue_measure]_(x in `[0%R, 1%R]) ((ubeta_nat_pdf a b x)%:E - (ubeta_nat_pdf (a+a') (b+b') x)%:E) : \bar R)%E. - congr (_ * _)%E. - apply: eq_integral => x x01. - rewrite /onem -EFinM mulrBl mul1r EFinB. - congr (_ - _)%E. - rewrite /ubeta_nat_pdf /ubeta_nat_pdf'/=. - rewrite mulrCA -mulrA -exprD mulrA -exprD. - congr (_ ^+ _ * _ ^+ _)%:E. - rewrite addnC -!subn1 subDnCA//. - rewrite -!subn1 subDnCA//. -rewrite integralB_EFin//=; last 2 first. - admit. - admit. -rewrite -!beta_nat_normE -EFinM mulrBr /onem mulVf; last first. - rewrite /B mulf_eq0 negb_or. - apply/andP; split. - rewrite mulf_eq0 negb_or. - rewrite gt_eqF ?ltr0n ?fact_gt0//=. - rewrite gt_eqF ?ltr0n ?fact_gt0//=. - rewrite invr_eq0 gt_eqF ?ltr0n ?fact_gt0//=. -congr (_ - _)%:E. -by rewrite mulrC. -Admitted. - -End integral_beta. - Section mscore. Context d (T : measurableType d) (R : realType). Variable f : T -> R. -Definition mscore t : {measure set _ -> \bar R} := - let p := NngNum (normr_ge0 (f t)) in - [the measure _ _ of mscale p [the measure _ _ of dirac tt]]. +Definition mscore t : {measure set unit -> \bar R} := + let p := NngNum (normr_ge0 (f t)) in mscale p \d_tt. Lemma mscoreE t U : mscore t U = if U == set0 then 0 else `| (f t)%:E |. Proof. @@ -1174,7 +364,7 @@ Section score. Context d (T : measurableType d) (R : realType). Variable f : T -> R. -Definition k (mf : measurable_fun setT f) i t U := +Definition k (mf : measurable_fun [set: T] f) i t U := if i%:R%:E <= mscore f t U < i.+1%:R%:E then mscore f t U else @@ -1317,7 +507,7 @@ Section kiteT. Variable k : R.-ker X ~> Y. Definition kiteT : X * bool -> {measure set Y -> \bar R} := - fun xb => if xb.2 then k xb.1 else [the measure _ _ of mzero]. + fun xb => if xb.2 then k xb.1 else mzero. Let measurable_fun_kiteT U : measurable U -> measurable_fun setT (kiteT ^~ U). Proof. @@ -1373,7 +563,7 @@ Section kiteF. Variable k : R.-ker X ~> Y. Definition kiteF : X * bool -> {measure set Y -> \bar R} := - fun xb => if ~~ xb.2 then k xb.1 else [the measure _ _ of mzero]. + fun xb => if ~~ xb.2 then k xb.1 else mzero. Let measurable_fun_kiteF U : measurable U -> measurable_fun setT (kiteF ^~ U). Proof. @@ -1437,7 +627,7 @@ Variables (f : T -> bool) (u1 u2 : R.-sfker T ~> T'). Definition mite (mf : measurable_fun setT f) : T -> set T' -> \bar R := fun t => if f t then u1 t else u2 t. -Variables mf : measurable_fun setT f. +Hypothesis mf : measurable_fun [set: T] f. Let mite0 t : mite mf t set0 = 0. Proof. by rewrite /mite; case: ifPn. Qed. @@ -1459,23 +649,20 @@ Import ITE. Definition kite : R.-sfker T ~> T' := kdirac mf \; kadd (kiteT u1) (kiteF u2). *) -Definition kite := - [the R.-sfker _ ~> _ of kdirac mf] \; - [the R.-sfker _ ~> _ of kadd - [the R.-sfker _ ~> T' of kiteT u1] - [the R.-sfker _ ~> T' of kiteF u2] ]. +Definition kite : R.-sfker T ~> T' := + kdirac mf \; kadd (kiteT u1) (kiteF u2). End ite. Section insn2. Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Definition ret (f : X -> Y) (mf : measurable_fun setT f) - : R.-pker X ~> Y := [the R.-pker _ ~> _ of kdirac mf]. +Definition ret (f : X -> Y) (mf : measurable_fun [set: X] f) + : R.-pker X ~> Y := kdirac mf. -Definition sample (P : X -> pprobability Y R) (mP : measurable_fun setT P) +Definition sample (P : X -> pprobability Y R) (mP : measurable_fun [set: X] P) : R.-pker X ~> Y := - [the R.-pker _ ~> _ of kprobability mP]. + kprobability mP. Definition sample_cst (P : pprobability Y R) : R.-pker X ~> Y := sample (measurable_cst P). @@ -1575,7 +762,7 @@ Context d d' d3 (X : measurableType d) (Y : measurableType d') (Z : measurableType d3) (R : realType). Lemma letin_kret (k : R.-sfker X ~> Y) - (f : X * Y -> Z) (mf : measurable_fun setT f) x U : + (f : X * Y -> Z) (mf : measurable_fun [set: X * Y] f) x U : measurable U -> letin k (ret mf) x U = k x (curry f x @^-1` U). Proof. @@ -1608,9 +795,9 @@ End insn1. Section hard_constraint. Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Definition fail := - letin (score (@measurable_cst _ _ X _ setT (0%R : R))) - (ret (@measurable_cst _ _ _ Y setT point)). +Definition fail : R.-sfker X ~> Y := + letin (score (measurable_cst (0%R : R))) + (ret (measurable_cst point)). Lemma failE x U : fail x U = 0. Proof. by rewrite /fail letinE ge0_integral_mscale//= normr0 mul0e. Qed. @@ -1638,15 +825,13 @@ Arguments kb {d T}. Arguments kn {d T}. Section iter_mprod. -Import Notations. +Local Open Scope type_scope. -Fixpoint iter_mprod (l : list {d & measurableType d}) - : {d & measurableType d} := +Fixpoint iter_mprod (l : seq {d & measurableType d}) : {d & measurableType d} := match l with - | [::] => existT measurableType _ munit + | [::] => existT measurableType _ unit | h :: t => let t' := iter_mprod t in - existT _ _ [the measurableType (projT1 h, projT1 t').-prod of - (projT2 h * projT2 t')%type] + existT _ _ [the measurableType _ of projT2 h * projT2 t'] end. End iter_mprod. @@ -1655,15 +840,13 @@ Section acc. Import Notations. Context {R : realType}. -Fixpoint acc (l : seq {d & measurableType d}) n : - projT2 (iter_mprod l) -> projT2 (nth (existT _ _ munit) l n) := - match l return - projT2 (iter_mprod l) -> projT2 (nth (existT _ _ munit) l n) - with - | [::] => match n with | O => id | m.+1 => id end - | _ :: _ => match n with +Fixpoint acc (l : seq {d & measurableType d}) k : + projT2 (iter_mprod l) -> projT2 (nth (existT _ _ munit) l k) := + match l with + | [::] => match k with O => id | _ => id end + | _ :: _ => match k with | O => fst - | m.+1 => fun H => acc m H.2 + | m.+1 => fun x => acc m x.2 end end. @@ -1680,31 +863,28 @@ Section rpair_pairA. Context d0 d1 d2 (T0 : measurableType d0) (T1 : measurableType d1) (T2 : measurableType d2). -Definition rpair d (T : measurableType d) t : - ([the measurableType _ of T0] -> [the measurableType _ of T0 * T])%type := +Definition rpair d (T : measurableType d) t : T0 -> T0 * T := fun x => (x, t). Lemma mrpair d (T : measurableType d) t : measurable_fun setT (@rpair _ T t). -Proof. exact: measurable_fun_prod. Qed. +Proof. exact: measurable_fun_pair. Qed. -Definition pairA : ([the measurableType _ of T0 * T1 * T2] -> - [the measurableType _ of T0 * (T1 * T2)])%type := +Definition pairA : T0 * T1 * T2 -> T0 * (T1 * T2) := fun x => (x.1.1, (x.1.2, x.2)). -Definition mpairA : measurable_fun setT pairA. +Definition mpairA : measurable_fun [set: (T0 * T1) * T2] pairA. Proof. -apply: measurable_fun_prod => /=; first exact: measurableT_comp. -by apply: measurable_fun_prod => //=; exact: measurableT_comp. +apply: measurable_fun_pair => /=; first exact: measurableT_comp. +by apply: measurable_fun_pair => //=; exact: measurableT_comp. Qed. -Definition pairAi : ([the measurableType _ of T0 * (T1 * T2)] -> - [the measurableType _ of T0 * T1 * T2])%type := +Definition pairAi : T0 * (T1 * T2) -> T0 * T1 * T2 := fun x => (x.1, x.2.1, x.2.2). Definition mpairAi : measurable_fun setT pairAi. Proof. -apply: measurable_fun_prod => //=; last exact: measurableT_comp. -apply: measurable_fun_prod => //=; exact: measurableT_comp. +apply: measurable_fun_pair => //=; last exact: measurableT_comp. +by apply: measurable_fun_pair => //=; exact: measurableT_comp. Qed. End rpair_pairA. @@ -1723,39 +903,33 @@ Import Notations. Context d0 d1 d2 d3 (T0 : measurableType d0) (T1 : measurableType d1) (T2 : measurableType d2) (T3 : measurableType d3) (R : realType). -Definition pairAr d (T : measurableType d) t : - ([the measurableType _ of T0 * T1] -> - [the measurableType _ of T0 * (T1 * T)])%type := +Definition pairAr d (T : measurableType d) t : T0 * T1 -> T0 * (T1 * T) := pairA \o rpair T t. Arguments pairAr {d} T. Lemma mpairAr d (T : measurableType d) t : measurable_fun setT (pairAr T t). Proof. exact: measurableT_comp. Qed. -Definition pairAAr : ([the measurableType _ of T0 * T1 * T2] -> - [the measurableType _ of T0 * (T1 * (T2 * munit))])%type := - pairA \o pairA \o rpair munit tt. +Definition pairAAr : T0 * T1 * T2 -> T0 * (T1 * (T2 * unit)) := + pairA \o pairA \o rpair unit tt. Lemma mpairAAr : measurable_fun setT pairAAr. Proof. by do 2 apply: measurableT_comp => //. Qed. -Definition pairAAAr : ([the measurableType _ of T0 * T1 * T2 * T3] -> - [the measurableType _ of T0 * (T1 * (T2 * (T3 * munit)))])%type := - pairA \o pairA \o pairA \o rpair munit tt. +Definition pairAAAr : T0 * T1 * T2 * T3 -> T0 * (T1 * (T2 * (T3 * unit))) := + pairA \o pairA \o pairA \o rpair unit tt. Lemma mpairAAAr : measurable_fun setT pairAAAr. Proof. by do 3 apply: measurableT_comp => //. Qed. -Definition pairAArAi : ([the measurableType _ of T0 * (T1 * T2)] -> - [the measurableType _ of T0 * (T1 * (T2 * munit))])%type := +Definition pairAArAi : T0 * (T1 * T2) -> T0 * (T1 * (T2 * unit)) := pairAAr \o pairAi. Lemma mpairAArAi : measurable_fun setT pairAArAi. Proof. by apply: measurableT_comp => //=; exact: mpairAAr. Qed. -Definition pairAAArAAi : ([the measurableType _ of T3 * (T0 * (T1 * T2))] -> - [the measurableType _ of T3 * (T0 * (T1 * (T2 * munit)))])%type := - pairA \o pairA \o pairA \o rpair munit tt \o pairAi \o pairAi. +Definition pairAAArAAi : T3 * (T0 * (T1 * T2)) -> T3 * (T0 * (T1 * (T2 * unit))) := + pairA \o pairA \o pairA \o rpair unit tt \o pairAi \o pairAi. Lemma mpairAAARAAAi : measurable_fun setT pairAAArAAi. Proof. by do 5 apply: measurableT_comp => //=. Qed. @@ -1772,92 +946,91 @@ Import Notations. Context d0 d1 d2 d3 (T0 : measurableType d0) (T1 : measurableType d1) (T2 : measurableType d2) (T3 : measurableType d3) (R : realType). -Definition Of2 := [:: existT _ _ T0; existT _ _ T1]. +Definition T01 : seq {d & measurableType d} := [:: existT _ _ T0; existT _ _ T1]. -Definition acc0of2 : [the measurableType _ of (T0 * T1)%type] -> T0 := - @acc Of2 0 \o pairAr munit tt. +Definition acc0of2 : T0 * T1 -> T0 := + acc T01 0 \o pairAr unit tt. Lemma macc0of2 : measurable_fun setT acc0of2. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of2 0)|exact: mpairAr]. +by apply: measurableT_comp; [exact: (measurable_acc T01 0)|exact: mpairAr]. Qed. -Definition acc1of2 : [the measurableType _ of (T0 * T1)%type] -> T1 := - acc Of2 1 \o pairAr munit tt. +Definition acc1of2 : T0 * T1 -> T1 := + acc T01 1 \o pairAr unit tt. Lemma macc1of2 : measurable_fun setT acc1of2. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of2 1)|exact: mpairAr]. +by apply: measurableT_comp; [exact: (measurable_acc T01 1)|exact: mpairAr]. Qed. -Definition Of3 := [:: existT _ _ T0; existT _ _ T1; existT _ d2 T2]. +Definition T02 := [:: existT _ _ T0; existT _ _ T1; existT _ _ T2]. -Definition acc1of3 : [the measurableType _ of (T0 * T1 * T2)%type] -> T1 := - acc Of3 1 \o pairAAr. +Definition acc1of3 : T0 * T1 * T2 -> T1 := + acc T02 1 \o pairAAr. Lemma macc1of3 : measurable_fun setT acc1of3. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of3 1)|exact: mpairAAr]. +by apply: measurableT_comp; [exact: (measurable_acc T02 1)|exact: mpairAAr]. Qed. -Definition acc2of3 : [the measurableType _ of (T0 * T1 * T2)%type] -> T2 := - acc Of3 2 \o pairAAr. +Definition acc2of3 : T0 * T1 * T2 -> T2 := + acc T02 2 \o pairAAr. Lemma macc2of3 : measurable_fun setT acc2of3. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of3 2)|exact: mpairAAr]. +by apply: measurableT_comp; [exact: (measurable_acc T02 2)|exact: mpairAAr]. Qed. -Definition acc0of3' : [the measurableType _ of (T0 * (T1 * T2))%type] -> T0 := - acc Of3 0 \o pairAArAi. +Definition acc0of3' : T0 * (T1 * T2) -> T0 := + acc T02 0 \o pairAArAi. Lemma macc0of3' : measurable_fun setT acc0of3'. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of3 0)|exact: mpairAArAi]. +by apply: measurableT_comp; [exact: (measurable_acc T02 0)|exact: mpairAArAi]. Qed. -Definition acc1of3' : [the measurableType _ of (T0 * (T1 * T2))%type] -> T1 := - acc Of3 1 \o pairAArAi. +Definition acc1of3' : T0 * (T1 * T2) -> T1 := + acc T02 1 \o pairAArAi. Lemma macc1of3' : measurable_fun setT acc1of3'. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of3 1)|exact: mpairAArAi]. +by apply: measurableT_comp; [exact: (measurable_acc T02 1)|exact: mpairAArAi]. Qed. -Definition acc2of3' : [the measurableType _ of (T0 * (T1 * T2))%type] -> T2 := - acc Of3 2 \o pairAArAi. +Definition acc2of3' : T0 * (T1 * T2) -> T2 := + acc T02 2 \o pairAArAi. Lemma macc2of3' : measurable_fun setT acc2of3'. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of3 2)|exact: mpairAArAi]. +by apply: measurableT_comp; [exact: (measurable_acc T02 2)|exact: mpairAArAi]. Qed. -Definition Of4 := +Definition T03 := [:: existT _ _ T0; existT _ _ T1; existT _ d2 T2; existT _ d3 T3]. -Definition acc1of4 : [the measurableType _ of (T0 * T1 * T2 * T3)%type] -> T1 := - acc Of4 1 \o pairAAAr. +Definition acc1of4 : T0 * T1 * T2 * T3 -> T1 := + acc T03 1 \o pairAAAr. Lemma macc1of4 : measurable_fun setT acc1of4. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of4 1)|exact: mpairAAAr]. +by apply: measurableT_comp; [exact: (measurable_acc T03 1)|exact: mpairAAAr]. Qed. -Definition acc2of4' : - [the measurableType _ of (T0 * (T1 * (T2 * T3)))%type] -> T2 := - acc Of4 2 \o pairAAArAAi. +Definition acc2of4' : T0 * (T1 * (T2 * T3)) -> T2 := + acc T03 2 \o pairAAArAAi. Lemma macc2of4' : measurable_fun setT acc2of4'. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of4 2)|exact: mpairAAARAAAi]. +by apply: measurableT_comp; [exact: (measurable_acc T03 2)|exact: mpairAAARAAAi]. Qed. -Definition acc3of4 : [the measurableType _ of (T0 * T1 * T2 * T3)%type] -> T3 := - acc Of4 3 \o pairAAAr. +Definition acc3of4 : T0 * T1 * T2 * T3 -> T3 := + acc T03 3 \o pairAAAr. Lemma macc3of4 : measurable_fun setT acc3of4. Proof. -by apply: measurableT_comp; [exact: (measurable_acc Of4 3)|exact: mpairAAAr]. +by apply: measurableT_comp; [exact: (measurable_acc T03 3)|exact: mpairAAAr]. Qed. End accessor_functions. @@ -1872,51 +1045,486 @@ Arguments macc1of4 {d0 d1 d2 d3 _ _ _ _}. Arguments macc2of4' {d0 d1 d2 d3 _ _ _ _}. Arguments macc3of4 {d0 d1 d2 d3 _ _ _ _}. -Section insn1_lemmas. -Import Notations. -Context d (T : measurableType d) (R : realType). +(* sample programs *) +Section poisson. +Variable R : realType. +Local Open Scope ring_scope. -Let kcomp_scoreE d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) - (g : R.-sfker [the measurableType _ of (T1 * unit)%type] ~> T2) - f (mf : measurable_fun setT f) r U : - (score mf \; g) r U = `|f r|%:E * g (r, tt) U. +(* density function for Poisson *) +Definition poisson k r : R := + if r > 0 then r ^+ k / k`!%:R^-1 * expR (- r) else 1%:R. + +Lemma poisson_ge0 k r : 0 <= poisson k r. Proof. -rewrite /= /kcomp /kscore /= ge0_integral_mscale//=. -by rewrite integral_dirac// diracT mul1e. +rewrite /poisson; case: ifPn => r0//. +by rewrite mulr_ge0 ?expR_ge0// mulr_ge0// exprn_ge0 ?ltW. Qed. -Lemma scoreE d' (T' : measurableType d') (x : T * T') (U : set T') (f : R -> R) - (r : R) (r0 : (0 <= r)%R) - (f0 : (forall r, 0 <= r -> 0 <= f r)%R) (mf : measurable_fun setT f) : - score (measurableT_comp mf (@macc1of2 _ _ _ _)) - (x, r) (curry (snd \o fst) x @^-1` U) = - (f r)%:E * \d_x.2 U. +Lemma poisson_gt0 k r : 0 < r -> 0 < poisson k.+1 r. Proof. -by rewrite /score/= /mscale/= ger0_norm//= f0. +move=> r0; rewrite /poisson r0 mulr_gt0 ?expR_gt0//. +by rewrite divr_gt0// ?exprn_gt0// invr_gt0 ltr0n fact_gt0. Qed. -Lemma score_score (f : R -> R) (g : R * unit -> R) - (mf : measurable_fun setT f) - (mg : measurable_fun setT g) : - letin (score mf) (score mg) = - score (measurable_funM mf (measurableT_comp mg (measurable_pair2 tt))). +Lemma measurable_poisson k : measurable_fun setT (poisson k). Proof. -apply/eq_sfkernel => x U. +rewrite /poisson; apply: measurable_fun_if => //. + exact: measurable_fun_ltr. +by apply: measurable_funM => /=; + [exact: measurable_funM|exact: measurableT_comp]. +Qed. + +Definition poisson3 := poisson 4 3%:R. (* 0.168 *) +Definition poisson10 := poisson 4 10%:R. (* 0.019 *) + +End poisson. + +Section exponential. +Variable R : realType. +Local Open Scope ring_scope. + +(* density function for exponential *) +Definition exp_density x r : R := r * expR (- r * x). + +Lemma exp_density_gt0 x r : 0 < r -> 0 < exp_density x r. +Proof. by move=> r0; rewrite /exp_density mulr_gt0// expR_gt0. Qed. + +Lemma exp_density_ge0 x r : 0 <= r -> 0 <= exp_density x r. +Proof. by move=> r0; rewrite /exp_density mulr_ge0// expR_ge0. Qed. + +Lemma mexp_density x : measurable_fun setT (exp_density x). +Proof. +apply: measurable_funM => //=; apply: measurableT_comp => //. +exact: measurable_funM. +Qed. + +End exponential. + +Module CASE_NAT. +Section case_nat. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). + +Section case_nat_ker. +Variable k : R.-ker X ~> Y. + +Definition case_nat_ m (xn : X * nat) : {measure set Y -> \bar R} := + if xn.2 == m then k xn.1 else mzero. + +Let measurable_fun_case_nat_ m U : measurable U -> + measurable_fun setT (case_nat_ m ^~ U). +Proof. +move=> mU; rewrite /case_nat_ (_ : (fun _ => _) = + (fun x => if x.2 == m then k x.1 U else mzero U)) /=; last first. + by apply/funext => -[t b]/=; case: ifPn. +apply: (@measurable_fun_if_pair_nat _ _ _ _ (k ^~ U) (fun=> mzero U)) => //. +exact/measurable_kernel. +Qed. + +#[export] +HB.instance Definition _ m := isKernel.Build _ _ _ _ _ + (case_nat_ m) (measurable_fun_case_nat_ m). +End case_nat_ker. + +Section sfcase_nat. +Variable k : R.-sfker X ~> Y. + +Let sfcase_nat_ m : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> case_nat_ k m x U = mseries (k_ ^~ x) 0 U. +Proof. +have [k_ hk /=] := sfinite_kernel k. +exists (fun n => case_nat_ (k_ n) m) => /=. + move=> n; have /measure_fam_uubP[r k_r] := measure_uub (k_ n). + exists r%:num => /= -[x [|n']]; rewrite /case_nat_//= /mzero//. + by case: ifPn => //= ?; rewrite /mzero. + by case: ifPn => // ?; rewrite /= /mzero. +move=> [x b] U mU; rewrite /case_nat_; case: ifPn => hb; first by rewrite hk. +by rewrite /mseries eseries0. +Qed. + +#[export] +HB.instance Definition _ m := @isSFiniteKernel_subdef.Build _ _ _ _ _ + (case_nat_ k m) (sfcase_nat_ m). +End sfcase_nat. + +Section fkcase_nat. +Variable k : R.-fker X ~> Y. + +Let case_nat_uub n : measure_fam_uub (case_nat_ k n). +Proof. +have /measure_fam_uubP[M hM] := measure_uub k. +exists M%:num => /= -[]; rewrite /case_nat_ => t [|n']/=. + by case: ifPn => //= ?; rewrite /mzero. +by case: ifPn => //= ?; rewrite /mzero. +Qed. + +#[export] +HB.instance Definition _ n := Kernel_isFinite.Build _ _ _ _ _ + (case_nat_ k n) (case_nat_uub n). +End fkcase_nat. + +End case_nat. +End CASE_NAT. + +Import CASE_NAT. + +Section case_nat. +Context d d' (T : measurableType d) (T' : measurableType d') (R : realType). + +Import CASE_NAT. + +Definition case_nat (t : R.-sfker T ~> nat) (u_ : (R.-sfker T ~> T')^nat) + : R.-sfker T ~> T' := + t \; kseries (fun n => case_nat_ (u_ n) n). + +End case_nat. + +Definition measure_sum_display : + (measure_display * measure_display) -> measure_display. +Proof. exact. Qed. + +Definition image_classes d1 d2 + (T1 : measurableType d1) (T2 : measurableType d2) (T : Type) + (f1 : T1 -> T) (f2 : T2 -> T) := + <>. + +Section sum_salgebra_instance. +Context d1 d2 (T1 : measurableType d1) (T2 : measurableType d2). +Let f1 : T1 -> T1 + T2 := @inl T1 T2. +Let f2 : T2 -> T1 + T2 := @inr T1 T2. + +Lemma sum_salgebra_set0 : image_classes f1 f2 (set0 : set (T1 + T2)). +Proof. exact: sigma_algebra0. Qed. + +Lemma sum_salgebra_setC A : image_classes f1 f2 A -> + image_classes f1 f2 (~` A). +Proof. exact: sigma_algebraC. Qed. + +Lemma sum_salgebra_bigcup (F : _^nat) : (forall i, image_classes f1 f2 (F i)) -> + image_classes f1 f2 (\bigcup_i (F i)). +Proof. exact: sigma_algebra_bigcup. Qed. + +HB.instance Definition sum_salgebra_mixin := + @isMeasurable.Build (measure_sum_display (d1, d2)) + (T1 + T2)%type (image_classes f1 f2) + sum_salgebra_set0 sum_salgebra_setC sum_salgebra_bigcup. + +End sum_salgebra_instance. +Reserved Notation "p .-sum" (at level 1, format "p .-sum"). +Reserved Notation "p .-sum.-measurable" + (at level 2, format "p .-sum.-measurable"). +Notation "p .-sum" := (measure_sum_display p) : measure_display_scope. +Notation "p .-sum.-measurable" := + ((p.-sum).-measurable : set (set (_ + _))) : + classical_set_scope. + +Module CASE_SUM. + +Section case_sum'. + +Section kcase_sum'. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). +Let A : measurableType _ := unit. +Let B : measurableType _ := bool. +Variables (k1 : A -> R.-sfker X ~> Y) (k2 : B -> R.-sfker X ~> Y). + +Definition case_sum' : X * (A + B) -> {measure set Y -> \bar R} := + fun xab => match xab with + | (x, inl a) => k1 a x + | (x, inr b) => k2 b x + end. + +Let measurable_fun_case_sum' U : measurable U -> + measurable_fun setT (case_sum' ^~ U). +Proof. +rewrite /= => mU. +apply: (measurability _ (ErealGenInftyO.measurableE R)) => //. +move=> /= _ [_ [x ->] <-]; apply: measurableI => //. +rewrite /case_sum'/= (_ : _ @^-1` _ = + ([set x1 | k1 tt x1 U < x%:E] `*` inl @` [set tt]) `|` + ([set x1 | k2 false x1 U < x%:E] `*` inr @` [set false]) `|` + ([set x1 | k2 true x1 U < x%:E] `*` inr @` [set true])); last first. + apply/seteqP; split. + - move=> z /=; rewrite in_itv/=; move: z => [z [[]|[|]]]//= ?. + + by do 2 left; split => //; exists tt. + + by right; split => //; exists true. + + by left; right; split => //; exists false. + - move=> z /=; rewrite in_itv/=; move: z => [z [[]|[|]]]//=. + - move=> [[[]//|]|]. + + by move=> [_ []]. + + by move=> [_ []]. + - move=> [[|]|[]//]. + + by move=> [_ []]. + + by move=> [_ [] [|]]. + - move=> [[|[]//]|]. + + by move=> [_ []]. + + by move=> [_ [] [|]]. +pose h1 := [set xub : X * (unit + bool) | k1 tt xub.1 U < x%:E]. +have mh1 : measurable h1. + rewrite -[X in measurable X]setTI; apply: emeasurable_fun_infty_o => //=. + have H : measurable_fun [set: X] (fun x => k1 tt x U) by exact/measurable_kernel. + move=> _ /= C mC; rewrite setTI. + have := H measurableT _ mC; rewrite setTI => {}H. + rewrite [X in measurable X](_ : _ = ((fun x => k1 tt x U) @^-1` C) `*` setT)//. + exact: measurableX. + by apply/seteqP; split => [z//=| z/= []]. +set h2 := [set xub : X * (unit + bool)| k2 false xub.1 U < x%:E]. +have mh2 : measurable h2. + rewrite -[X in measurable X]setTI. + apply: emeasurable_fun_infty_o => //=. + have H : measurable_fun [set: X] (fun x => k2 false x U) by exact/measurable_kernel. + move=> _ /= C mC; rewrite setTI. + have := H measurableT _ mC; rewrite setTI => {}H. + rewrite [X in measurable X](_ : _ = ((fun x => k2 false x U) @^-1` C) `*` setT)//. + exact: measurableX. + by apply/seteqP; split => [z //=|z/= []]. +set h3 := [set xub : X * (unit + bool)| k2 true xub.1 U < x%:E]. +have mh3 : measurable h3. + rewrite -[X in measurable X]setTI. + apply: emeasurable_fun_infty_o => //=. + have H : measurable_fun [set: X] (fun x => k2 (true) x U) by exact/measurable_kernel. + move=> _ /= C mC; rewrite setTI. + have := H measurableT _ mC; rewrite setTI => {}H. + rewrite [X in measurable X](_ : _ = ((fun x => k2 (true) x U) @^-1` C) `*` setT)//. + exact: measurableX. + by apply/seteqP; split=> [z//=|z/= []]. +apply: measurableU. +- apply: measurableU. + + apply: measurableX => //. + rewrite [X in measurable X](_ : _ = ysection h1 (inl tt))//. + * by apply: measurable_ysection. + * by apply/seteqP; split => z /=; rewrite /ysection /= inE. + + apply: measurableX => //. + rewrite [X in measurable X](_ : _ = ysection h2 (inr false))//. + * by apply: measurable_ysection. + * by apply/seteqP; split => z /=; rewrite /ysection /= inE. +- apply: measurableX => //. + rewrite [X in measurable X](_ : _ = ysection h3 (inr true))//. + + by apply: measurable_ysection. + + by apply/seteqP; split => z /=; rewrite /ysection /= inE. +Qed. + +#[export] +HB.instance Definition _ := isKernel.Build _ _ _ _ _ + case_sum' measurable_fun_case_sum'. +End kcase_sum'. + +Section sfkcase_sum'. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). +Let A : measurableType _ := unit. +Let B : measurableType _ := bool. +Variables (k1 : A -> R.-sfker X ~> Y) (k2 : B-> R.-sfker X ~> Y). + +Let sfinite_case_sum' : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> case_sum' k1 k2 x U = mseries (k_ ^~ x) 0 U. +Proof. +rewrite /=. +set f1 : A -> (R.-fker _ ~> _)^nat := + fun ab : A => sval (cid (sfinite_kernel (k1 ab))). +set Hf1 := fun ab : A => svalP (cid (sfinite_kernel (k1 ab))). +rewrite /= in Hf1. +set f2 : B -> (R.-fker _ ~> _)^nat := + fun ab : B => sval (cid (sfinite_kernel (k2 ab))). +set Hf2 := fun ab : B => svalP (cid (sfinite_kernel (k2 ab))). +rewrite /= in Hf2. +exists (fun n => case_sum' (f1 ^~ n) (f2 ^~ n)). + move=> n /=. + have [rtt Hrtt] := measure_uub (f1 tt n). + have [rfalse Hrfalse] := measure_uub (f2 false n). + have [rtrue Hrtrue] := measure_uub (f2 true n). + exists (maxr rtt (maxr rfalse rtrue)) => //= -[x [[]|[|]]] /=. + by rewrite 2!EFin_max lt_max Hrtt. + by rewrite 2!EFin_max 2!lt_max Hrtrue 2!orbT. + by rewrite 2!EFin_max 2!lt_max Hrfalse orbT. +move=> [x [[]|[|]]] U mU/=-. +by rewrite (Hf1 tt x _ mU). +by rewrite (Hf2 true x _ mU). +by rewrite (Hf2 false x _ mU). +Qed. + +#[export] +HB.instance Definition _ := @isSFiniteKernel_subdef.Build _ _ _ _ _ + (case_sum' k1 k2) (sfinite_case_sum'). +End sfkcase_sum'. + +End case_sum'. + +Section case_sum. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). +Let A : measurableType _ := unit. +Let B : measurableType _ := bool. + +(* case analysis on the datatype unit + bool *) +Definition case_sum (f : R.-sfker X ~> (A + B)%type) + (k1 : A -> R.-sfker X ~> Y) (k2 : B -> R.-sfker X ~> Y) : R.-sfker X ~> Y := + f \; case_sum' k1 k2. + +End case_sum. + +End CASE_SUM. + +(* counting measure as a kernel *) +Section kcounting. +Context d (G : measurableType d) (R : realType). + +Definition kcounting : G -> {measure set nat -> \bar R} := fun=> counting. + +Let mkcounting U : measurable U -> measurable_fun setT (kcounting ^~ U). +Proof. by []. Qed. + +HB.instance Definition _ := isKernel.Build _ _ _ _ _ kcounting mkcounting. + +Let sfkcounting : exists2 k_ : (R.-ker _ ~> _)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> kcounting x U = mseries (k_ ^~ x) 0 U. +Proof. +exists (fun n => [the R.-fker _ ~> _ of + @kdirac _ _ G nat R _ (@measurable_cst _ _ _ _ setT n)]). + by move=> n /=; exact: measure_uub. +by move=> g U mU; rewrite /kcounting/= counting_dirac. +Qed. + +HB.instance Definition _ := + isSFiniteKernel_subdef.Build _ _ _ _ R kcounting sfkcounting. + +End kcounting. + +(* formalization of the iterate construct of Staton ESOP 2017, Sect. 4.2 *) +Section iterate. +Context d {G : measurableType d} {R : realType}. +Let A : measurableType _ := unit. +Let B : measurableType _ := bool. + +Import CASE_SUM. + +(* formalization of iterate^n + Gamma |-p iterate^n t from x = u : B *) +Variables (t : R.-sfker (G * A) ~> (A + B)%type) + (u : G -> A) (mu : measurable_fun setT u). + +Fixpoint iterate_ n : R.-sfker G ~> B := + match n with + | 0%N => case_sum (letin (ret mu) t) + (fun u' => fail) + (fun v => ret (measurable_cst v)) + | m.+1 => case_sum (letin (ret mu) t) + (fun u' => iterate_ m) + (fun v => fail) + end. + +(* formalization of iterate (A = unit, B = bool) + Gamma, x : A |-p t : A + B Gamma |-d u : A +----------------------------------------------- + Gamma |-p iterate t from x = u : B *) +Definition iterate : R.-sfker G ~> B := case_nat (kcounting R) iterate_. + +End iterate. + +(* an s-finite kernel to test that two expressions are different *) +Section lift_neq. +Context {R : realType} d (G : measurableType d). +Variables (f : G -> bool) (g : G -> bool). + +Definition flift_neq : G -> bool := fun x' => f x' != g x'. + +Hypotheses (mf : measurable_fun setT f) (mg : measurable_fun setT g). + +(* see also emeasurable_fun_neq *) +Lemma measurable_fun_flift_neq : measurable_fun setT flift_neq. +Proof. +apply: (@measurable_fun_bool _ _ _ _ true). +rewrite setTI. +rewrite /flift_neq /= (_ : _ @^-1` _ = ([set x | f x] `&` [set x | ~~ g x]) `|` + ([set x | ~~ f x] `&` [set x | g x])). + apply: measurableU; apply: measurableI. + - by rewrite -[X in measurable X]setTI; exact: mf. + - rewrite [X in measurable X](_ : _ = ~` [set x | g x]); last first. + by apply/seteqP; split => x /= /negP. + by apply: measurableC; rewrite -[X in measurable X]setTI; exact: mg. + - rewrite [X in measurable X](_ : _ = ~` [set x | f x]); last first. + by apply/seteqP; split => x /= /negP. + by apply: measurableC; rewrite -[X in measurable X]setTI; exact: mf. + - by rewrite -[X in measurable X]setTI; exact: mg. +by apply/seteqP; split => x /=; move: (f x) (g x) => [|] [|]//=; intuition. +Qed. + +Definition lift_neq : R.-sfker G ~> bool := ret measurable_fun_flift_neq. + +End lift_neq. + +Section von_neumann_trick. +Context d {T : measurableType d} {R : realType}. + +Definition minltt {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} := + @measurable_cst _ _ T1 _ setT (@inl unit T2 tt). + +Definition finrb d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) : + T1 * bool -> T2 + bool := fun t1b => inr t1b.2. + +Lemma minrb {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} : + measurable_fun setT (@finrb _ _ T1 T2). +Proof. exact: measurableT_comp. Qed. + +(* biased coin *) +Variable (D : pprobability bool R). + +Definition trick : R.-sfker (T * unit) ~> (unit + bool)%type := + letin (sample_cst D) + (letin (sample_cst D) + (letin (lift_neq macc1of3 macc2of3) + (ite macc3of4 + (letin (ret macc1of4) (ret minrb)) + (ret minltt)))). + +Definition von_neumann_trick : R.-sfker T ~> bool := iterate trick ktt. + +End von_neumann_trick. + +Section insn1_lemmas. +Import Notations. +Context d (T : measurableType d) (R : realType). + +Let kcomp_scoreE d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) + (g : R.-sfker [the measurableType _ of (T1 * unit)%type] ~> T2) + f (mf : measurable_fun setT f) r U : + (score mf \; g) r U = `|f r|%:E * g (r, tt) U. +Proof. +rewrite /= /kcomp /kscore /= ge0_integral_mscale//=. +by rewrite integral_dirac// diracT mul1e. +Qed. + +Lemma scoreE d' (T' : measurableType d') (x : T * T') (U : set T') (f : R -> R) + (r : R) (r0 : (0 <= r)%R) + (f0 : (forall r, 0 <= r -> 0 <= f r)%R) (mf : measurable_fun setT f) : + score (measurableT_comp mf (@macc1of2 _ _ _ _)) + (x, r) (curry (snd \o fst) x @^-1` U) = + (f r)%:E * \d_x.2 U. +Proof. +by rewrite /score/= /mscale/= ger0_norm//= f0. +Qed. + +Lemma score_score (f : R -> R) (g : R * unit -> R) + (mf : measurable_fun setT f) + (mg : measurable_fun setT g) : + letin (score mf) (score mg) = + score (measurable_funM mf (measurableT_comp mg (pair2_measurable tt))). +Proof. +apply/eq_sfkernel => x U. rewrite {1}/letin; unlock. by rewrite kcomp_scoreE/= /mscale/= diracE normrM muleA EFinM. Qed. (* hard constraints to express score below 1 *) -Lemma score_fail (r : {nonneg R}) (r1 : (r%:num <= 1)%R) : - score (kr r%:num) = - letin (sample_cst (bernoulli r1) : R.-pker T ~> _) +Lemma score_fail (r : R) : (0 <= r <= 1)%R -> + score (kr r) = + letin (sample_cst (bernoulli r) : R.-pker T ~> _) (ite (@macc1of2 _ _ _ _) (ret ktt) fail). Proof. -apply/eq_sfkernel => x U. +move=> /andP[r0 r1]; apply/eq_sfkernel => x U. rewrite letinE/= /sample; unlock. -rewrite ge0_integral_measure_add// ge0_integral_mscale//= ge0_integral_mscale//=. -rewrite integral_dirac//= integral_dirac//= !diracT/= !mul1e. -by rewrite /mscale/= iteE//= iteE//= failE mule0 adde0 ger0_norm. +by rewrite /mscale/= ger0_norm// integral_bernoulli ?r0//= 2!iteE//= failE mule0 adde0. Qed. End insn1_lemmas. @@ -1931,17 +1539,13 @@ Variables (k1 k2 : R.-sfker T ~> Z) Lemma letin_iteT : f t -> letin (ite mf k1 k2) u t U = letin k1 u t U. Proof. -move=> ftT. -rewrite !letinE/=. -apply: eq_measure_integral => V mV _. +move=> ftT; rewrite !letinE/=; apply: eq_measure_integral => V mV _. by rewrite iteE ftT. Qed. Lemma letin_iteF : ~~ f t -> letin (ite mf k1 k2) u t U = letin k2 u t U. Proof. -move=> ftF. -rewrite !letinE/=. -apply: eq_measure_integral => V mV _. +move=> ftF; rewrite !letinE/=; apply: eq_measure_integral => V mV _. by rewrite iteE (negbTE ftF). Qed. @@ -1953,9 +1557,9 @@ Context d d' d1 d2 d3 (X : measurableType d) (Y : measurableType d') (R : realType). Import Notations. Variables (t : R.-sfker X ~> T1) - (u : R.-sfker [the measurableType _ of (X * T1)%type] ~> T2) - (v : R.-sfker [the measurableType _ of (X * T2)%type] ~> Y) - (v' : R.-sfker [the measurableType _ of (X * T1 * T2)%type] ~> Y) + (u : R.-sfker (X * T1) ~> T2) + (v : R.-sfker (X * T2) ~> Y) + (v' : R.-sfker (X * T1 * T2) ~> Y) (vv' : forall y, v =1 fun xz => v' (xz.1, y, xz.2)). Lemma letinA x A : measurable A -> @@ -1970,7 +1574,7 @@ rewrite integral_kcomp; [|by []|]. - apply: eq_integral => y _. apply: eq_integral => z _. by rewrite (vv' y). -exact: (measurableT_comp (measurable_kernel v _ mA)). +- exact: (measurableT_comp (measurable_kernel v _ mA)). Qed. End letinA. @@ -2013,10 +1617,10 @@ HB.instance Definition _ z := @isSFinite.Build _ Y R (U z) (sfinU z). Lemma letinC z A : measurable A -> letin t (letin u' - (ret (measurable_fun_prod macc1of3 macc2of3))) z A = + (ret (measurable_fun_pair macc1of3 macc2of3))) z A = letin u (letin t' - (ret (measurable_fun_prod macc2of3 macc1of3))) z A. + (ret (measurable_fun_pair macc2of3 macc1of3))) z A. Proof. move=> mA. rewrite !letinE. @@ -2038,71 +1642,15 @@ Qed. End letinC. -(* sample programs *) -Section poisson. -Variable R : realType. -Local Open Scope ring_scope. - -(* density function for Poisson *) -Definition poisson k r : R := - if r > 0 then r ^+ k / k`!%:R^-1 * expR (- r) else 1%:R. - -Lemma poisson_ge0 k r : 0 <= poisson k r. -Proof. -rewrite /poisson; case: ifPn => r0//. -by rewrite mulr_ge0 ?expR_ge0// mulr_ge0// exprn_ge0 ?ltW. -Qed. - -Lemma poisson_gt0 k r : 0 < r -> 0 < poisson k.+1 r. -Proof. -move=> r0; rewrite /poisson r0 mulr_gt0 ?expR_gt0//. -by rewrite divr_gt0// ?exprn_gt0// invr_gt0 ltr0n fact_gt0. -Qed. - -Lemma measurable_poisson k : measurable_fun setT (poisson k). -Proof. -rewrite /poisson; apply: measurable_fun_if => //. - exact: measurable_fun_ltr. -by apply: measurable_funM => /=; - [exact: measurable_funM|exact: measurableT_comp]. -Qed. - -Definition poisson3 := poisson 4 3%:R. (* 0.168 *) -Definition poisson10 := poisson 4 10%:R. (* 0.019 *) - -End poisson. - -Section exponential. -Variable R : realType. -Local Open Scope ring_scope. - -(* density function for exponential *) -Definition exp_density x r : R := r * expR (- r * x). - -Lemma exp_density_gt0 x r : 0 < r -> 0 < exp_density x r. -Proof. by move=> r0; rewrite /exp_density mulr_gt0// expR_gt0. Qed. - -Lemma exp_density_ge0 x r : 0 <= r -> 0 <= exp_density x r. -Proof. by move=> r0; rewrite /exp_density mulr_ge0// expR_ge0. Qed. - -Lemma mexp_density x : measurable_fun setT (exp_density x). -Proof. -apply: measurable_funM => //=; apply: measurableT_comp => //. -exact: measurable_funM. -Qed. - -End exponential. +(* examples *) Lemma letin_sample_bernoulli d d' (T : measurableType d) - (T' : measurableType d') (R : realType)(r : {nonneg R}) (r1 : (r%:num <= 1)%R) + (T' : measurableType d') (R : realType) (r : R) (u : R.-sfker [the measurableType _ of (T * bool)%type] ~> T') x y : - letin (sample_cst (bernoulli r1)) u x y = - r%:num%:E * u (x, true) y + (`1- (r%:num))%:E * u (x, false) y. -Proof. -rewrite letinE/=. -rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. -by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT 2!mul1e. -Qed. + (0 <= r <= 1)%R -> + letin (sample_cst (bernoulli r)) u x y = + r%:E * u (x, true) y + (`1- r)%:E * u (x, false) y. +Proof. by move=> r01; rewrite letinE/= integral_bernoulli. Qed. Section sample_and_return. Import Notations. @@ -2110,18 +1658,18 @@ Context d (T : measurableType d) (R : realType). Definition sample_and_return : R.-sfker T ~> _ := letin - (sample_cst [the probability _ _ of bernoulli p27]) (* T -> B *) + (sample_cst (bernoulli (2 / 7))) (* T -> B *) (ret macc1of2) (* T * B -> B *). Lemma sample_and_returnE t U : sample_and_return t U = (2 / 7%:R)%:E * \d_true U + (5%:R / 7%:R)%:E * \d_false U. Proof. -by rewrite /sample_and_return letin_sample_bernoulli !retE onem27. +rewrite /sample_and_return letin_sample_bernoulli; last lra. +by rewrite !retE onem27. Qed. End sample_and_return. -(* trivial example *) Section sample_and_branch. Import Notations. Context d (T : measurableType d) (R : realType). @@ -2130,16 +1678,16 @@ Context d (T : measurableType d) (R : realType). let r = case x of {(1, _) => return (k3()), (2, _) => return (k10())} in return r *) -Definition sample_and_branch : R.-sfker T ~> R := +Definition sample_and_branch : R.-sfker T ~> _ := letin - (sample_cst [the probability _ _ of bernoulli p27]) (* T -> B *) + (sample_cst (bernoulli (2 / 7))) (* T -> B *) (ite macc1of2 (ret (@k3 _ _ R)) (ret k10)). Lemma sample_and_branchE t U : sample_and_branch t U = - (2 / 7%:R)%:E * \d_(3%:R : R) U + - (5%:R / 7%:R)%:E * \d_(10%:R : R) U. + (2 / 7)%:E * \d_(3%R : R) U + (5 / 7)%:E * \d_(10%R : R) U. Proof. -by rewrite /sample_and_branch letin_sample_bernoulli//= !iteE/= onem27. +rewrite /sample_and_branch letin_sample_bernoulli/=; last lra. +by rewrite !iteE/= onem27. Qed. End sample_and_branch. @@ -2149,43 +1697,24 @@ Context d (T : measurableType d) (R : realType). Import Notations. Definition bernoulli_and : R.-sfker T ~> mbool := - (letin (sample_cst [the probability _ _ of bernoulli p12]) - (letin (sample_cst [the probability _ _ of bernoulli p12]) + (letin (sample_cst (bernoulli (1 / 2))) + (letin (sample_cst (bernoulli (1 / 2))) (ret (measurable_and macc1of3 macc2of3)))). -(* -Definition mand (x y : T * mbool * mbool -> mbool) - (t : T * mbool * mbool) : mbool := x t && y t. - -Lemma measurable_fun_mand (x y : T * mbool * mbool -> mbool) : - measurable_fun setT x -> measurable_fun setT y -> - measurable_fun setT (mand x y). -Proof. -move=> /= mx my; apply: (measurable_fun_bool true). -rewrite [X in measurable X](_ : _ = - (x @^-1` [set true]) `&` (y @^-1` [set true])); last first. - by rewrite /mand; apply/seteqP; split => z/= /andP. -apply: measurableI. -- by rewrite -[X in measurable X]setTI; exact: mx. -- by rewrite -[X in measurable X]setTI; exact: my. -Qed. - -Definition bernoulli_and : R.-sfker T ~> mbool := - (letin (sample_cst [the probability _ _ of bernoulli p12]) - (letin (sample_cst [the probability _ _ of bernoulli p12]) - (ret (measurable_fun_mand macc1of3 macc2of3)))). -*) Lemma bernoulli_andE t U : - bernoulli_and t U = - sample_cst (bernoulli p14) t U. + bernoulli_and t U = sample_cst (bernoulli (1 / 4)) t U. Proof. -rewrite /bernoulli_and 3!letin_sample_bernoulli/= muleDr//= -muleDl//. +rewrite /bernoulli_and. +rewrite letin_sample_bernoulli; last lra. +rewrite (letin_sample_bernoulli (r := 1 / 2)); last lra. +rewrite (letin_sample_bernoulli (r := 1 / 2)); last lra. +rewrite muleDr//= -muleDl//. rewrite !muleA -addeA -muleDl// -!EFinM !onem1S/= -splitr mulr1. have -> : (1 / 2 * (1 / 2) = 1 / 4%:R :> R)%R by rewrite mulf_div mulr1// -natrM. -rewrite /bernoulli/= measure_addE/= /mscale/= -!EFinM; congr( _ + (_ * _)%:E). -have -> : (1 / 2 = 2 / 4%:R :> R)%R. - by apply/eqP; rewrite eqr_div// ?pnatr_eq0// mul1r -natrM. -by rewrite onem1S// -mulrDl. +rewrite [in RHS](_ : 1 / 4 = (1 / 4)%:nng%:num)%R//. +rewrite bernoulliE/=; last lra. +rewrite -!EFinM; congr( _ + (_ * _)%:E). +by rewrite /onem; lra. Qed. End bernoulli_and. @@ -2195,7 +1724,7 @@ Import Notations. Context d (T : measurableType d) (R : realType) (h : R -> R). Hypothesis mh : measurable_fun setT h. Definition kstaton_bus : R.-sfker T ~> mbool := - letin (sample_cst [the probability _ _ of bernoulli p27]) + letin (sample_cst (bernoulli (2 / 7))) (letin (letin (ite macc1of2 (ret k3) (ret k10)) (score (measurableT_comp mh macc2of3))) @@ -2212,43 +1741,43 @@ End staton_bus. Section staton_bus_poisson. Import Notations. Context d (T : measurableType d) (R : realType). -Let poisson4 := @poisson R 4%N. -Let mpoisson4 := @measurable_poisson R 4%N. +Let poisson4 := @poisson_pdf R 4%N. +Let mpoisson4 := @measurable_poisson_pdf R 4%N. Definition kstaton_bus_poisson : R.-sfker R ~> mbool := kstaton_bus _ mpoisson4. Let kstaton_bus_poissonE t U : kstaton_bus_poisson t U = - (2 / 7%:R)%:E * (poisson4 3%:R)%:E * \d_true U + - (5%:R / 7%:R)%:E * (poisson4 10%:R)%:E * \d_false U. + (2 / 7)%:E * (poisson4 3)%:E * \d_true U + + (5 / 7)%:E * (poisson4 10)%:E * \d_false U. Proof. -rewrite /kstaton_bus. -rewrite letin_sample_bernoulli. +rewrite /kstaton_bus_poisson /kstaton_bus. +rewrite letin_sample_bernoulli; last lra. rewrite -!muleA; congr (_ * _ + _ * _). - rewrite letin_kret//. rewrite letin_iteT//. rewrite letin_retk//. - by rewrite scoreE//= => r r0; exact: poisson_ge0. + by rewrite scoreE//= => r r0; exact: poisson_pdf_ge0. - by rewrite onem27. rewrite letin_kret//. rewrite letin_iteF//. rewrite letin_retk//. - by rewrite scoreE//= => r r0; exact: poisson_ge0. + by rewrite scoreE//= => r r0; exact: poisson_pdf_ge0. Qed. (* true -> 2/7 * 0.168 = 2/7 * 3^4 e^-3 / 4! *) (* false -> 5/7 * 0.019 = 5/7 * 10^4 e^-10 / 4! *) Lemma staton_busE P (t : R) U : - let N := ((2 / 7%:R) * poisson4 3%:R + - (5%:R / 7%:R) * poisson4 10%:R)%R in + let N := ((2 / 7) * poisson4 3 + + (5 / 7) * poisson4 10)%R in staton_bus mpoisson4 P t U = - ((2 / 7%:R)%:E * (poisson4 3%:R)%:E * \d_true U + - (5%:R / 7%:R)%:E * (poisson4 10%:R)%:E * \d_false U) * N^-1%:E. + ((2 / 7)%:E * (poisson4 3)%:E * \d_true U + + (5 / 7)%:E * (poisson4 10)%:E * \d_false U) * N^-1%:E. Proof. rewrite /staton_bus normalizeE !kstaton_bus_poissonE !diracT !mule1 ifF //. apply/negbTE; rewrite gt_eqF// lte_fin. -by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// poisson_gt0// ltr0n. +by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// poisson_pdf_gt0// ltr0n. Qed. End staton_bus_poisson. @@ -2260,8 +1789,8 @@ End staton_bus_poisson. Section staton_bus_exponential. Import Notations. Context d (T : measurableType d) (R : realType). -Let exp1560 := @exp_density R (ratr (15%:Q / 60%:Q)). -Let mexp1560 := @mexp_density R (ratr (15%:Q / 60%:Q)). +Let exp1560 := @exponential_pdf R (ratr (15%:Q / 60%:Q)). +Let mexp1560 := @measurable_exponential_pdf R (ratr (15%:Q / 60%:Q)). (* 15/60 = 0.25 *) @@ -2269,560 +1798,303 @@ Definition kstaton_bus_exponential : R.-sfker R ~> mbool := kstaton_bus _ mexp1560. Let kstaton_bus_exponentialE t U : kstaton_bus_exponential t U = - (2 / 7%:R)%:E * (exp1560 3%:R)%:E * \d_true U + - (5%:R / 7%:R)%:E * (exp1560 10%:R)%:E * \d_false U. + (2 / 7)%:E * (exp1560 3)%:E * \d_true U + + (5 / 7)%:E * (exp1560 10)%:E * \d_false U. Proof. rewrite /kstaton_bus. -rewrite letin_sample_bernoulli. +rewrite letin_sample_bernoulli; last lra. rewrite -!muleA; congr (_ * _ + _ * _). - rewrite letin_kret//. rewrite letin_iteT//. rewrite letin_retk//. - rewrite scoreE//= => r r0; exact: exp_density_ge0. + rewrite scoreE//= => r r0; exact: exponential_pdf_ge0. - by rewrite onem27. rewrite letin_kret//. rewrite letin_iteF//. - rewrite letin_retk//. - by rewrite scoreE//= => r r0; exact: exp_density_ge0. -Qed. - -(* true -> 5/7 * 0.019 = 5/7 * 10^4 e^-10 / 4! *) -(* false -> 2/7 * 0.168 = 2/7 * 3^4 e^-3 / 4! *) - -Lemma staton_bus_exponentialE P (t : R) U : - let N := ((2 / 7%:R) * exp1560 3%:R + - (5%:R / 7%:R) * exp1560 10%:R)%R in - staton_bus mexp1560 P t U = - ((2 / 7%:R)%:E * (exp1560 3%:R)%:E * \d_true U + - (5%:R / 7%:R)%:E * (exp1560 10%:R)%:E * \d_false U) * N^-1%:E. -Proof. -rewrite /staton_bus. -rewrite normalizeE /= !kstaton_bus_exponentialE !diracT !mule1 ifF //. -apply/negbTE; rewrite gt_eqF// lte_fin. -by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// exp_density_gt0 ?ltr0n. -Qed. - -End staton_bus_exponential. - -Module CASE_NAT. -Section case_nat. -Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). - -Section case_nat_ker. -Variable k : R.-ker X ~> Y. - -Definition case_nat_ j : X * nat -> {measure set Y -> \bar R} := - fun xn => if xn.2 == j then k xn.1 else [the measure _ _ of mzero]. - -Let measurable_fun_case_nat_ m U : measurable U -> - measurable_fun setT (case_nat_ m ^~ U). -Proof. -move=> mU; rewrite /case_nat_ (_ : (fun _ => _) = - (fun x => if x.2 == m then k x.1 U else mzero U)) /=; last first. - by apply/funext => -[t b]/=; case: ifPn. -apply: (@measurable_fun_if_pair_nat _ _ _ _ (k ^~ U) (fun=> mzero U)) => //. -exact/measurable_kernel. -Qed. - -#[export] -HB.instance Definition _ j := isKernel.Build _ _ _ _ _ - (case_nat_ j) (measurable_fun_case_nat_ j). -End case_nat_ker. - -Section sfcase_nat. -Variable k : R.-sfker X ~> Y. - -Let sfcase_nat_ j : exists2 k_ : (R.-ker _ ~> _)^nat, - forall n, measure_fam_uub (k_ n) & - forall x U, measurable U -> case_nat_ k j x U = mseries (k_ ^~ x) 0 U. -Proof. -have [k_ hk /=] := sfinite_kernel k. -exists (fun n => [the _.-ker _ ~> _ of case_nat_ (k_ n) j]) => /=. - move=> n; have /measure_fam_uubP[r k_r] := measure_uub (k_ n). - exists r%:num => /= -[x [|n']]; rewrite /case_nat_//= /mzero//. - by case: ifPn => //= ?; rewrite /mzero. - by case: ifPn => // ?; rewrite /= /mzero. -move=> [x b] U mU; rewrite /case_nat_; case: ifPn => hb; first by rewrite hk. -by rewrite /mseries eseries0. -Qed. - -#[export] -HB.instance Definition _ j := @isSFiniteKernel_subdef.Build _ _ _ _ _ - (case_nat_ k j) (sfcase_nat_ j). -End sfcase_nat. - -Section fkcase_nat. -Variable k : R.-fker X ~> Y. - -Let case_nat_uub (m : nat) : measure_fam_uub (case_nat_ k m). -Proof. -have /measure_fam_uubP[M hM] := measure_uub k. -exists M%:num => /= -[]; rewrite /case_nat_ => t [|m']/=. - by case: ifPn => //= ?; rewrite /mzero//=. -by case: ifPn => //= ?; rewrite /mzero//=. -Qed. - -#[export] -HB.instance Definition _ j := Kernel_isFinite.Build _ _ _ _ _ - (case_nat_ k j) (case_nat_uub j). -End fkcase_nat. - -End case_nat. -End CASE_NAT. - -Import CASE_NAT. - -Section case_nat. -Context d d' (T : measurableType d) (T' : measurableType d') (R : realType). - -Import CASE_NAT. - -(* case analysis on the nat datatype *) -Definition case_nat (t : R.-sfker T ~> nat) (u_ : (R.-sfker T ~> T')^nat) - : R.-sfker T ~> T' := - [the R.-sfker T ~> nat of t] \; - [the R.-sfker T * nat ~> T' of - kseries (fun n => [the R.-sfker T * nat ~> T' of case_nat_ (u_ n) n])]. - -End case_nat. - -Definition measure_sum_display : - (measure_display * measure_display) -> measure_display. -Proof. exact. Qed. - -Definition image_classes d1 d2 - (T1 : measurableType d1) (T2 : measurableType d2) (T : Type) - (f1 : T1 -> T) (f2 : T2 -> T) := - <>. - -Section sum_salgebra_instance. -Context d1 d2 (T1 : measurableType d1) (T2 : measurableType d2). -Let f1 : T1 -> T1 + T2 := @inl T1 T2. -Let f2 : T2 -> T1 + T2 := @inr T1 T2. - -Lemma sum_salgebra_set0 : image_classes f1 f2 (set0 : set (T1 + T2)). -Proof. exact: sigma_algebra0. Qed. - -Lemma sum_salgebra_setC A : image_classes f1 f2 A -> - image_classes f1 f2 (~` A). -Proof. exact: sigma_algebraC. Qed. - -Lemma sum_salgebra_bigcup (F : _^nat) : (forall i, image_classes f1 f2 (F i)) -> - image_classes f1 f2 (\bigcup_i (F i)). -Proof. exact: sigma_algebra_bigcup. Qed. - -HB.instance Definition sum_salgebra_mixin := - @isMeasurable.Build (measure_sum_display (d1, d2)) - (T1 + T2)%type (image_classes f1 f2) - sum_salgebra_set0 sum_salgebra_setC sum_salgebra_bigcup. - -End sum_salgebra_instance. -Reserved Notation "p .-sum" (at level 1, format "p .-sum"). -Reserved Notation "p .-sum.-measurable" - (at level 2, format "p .-sum.-measurable"). -Notation "p .-sum" := (measure_sum_display p) : measure_display_scope. -Notation "p .-sum.-measurable" := - ((p.-sum).-measurable : set (set (_ + _))) : - classical_set_scope. - -Module CASE_SUM. - -Section case_suml. -Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Let A : measurableType _ := unit. - -Section kcase_suml. -Variable k : R.-ker X ~> Y. - -Definition case_suml (a : A) : X * A -> {measure set Y -> \bar R} := - fun xa => k xa.1. - -Let measurable_fun_case_suml a U : measurable U -> - measurable_fun setT (case_suml a ^~ U). -Proof. -move=> /= mU; rewrite /case_suml. -have h := measurable_kernel k _ mU. -rewrite (_ : (fun x : X * unit => k x.1 U) = (fun x : X => k x U) \o fst) //. -by apply: measurableT_comp => //. + rewrite letin_retk//. + by rewrite scoreE//= => r r0; exact: exponential_pdf_ge0. Qed. -#[export] -HB.instance Definition _ a := isKernel.Build _ _ _ _ _ - (case_suml a) (measurable_fun_case_suml a). -End kcase_suml. - -Section sfkcase_suml. -Variable k : R.-sfker X ~> Y. +(* true -> 5/7 * 0.019 = 5/7 * 10^4 e^-10 / 4! *) +(* false -> 2/7 * 0.168 = 2/7 * 3^4 e^-3 / 4! *) -Let sfinite_case_suml (a : A) : exists2 k_ : (R.-ker _ ~> _)^nat, - forall n, measure_fam_uub (k_ n) & - forall x U, measurable U -> case_suml k a x U = mseries (k_ ^~ x) 0 U. +Lemma staton_bus_exponentialE P (t : R) U : + let N := ((2 / 7) * exp1560 3 + + (5 / 7) * exp1560 10)%R in + staton_bus mexp1560 P t U = + ((2 / 7)%:E * (exp1560 3)%:E * \d_true U + + (5 / 7)%:E * (exp1560 10)%:E * \d_false U) * N^-1%:E. Proof. -have [k_ hk /=] := sfinite_kernel k. -exists (fun n => [the _.-ker _ ~> _ of case_suml (k_ n) a]) => /=. - move=> n; have /measure_fam_uubP[r k_r] := measure_uub (k_ n). - by exists r%:num => /= -[x []]; rewrite /case_suml//= /mzero//. -move=> [x b] U mU; rewrite /case_suml /=. -by rewrite /mseries hk. +rewrite /staton_bus. +rewrite normalizeE /= !kstaton_bus_exponentialE !diracT !mule1 ifF //. +apply/negbTE; rewrite gt_eqF// lte_fin. +by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// exponential_pdf_gt0 ?ltr0n. Qed. -#[export] -HB.instance Definition _ (a : A) := @isSFiniteKernel_subdef.Build _ _ _ _ _ - (case_suml k a) (sfinite_case_suml a). -End sfkcase_suml. - -Section fkcase_suml. -Variable k : R.-fker X ~> Y. - -Let case_suml_uub (a : A) : measure_fam_uub (case_suml k a). -Proof. -have /measure_fam_uubP[M hM] := measure_uub k. -by exists M%:num => /= -[]; rewrite /case_suml. -Qed. +End staton_bus_exponential. -#[export] -HB.instance Definition _ a := Kernel_isFinite.Build _ _ _ _ _ - (case_suml k a) (case_suml_uub a). -End fkcase_suml. +(**md + letin' variants +*) -End case_suml. +Section mswap. +Context d d' d3 (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). +Variable k : R.-ker Y * X ~> Z. -Section case_sumr. -Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Let B : measurableType _ := bool. +Definition mswap xy U := k (swap xy) U. -Section kcase_sumr. -Variable k : R.-ker X ~> Y. +Let mswap0 xy : mswap xy set0 = 0. +Proof. done. Qed. -Definition case_sumr (b : B) : X * B -> {measure set Y -> \bar R} := - fun xa => if xa.2 == b then k xa.1 else [the measure _ _ of mzero]. +Let mswap_ge0 x U : 0 <= mswap x U. +Proof. done. Qed. -Let measurable_fun_case_sumr b U : measurable U -> - measurable_fun setT (case_sumr b ^~ U). -Proof. -move=> /= mU; rewrite /case_suml. -have h := measurable_kernel k _ mU. -rewrite /case_sumr. -rewrite (_ : (fun x : X * bool => case_sumr b x U) = - (fun x : X * bool => (if x.2 == b then k x.1 U else [the {measure set Y -> \bar R} of mzero] U))); last first. - apply/funext => x. - rewrite /case_sumr. - by case: ifPn. -apply: measurable_fun_ifT => //=. - rewrite (_ : (fun t : X * bool => t.2 == b) = (fun t : bool => t == b) \o snd)//. - apply: measurableT_comp => //. -rewrite (_ : (fun t : X * bool => k t.1 U) = (fun t : X => k t U) \o fst)//. -by apply: measurableT_comp => //. -Qed. +Let mswap_sigma_additive x : semi_sigma_additive (mswap x). +Proof. exact: measure_semi_sigma_additive. Qed. -#[export] -HB.instance Definition _ b := isKernel.Build _ _ _ _ _ - (case_sumr b) (measurable_fun_case_sumr b). -End kcase_sumr. +HB.instance Definition _ x := isMeasure.Build _ _ R + (mswap x) (mswap0 x) (mswap_ge0 x) (@mswap_sigma_additive x). -Section sfkcase_sumr. -Variable k : R.-sfker X ~> Y. +Definition mkswap : _ -> {measure set Z -> \bar R} := + fun x => mswap x. -Let sfinite_case_sumr b : exists2 k_ : (R.-ker _ ~> _)^nat, - forall n, measure_fam_uub (k_ n) & - forall x U, measurable U -> case_sumr k b x U = mseries (k_ ^~ x) 0 U. +Let measurable_fun_kswap U : + measurable U -> measurable_fun setT (mkswap ^~ U). Proof. -have [k_ hk /=] := sfinite_kernel k. -exists (fun n => [the _.-ker _ ~> _ of case_sumr (k_ n) b]) => /=. - move=> n; have /measure_fam_uubP[r k_r] := measure_uub (k_ n). - by exists r%:num => /= -[x []]; rewrite /case_sumr//=; case: ifPn => // _; - rewrite /= (le_lt_trans _ (k_r x))// /mzero//. -move=> [x [|]] U mU; rewrite /case_sumr /=; case: b => //=; rewrite ?hk//; -by rewrite /mseries/= eseries0. +move=> mU. +rewrite [X in measurable_fun _ X](_ : _ = k ^~ U \o @swap _ _)//. +apply measurableT_comp => //=; first exact: measurable_kernel. +exact: measurable_swap. Qed. -#[export] -HB.instance Definition _ b := @isSFiniteKernel_subdef.Build _ _ _ _ _ - (case_sumr k b) (sfinite_case_sumr b). -End sfkcase_sumr. +HB.instance Definition _ := isKernel.Build _ _ + (X * Y)%type Z R mkswap measurable_fun_kswap. -Section fkcase_sumr. -Variable k : R.-fker X ~> Y. +End mswap. + +Section mswap_sfinite_kernel. +Variables (d d' d3 : _) (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). +Variable k : R.-sfker Y * X ~> Z. -Let case_sumr_uub b : measure_fam_uub (case_sumr k b). +Let mkswap_sfinite : + exists2 k_ : (R.-ker X * Y ~> Z)^nat, + forall n, measure_fam_uub (k_ n) & + forall x U, measurable U -> mkswap k x U = kseries k_ x U. Proof. -have /measure_fam_uubP[M hM] := measure_uub k. -by exists M%:num => /= -[]; rewrite /case_sumr => x [|] /=; case: b => //=; - rewrite (le_lt_trans _ (hM x))// /mzero. +have [k_ /= kE] := sfinite_kernel k. +exists (fun n => mkswap (k_ n)). + move=> n. + have /measure_fam_uubP[M hM] := measure_uub (k_ n). + by exists M%:num => x/=; exact: hM. +move=> xy U mU. +by rewrite /mswap/= kE. Qed. -#[export] -HB.instance Definition _ b := Kernel_isFinite.Build _ _ _ _ _ - (case_sumr k b) (case_sumr_uub b). -End fkcase_sumr. - -End case_sumr. -End CASE_SUM. - -Section case_sum'. +HB.instance Definition _ := + isSFiniteKernel_subdef.Build _ _ _ Z R (mkswap k) mkswap_sfinite. -Section kcase_sum'. -Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Let A : measurableType _ := unit. -Let B : measurableType _ := bool. -Variables (k : (A + B)%type -> R.-sfker X ~> Y). +End mswap_sfinite_kernel. -Definition case_sum' : X * (A + B)%type -> {measure set Y -> \bar R} := - fun xab => match xab with - | (x, inl a) => CASE_SUM.case_suml (k xab.2) a (x, a) - | (x, inr b) => CASE_SUM.case_sumr (k xab.2) b (x, b) - end. +Section kswap_finite_kernel_finite. +Context d d' d3 (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType) + (k : R.-fker Y * X ~> Z). -Let measurable_fun_case_sum' U : measurable U -> - measurable_fun setT (case_sum' ^~ U). +Let mkswap_finite : measure_fam_uub (mkswap k). Proof. -rewrite /= => mU. -apply: (measurability _ (ErealGenInftyO.measurableE R)) => //. -move=> /= _ [_ [x ->] <-]; apply: measurableI => //. -rewrite /case_sum' /CASE_SUM.case_suml /CASE_SUM.case_sumr /=. -rewrite (_ : - (fun x : X * (unit + bool) => (let (x0, s) := x in - match s with inl _ => k x.2 x0 | inr b => if b == b then k x.2 x0 else mzero - end) U) = - (fun x : X * (unit + bool) => k x.2 x.1 U)); last first. - by apply/funext => -[x1 [a|b]] //; rewrite eqxx. -rewrite (_ : _ @^-1` _ = - ([set x1 | k (inl tt) x1 U < x%:E] `*` inl @` [set tt]) `|` - ([set x1 | k (inr false) x1 U < x%:E] `*` inr @` [set false]) `|` - ([set x1 | k (inr true) x1 U < x%:E] `*` inr @` [set true])); last first. - apply/seteqP; split. - - move=> z /=; rewrite in_itv/=; move: z => [z [[]|[|]]]//= ?. - + by do 2 left; split => //; exists tt. - + by right; split => //; exists true. - + by left; right; split => //; exists false. - - move=> z /=; rewrite in_itv/=; move: z => [z [[]|[|]]]//=. - - move=> [[[]//|]|]. - + by move=> [_ []]. - + by move=> [_ []]. - - move=> [[|]|[]//]. - + by move=> [_ []]. - + by move=> [_ [] [|]]. - - move=> [[|[]//]|]. - + by move=> [_ []]. - + by move=> [_ [] [|]]. -pose h1 := [set xub : X * (unit + bool) | k (inl tt) xub.1 U < x%:E]. -have mh1 : measurable h1. - rewrite -[X in measurable X]setTI; apply: emeasurable_fun_infty_o => //=. - have H : measurable_fun [set: X] (fun x => k (inl tt) x U) by exact/measurable_kernel. - move=> _ /= C mC; rewrite setTI. - have := H measurableT _ mC; rewrite setTI => {}H. - rewrite [X in measurable X](_ : _ = ((fun x => k (inl tt) x U) @^-1` C) `*` setT)//. - exact: measurableX. - by apply/seteqP; split => [z//=| z/= []]. -set h2 := [set xub : X * (unit + bool)| k (inr false) xub.1 U < x%:E]. -have mh2 : measurable h2. - rewrite -[X in measurable X]setTI. - apply: emeasurable_fun_infty_o => //=. - have H : measurable_fun [set: X] (fun x => k (inr false) x U) by exact/measurable_kernel. - move=> _ /= C mC; rewrite setTI. - have := H measurableT _ mC; rewrite setTI => {}H. - rewrite [X in measurable X](_ : _ = ((fun x => k (inr false) x U) @^-1` C) `*` setT)//. - exact: measurableX. - by apply/seteqP; split => [z //=|z/= []]. -set h3 := [set xub : X * (unit + bool)| k (inr true) xub.1 U < x%:E]. -have mh3 : measurable h3. - rewrite -[X in measurable X]setTI. - apply: emeasurable_fun_infty_o => //=. - have H : measurable_fun [set: X] (fun x => k (inr true) x U) by exact/measurable_kernel. - move=> _ /= C mC; rewrite setTI. - have := H measurableT _ mC; rewrite setTI => {}H. - rewrite [X in measurable X](_ : _ = ((fun x => k (inr true) x U) @^-1` C) `*` setT)//. - exact: measurableX. - by apply/seteqP; split=> [z//=|z/= []]. -apply: measurableU. -- apply: measurableU. - + apply: measurableX => //. - rewrite [X in measurable X](_ : _ = ysection h1 (inl tt))//. - * by apply: measurable_ysection. - * by apply/seteqP; split => z /=; rewrite /ysection /= inE. - + apply: measurableX => //. - rewrite [X in measurable X](_ : _ = ysection h2 (inr false))//. - * by apply: measurable_ysection. - * by apply/seteqP; split => z /=; rewrite /ysection /= inE. -- apply: measurableX => //. - rewrite [X in measurable X](_ : _ = ysection h3 (inr true))//. - + by apply: measurable_ysection. - + by apply/seteqP; split => z /=; rewrite /ysection /= inE. +have /measure_fam_uubP[r hr] := measure_uub k. +apply/measure_fam_uubP; exists (PosNum [gt0 of r%:num%R]) => x /=. +exact: hr. Qed. -#[export] -HB.instance Definition _ := isKernel.Build _ _ _ _ _ - (case_sum') (measurable_fun_case_sum'). -End kcase_sum'. +HB.instance Definition _ := + Kernel_isFinite.Build _ _ _ Z R (mkswap k) mkswap_finite. -Section sfkcase_sum'. -Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Let A : measurableType _ := unit. -Let B : measurableType _ := bool. -Variables (k : (A + B)%type -> R.-sfker X ~> Y). +End kswap_finite_kernel_finite. -Let sfinite_case_sum' : exists2 k_ : (R.-ker _ ~> _)^nat, - forall n, measure_fam_uub (k_ n) & - forall x U, measurable U -> case_sum' k x U = mseries (k_ ^~ x) 0 U. -Proof. -rewrite /=. -set f : A + B -> (R.-fker _ ~> _)^nat := - fun ab : A + B => sval (cid (sfinite_kernel (k ab))). -set Hf := fun ab : A + B => svalP (cid (sfinite_kernel (k ab))). -rewrite /= in Hf. -exists (fun n => [the R.-ker _ ~> _ of case_sum' (fun ab => [the R.-fker _ ~> _ of f ab n])]). - move=> n /=. - have [rtt Hrtt] := measure_uub (f (inl tt) n). - have [rfalse Hrfalse] := measure_uub (f (inr false) n). - have [rtrue Hrtrue] := measure_uub (f (inr true) n). - exists (maxr rtt (maxr rfalse rtrue)) => //= -[x [[]|[|]]] /=. - by rewrite 2!EFin_max lt_max Hrtt. - by rewrite /CASE_SUM.case_sumr /= 2!EFin_max 2!lt_max Hrtrue 2!orbT. - by rewrite /CASE_SUM.case_sumr /= 2!EFin_max 2!lt_max Hrfalse orbT. -move=> [x [[]|[|]]] U mU/=-. -by rewrite (Hf (inl tt) x _ mU). -by rewrite (Hf (inr true) x _ mU). -by rewrite (Hf (inr false) x _ mU). -Qed. +Reserved Notation "f .; g" (at level 60, right associativity, + format "f .; '/ ' g"). -#[export] -HB.instance Definition _ := @isSFiniteKernel_subdef.Build _ _ _ _ _ - (case_sum' k) sfinite_case_sum'. -End sfkcase_sum'. +Notation "l .; k" := (mkcomp l (mkswap k)) : ereal_scope. -End case_sum'. +Section letin'. +Variables (d d' d3 : _) (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). -Section case_sum. -Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Let A : measurableType _ := unit. -Let B : measurableType _ := bool. +Definition letin' (l : R.-sfker X ~> Y) (k : R.-sfker Y * X ~> Z) := + locked [the R.-sfker X ~> Z of l .; k]. -Import CASE_SUM. +Lemma letin'E (l : R.-sfker X ~> Y) (k : R.-sfker Y * X ~> Z) x U : + letin' l k x U = \int[l x]_y k (y, x) U. +Proof. by rewrite /letin'; unlock. Qed. -(* case analysis on the datatype unit + bool *) -Definition case_sum (f : R.-sfker X ~> (A + B)%type) - (k : (A + B)%type-> R.-sfker X ~> Y) : R.-sfker X ~> Y := - [the R.-sfker X ~> (A + B)%type of f] \; - [the R.-sfker X * (A + B) ~> Y of case_sum' k]. +Lemma letin'_letin (l : R.-sfker X ~> Y) (k : R.-sfker Y * X ~> Z) : + letin' l k = letin l (mkswap k). +Proof. by rewrite /letin'; unlock. Qed. -End case_sum. +End letin'. -(* counting measure as a kernel *) -Section kcounting. -Context d (G : measurableType d) (R : realType). +Section letin'C. +Import Notations. +Context d d1 d' (X : measurableType d) (Y : measurableType d1) + (Z : measurableType d') (R : realType). +Variables (t : R.-sfker Z ~> X) + (u' : R.-sfker X * Z ~> Y) + (u : R.-sfker Z ~> Y) + (t' : R.-sfker Y * Z ~> X) + (tt' : forall y, t =1 fun z => t' (y, z)) + (uu' : forall x, u =1 fun z => u' (x, z)). + +Definition T' z : set X -> \bar R := t z. +Let T0 z : (T' z) set0 = 0. Proof. by []. Qed. +Let T_ge0 z x : 0 <= (T' z) x. Proof. by []. Qed. +Let T_semi_sigma_additive z : semi_sigma_additive (T' z). +Proof. exact: measure_semi_sigma_additive. Qed. +HB.instance Definition _ z := @isMeasure.Build _ X R (T' z) (T0 z) (T_ge0 z) + (@T_semi_sigma_additive z). -Definition kcounting : G -> {measure set nat -> \bar R} := fun=> counting. +Let sfinT z : sfinite_measure (T' z). Proof. exact: sfinite_kernel_measure. Qed. +HB.instance Definition _ z := @isSFinite.Build _ X R (T' z) (sfinT z). -Let mkcounting U : measurable U -> measurable_fun setT (kcounting ^~ U). -Proof. by []. Qed. +Definition U' z : set Y -> \bar R := u z. +Let U0 z : (U' z) set0 = 0. Proof. by []. Qed. +Let U_ge0 z x : 0 <= (U' z) x. Proof. by []. Qed. +Let U_semi_sigma_additive z : semi_sigma_additive (U' z). +Proof. exact: measure_semi_sigma_additive. Qed. +HB.instance Definition _ z := @isMeasure.Build _ Y R (U' z) (U0 z) (U_ge0 z) + (@U_semi_sigma_additive z). -HB.instance Definition _ := isKernel.Build _ _ _ _ _ kcounting mkcounting. +Let sfinU z : sfinite_measure (U' z). Proof. exact: sfinite_kernel_measure. Qed. +HB.instance Definition _ z := @isSFinite.Build _ Y R + (U' z) (sfinU z). -Let sfkcounting : exists2 k_ : (R.-ker _ ~> _)^nat, - forall n, measure_fam_uub (k_ n) & - forall x U, measurable U -> kcounting x U = mseries (k_ ^~ x) 0 U. +Lemma letin'C z A : measurable A -> + letin' t + (letin' u' + (ret (measurable_fun_pair macc1of3' macc0of3'))) z A = + letin' u + (letin' t' + (ret (measurable_fun_pair macc0of3' macc1of3'))) z A. Proof. -exists (fun n => [the R.-fker _ ~> _ of - @kdirac _ _ G nat R _ (@measurable_cst _ _ _ _ setT n)]). - by move=> n /=; exact: measure_uub. -by move=> g U mU; rewrite /kcounting/= counting_dirac. +move=> mA. +rewrite !letin'E. +under eq_integral. + move=> x _. + rewrite letin'E -uu'. + under eq_integral do rewrite retE /=. + over. +rewrite (sfinite_Fubini (T' z) (U' z) (fun x => \d_(x.1, x.2) A ))//; last first. + apply/measurable_EFinP => /=; rewrite (_ : (fun x => _) = mindic R mA)//. + by apply/funext => -[]. +rewrite /=. +apply: eq_integral => y _. +by rewrite letin'E/= -tt'; apply: eq_integral => // x _; rewrite retE. Qed. -HB.instance Definition _ := - isSFiniteKernel_subdef.Build _ _ _ _ R kcounting sfkcounting. - -End kcounting. - -(* formalization of the iterate construct of Staton ESOP 2017, Sect. 4.2 *) -Section iterate. -Context d {G : measurableType d} {R : realType}. -Let A : measurableType _ := unit. -Let B : measurableType _ := bool. - -(* formalization of iterate^n - Gamma |-p iterate^n t from x = u : B *) -Variables (t : R.-sfker (G * A) ~> (A + B)%type) - (u : G -> A) (mu : measurable_fun setT u). +End letin'C. +Arguments letin'C {d d1 d' X Y Z R} _ _ _ _. -Fixpoint iterate_ n : R.-sfker G ~> B := - match n with - | 0%N => case_sum (letin (ret mu) t) - (fun x => match x with - | inl a => fail - | inr b => ret (measurable_cst b) - end) - | m.+1 => case_sum (letin (ret mu) t) - (fun x => match x with - | inl a => iterate_ m - | inr b => fail - end) - end. +Section letin'A. +Context d d' d1 d2 d3 (X : measurableType d) (Y : measurableType d') + (T1 : measurableType d1) (T2 : measurableType d2) (T3 : measurableType d3) + (R : realType). +Import Notations. +Variables (t : R.-sfker X ~> T1) + (u : R.-sfker T1 * X ~> T2) + (v : R.-sfker T2 * X ~> Y) + (v' : R.-sfker T2 * (T1 * X) ~> Y) + (vv' : forall y, v =1 fun xz => v' (xz.1, (y, xz.2))). -(* formalization of iterate (A = unit, B = bool) - Gamma, x : A |-p t : A + B Gamma |-d u : A ------------------------------------------------ - Gamma |-p iterate t from x = u : B *) -Definition iterate : R.-sfker G ~> B := case_nat (kcounting R) iterate_. +Lemma letin'A x A : measurable A -> + letin' t (letin' u v') x A + = + (letin' (letin' t u) v) x A. +Proof. +move=> mA. +rewrite !letin'E. +under eq_integral do rewrite letin'E. +rewrite letin'_letin/=. +rewrite integral_kcomp; [|by []|]. + apply: eq_integral => z _. + apply: eq_integral => y _. + by rewrite (vv' z). +exact: measurableT_comp (@measurable_kernel _ _ _ _ _ v _ mA) _. +Qed. -End iterate. +End letin'A. -(* an s-finite kernel to test that two expressions are different *) -Section lift_neq. -Context {R : realType} d (G : measurableType d). -Variables (f : G -> bool) (g : G -> bool). +Lemma letin'_sample_bernoulli d d' (T : measurableType d) + (T' : measurableType d') (R : realType) (r : R) (r01 : (0 <= r <= 1)%R) + (u : R.-sfker bool * T ~> T') x y : + letin' (sample_cst (bernoulli r)) u x y = + r%:E * u (true, x) y + (`1- r)%:E * u (false, x) y. +Proof. by rewrite letin'_letin letin_sample_bernoulli. Qed. -Definition flift_neq : G -> bool := fun x' => f x' != g x'. +Section letin'_return. +Context d d' d3 (X : measurableType d) (Y : measurableType d') + (Z : measurableType d3) (R : realType). -Hypotheses (mf : measurable_fun setT f) (mg : measurable_fun setT g). +Lemma letin'_kret (k : R.-sfker X ~> Y) + (f : Y * X -> Z) (mf : measurable_fun setT f) x U : + measurable U -> + letin' k (ret mf) x U = k x (curry f ^~ x @^-1` U). +Proof. +move=> mU. +rewrite letin'E. +under eq_integral do rewrite retE. +rewrite integral_indic ?setIT// -[X in measurable X]setTI. +exact: (measurableT_comp mf). +Qed. -(* see also emeasurable_fun_neq *) -Lemma measurable_fun_flift_neq : measurable_fun setT flift_neq. +Lemma letin'_retk (f : X -> Y) (mf : measurable_fun setT f) + (k : R.-sfker Y * X ~> Z) x U : + measurable U -> letin' (ret mf) k x U = k (f x, x) U. Proof. -apply: (@measurable_fun_bool _ _ _ _ true). -rewrite setTI. -rewrite /flift_neq /= (_ : _ @^-1` _ = ([set x | f x] `&` [set x | ~~ g x]) `|` - ([set x | ~~ f x] `&` [set x | g x])). - apply: measurableU; apply: measurableI. - - by rewrite -[X in measurable X]setTI; exact: mf. - - rewrite [X in measurable X](_ : _ = ~` [set x | g x]); last first. - by apply/seteqP; split => x /= /negP. - by apply: measurableC; rewrite -[X in measurable X]setTI; exact: mg. - - rewrite [X in measurable X](_ : _ = ~` [set x | f x]); last first. - by apply/seteqP; split => x /= /negP. - by apply: measurableC; rewrite -[X in measurable X]setTI; exact: mf. - - by rewrite -[X in measurable X]setTI; exact: mg. -by apply/seteqP; split => x /=; move: (f x) (g x) => [|] [|]//=; intuition. +move=> mU; rewrite letin'E retE integral_dirac ?diracT ?mul1e//. +exact: (measurableT_comp (measurable_kernel k _ mU)). Qed. -Definition lift_neq : R.-sfker G ~> bool := ret measurable_fun_flift_neq. +End letin'_return. -End lift_neq. +Section letin'_ite. +Context d d2 d3 (T : measurableType d) (T2 : measurableType d2) + (Z : measurableType d3) (R : realType). +Variables (k1 k2 : R.-sfker T ~> Z) + (u : R.-sfker Z * T ~> T2) + (f : T -> bool) (mf : measurable_fun setT f) + (t : T) (U : set T2). -Section von_neumann_trick. -Context d {T : measurableType d} {R : realType}. +Lemma letin'_iteT : f t -> letin' (ite mf k1 k2) u t U = letin' k1 u t U. +Proof. by move=> ftT; rewrite !letin'_letin letin_iteT. Qed. -Definition kinrtt {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} := - @measurable_cst _ _ T1 _ setT (@inl unit T2 tt). +Lemma letin'_iteF : ~~ f t -> letin' (ite mf k1 k2) u t U = letin' k2 u t U. +Proof. by move=> ftF; rewrite !letin'_letin letin_iteF. Qed. -Definition finlb d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) - : T1 * bool -> T2 + bool := fun t1b => inr t1b.2. +End letin'_ite. -Lemma minlb {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} : - measurable_fun setT (@finlb _ _ T1 T2). -Proof. exact: measurableT_comp. Qed. +Section hard_constraint'. +Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Variable (D : pprobability bool R (* biased coin *)). +Definition fail' : R.-sfker X ~> Y := + letin' (score (measurable_cst (0%R : R))) + (ret (measurable_cst point)). -Definition von_neumann_trick' : R.-sfker (T * unit) ~> (unit + bool)%type := - letin (sample_cst D) - (letin (sample_cst D) - (letin (lift_neq macc1of3 macc2of3) - (ite macc3of4 - (letin (ret macc1of4) (ret minlb)) - (ret kinrtt)))). +Lemma fail'E x U : fail' x U = 0. +Proof. by rewrite /fail' letin'_letin failE. Qed. -Definition von_neumann_trick : R.-sfker T ~> bool := - iterate von_neumann_trick' ktt. +End hard_constraint'. +Arguments fail' {d d' X Y R}. -End von_neumann_trick. +Lemma score_fail' d (X : measurableType d) {R : realType} + (r : R) (r01 : (0 <= r <= 1)%R) : + score (kr r) = + letin' (sample_cst (bernoulli r) : R.-pker X ~> _) + (ite macc0of2 (ret ktt) fail'). +Proof. +move: r01 => /andP[r0 r1]; apply/eq_sfkernel => x U. +rewrite letin'E/= /sample; unlock. +rewrite integral_bernoulli ?r0//=. +by rewrite /mscale/= iteE//= iteE//= fail'E mule0 adde0 ger0_norm. +Qed. diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 79d2c24888..a9b321353a 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -3,13 +3,14 @@ From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. From mathcomp Require Import rat interval_inference. From mathcomp Require Import mathcomp_extra boolp classical_sets. From mathcomp Require Import functions cardinality fsbigop. -From mathcomp Require Import reals ereal topology normedtype sequences. -From mathcomp Require Import esum measure lebesgue_measure numfun. +From mathcomp Require Import interval_inference reals ereal topology normedtype. +From mathcomp Require Import sequences esum measure lebesgue_measure numfun. From mathcomp Require Import lebesgue_integral exp kernel trigo prob_lang. +From mathcomp Require Import realfun charge. (******************************************************************************) (* Semantics of a probabilistic programming language using s-finite kernels *) -(* (wip about definition of Lebesgue and counting measures) *) +(* (wip about the definition of Lebesgue measure) *) (******************************************************************************) Set Implicit Arguments. @@ -22,112 +23,187 @@ Local Open Scope classical_set_scope. Local Open Scope ring_scope. Local Open Scope ereal_scope. -Section gauss. -Variable R : realType. +Section gauss_pdf. +Context {R : realType}. Local Open Scope ring_scope. -(* density function for gauss *) -Definition gauss_density m s x : R := +Definition gauss_pdf m s x : R := (s * sqrtr (pi *+ 2))^-1 * expR (- ((x - m) / s) ^+ 2 / 2%:R). -Lemma gauss_density_ge0 m s x : 0 <= s -> 0 <= gauss_density m s x. +Lemma gauss_pdf_ge0 m s x : 0 <= s -> 0 <= gauss_pdf m s x. Proof. by move=> s0; rewrite mulr_ge0 ?expR_ge0// invr_ge0 mulr_ge0. Qed. -Lemma gauss_density_gt0 m s x : 0 < s -> 0 < gauss_density m s x. +Lemma gauss_pdf_gt0 m s x : 0 < s -> 0 < gauss_pdf m s x. Proof. move=> s0; rewrite mulr_gt0 ?expR_gt0// invr_gt0 mulr_gt0//. by rewrite sqrtr_gt0 pmulrn_rgt0// pi_gt0. Qed. -Definition gauss01_density : R -> R := gauss_density 0 1. +Lemma measurable_gauss_pdf m s : measurable_fun setT (gauss_pdf m s). +Proof. +apply: measurable_funM => //=; apply: measurableT_comp => //=. +apply: measurable_funM => //=; apply: measurableT_comp => //=. +apply: measurableT_comp (exprn_measurable _) _ => /=. +by apply: measurable_funM => //=; exact: measurable_funD. +Qed. -Hypothesis integral_gauss01_density : - (\int[lebesgue_measure]_x (gauss01_density x)%:E = 1%E)%E. +Definition gauss_pdf01 : R -> R := gauss_pdf 0 1. -Lemma gauss01_densityE x : - gauss01_density x = (sqrtr (pi *+ 2))^-1 * expR (- (x ^+ 2) / 2%:R). -Proof. by rewrite /gauss01_density /gauss_density mul1r subr0 divr1. Qed. +Lemma gauss_pdf01E x : + gauss_pdf01 x = (sqrtr (pi *+ 2))^-1 * expR (- (x ^+ 2) / 2%:R). +Proof. by rewrite /gauss_pdf01 /gauss_pdf mul1r subr0 divr1. Qed. -Definition mgauss01 (V : set R) := - (\int[lebesgue_measure]_(x in V) (gauss01_density x)%:E)%E. +Lemma gauss_pdf01_ub x : gauss_pdf01 x <= (Num.sqrt (pi *+ 2))^-1. +Proof. +rewrite -[leRHS]mulr1. +rewrite /gauss_pdf01 /gauss_pdf; last first. +rewrite mul1r subr0 ler_pM2l ?invr_gt0// ?sqrtr_gt0; last by rewrite mulrn_wgt0// pi_gt0. +by rewrite -[leRHS]expR0 ler_expR mulNr oppr_le0 mulr_ge0// sqr_ge0. +Qed. -Lemma measurable_fun_gauss_density m s : - measurable_fun setT (gauss_density m s). +Lemma continuous_gauss_pdf1 x : {for x, continuous gauss_pdf01}. Proof. -apply: measurable_funM => //=. -apply: measurableT_comp => //=. -apply: measurable_funM => //=. -apply: measurableT_comp => //=. -apply: measurableT_comp (exprn_measurable _) _ => /=. -apply: measurable_funM => //=. -exact: measurable_funD. +apply: continuousM => //=; first exact: cvg_cst. +apply: continuous_comp => /=; last exact: continuous_expR. +apply: continuousM => //=; last exact: cvg_cst. +apply: continuous_comp => //=; last exact: (@continuousN _ R^o). +apply: (@continuous_comp _ _ _ _ (fun x : R => x ^+ 2)%R); last exact: exprn_continuous. +apply: continuousM => //=; last exact: cvg_cst. +by apply: (@continuousD _ R^o) => //=; exact: cvg_cst. Qed. -Let mgauss010 : mgauss01 set0 = 0%E. -Proof. by rewrite /mgauss01 integral_set0. Qed. +End gauss_pdf. + +Definition gauss01 {R : realType} + of \int[@lebesgue_measure R]_x (gauss_pdf01 x)%:E = 1%E : set _ -> \bar R := + fun V => (\int[lebesgue_measure]_(x in V) (gauss_pdf01 x)%:E)%E. + +Section gauss. +Variable R : realType. +Local Open Scope ring_scope. + +Hypothesis integral_gauss_pdf01 : + (\int[@lebesgue_measure R]_x (gauss_pdf01 x)%:E = 1%E)%E. + +Local Notation gauss01 := (gauss01 integral_gauss_pdf01). + +Let gauss010 : gauss01 set0 = 0%E. +Proof. by rewrite /gauss01 integral_set0. Qed. -Let mgauss01_ge0 A : (0 <= mgauss01 A)%E. +Let gauss01_ge0 A : (0 <= gauss01 A)%E. Proof. -by rewrite /mgauss01 integral_ge0//= => x _; rewrite lee_fin gauss_density_ge0. +by rewrite /gauss01 integral_ge0//= => x _; rewrite lee_fin gauss_pdf_ge0. Qed. -Let mgauss01_sigma_additive : semi_sigma_additive mgauss01. +Let gauss01_sigma_additive : semi_sigma_additive gauss01. Proof. move=> /= F mF tF mUF. -rewrite /mgauss01/= integral_bigcup//=; last first. +rewrite /gauss01/= integral_bigcup//=; last first. apply/integrableP; split. - apply/measurable_EFinP. - exact: measurable_funS (measurable_fun_gauss_density 0 1). - rewrite (_ : (fun x => _) = (EFin \o gauss01_density)); last first. - by apply/funext => x; rewrite gee0_abs// lee_fin gauss_density_ge0. + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_gauss_pdf. + rewrite (_ : (fun x => _) = EFin \o gauss_pdf01); last first. + by apply/funext => x; rewrite gee0_abs// lee_fin gauss_pdf_ge0. apply: le_lt_trans. apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. - apply/measurable_EFinP. - exact: measurable_fun_gauss_density. - by move=> ? _; rewrite lee_fin gauss_density_ge0. - by rewrite integral_gauss01_density// ltey. + by apply/measurable_EFinP; exact: measurable_gauss_pdf. + by move=> ? _; rewrite lee_fin gauss_pdf_ge0. + by rewrite integral_gauss_pdf01 // ltey. apply: is_cvg_ereal_nneg_natsum_cond => n _ _. -by apply: integral_ge0 => /= x ?; rewrite lee_fin gauss_density_ge0. +by apply: integral_ge0 => /= x ?; rewrite lee_fin gauss_pdf_ge0. Qed. HB.instance Definition _ := isMeasure.Build _ _ _ - mgauss01 mgauss010 mgauss01_ge0 mgauss01_sigma_additive. - -Let mgauss01_setT : mgauss01 [set: _] = 1%E. -Proof. by rewrite /mgauss01 integral_gauss01_density. Qed. + gauss01 gauss010 gauss01_ge0 gauss01_sigma_additive. -HB.instance Definition _ := @Measure_isProbability.Build _ _ R mgauss01 mgauss01_setT. +Let gauss01_setT : gauss01 [set: _] = 1%E. +Proof. by rewrite /gauss01 integral_gauss_pdf01. Qed. -Definition gauss01 := [the probability _ _ of mgauss01]. +HB.instance Definition _ := @Measure_isProbability.Build _ _ R gauss01 gauss01_setT. End gauss. Section gauss_lebesgue. -Import Notations. Context d (T : measurableType d) (R : realType). -Hypothesis integral_gauss01_density : - (\int[@lebesgue_measure R]_x (gauss01_density x)%:E = 1%E)%E. +Notation mu := (@lebesgue_measure R). +Hypothesis integral_gauss_pdf01 : \int[mu]_x (gauss_pdf01 x)%:E = 1%E. + +Lemma gauss01_dom : gauss01 integral_gauss_pdf01 `<< mu. +Proof. +move=> A mA muA0; rewrite /gauss01. +apply/eqP; rewrite eq_le; apply/andP; split; last first. + by apply: integral_ge0 => x _; rewrite lee_fin gauss_pdf_ge0. +apply: (@le_trans _ _ (\int[mu]_(x in A) (Num.sqrt (pi *+ 2))^-1%:E))%E; last first. + by rewrite integral_cst//= muA0 mule0. +apply: ge0_le_integral => //=. +- by move=> x _; rewrite lee_fin gauss_pdf_ge0. +- apply/measurable_funTS/measurableT_comp => //. + exact: measurable_gauss_pdf. +- by move=> x _; rewrite lee_fin gauss_pdf01_ub. +Qed. -Let f1 (x : R) := ((gauss01_density x) ^-1)%R. +Let f1 (x : g_sigma_algebraType (R.-ocitv.-measurable)) := ((gauss_pdf01 x) ^-1)%R. -Hypothesis integral_mgauss01 : forall U, measurable U -> - \int[mgauss01 (R:=R)]_(y in U) (f1 y)%:E = - \int[lebesgue_measure]_(x0 in U) (gauss01_density x0 * f1 x0)%:E. +Lemma measurable_fun_f1 : measurable_fun setT f1. +Proof. +apply: continuous_measurable_fun => x. +apply: (@continuousV _ _ gauss_pdf01). + by rewrite gt_eqF// gauss_pdf_gt0. +exact: continuous_gauss_pdf1. +Qed. + +Lemma integrable_f1 U : measurable U -> + (gauss01 integral_gauss_pdf01).-integrable U (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (f1 x)%:E). +Proof. +Admitted. + +Lemma integral_mgauss01 : forall U, measurable U -> + \int[gauss01 integral_gauss_pdf01]_(y in U) (f1 y)%:E = + \int[mu]_(x0 in U) (gauss_pdf01 x0 * f1 x0)%:E. +Proof. +move=> U mU. +under [in RHS]eq_integral do rewrite EFinM/= muleC. +rewrite -(Radon_Nikodym_change_of_variables gauss01_dom _ (integrable_f1 mU))//=. +apply: ae_eq_integral => //=. +- apply: emeasurable_funM => //. + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_fun_f1. + apply: measurable_int. + apply: integrableS (Radon_Nikodym_integrable _) => //=. + exact: gauss01_dom. +- apply: emeasurable_funM => //. + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_fun_f1. + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_gauss_pdf. +- apply: ae_eqe_mul2l => /=. + rewrite Radon_NikodymE//=. + exact: gauss01_dom. + move=> gauss01_dom'. + case: cid => //= h [h1 h2 h3]. + apply: integral_ae_eq => //=. + + exact: integrableS h2. + + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_gauss_pdf. + + by move=> E EU mE; rewrite -(h3 _ mE). +Qed. + +(*Hypothesis integral_mgauss01 : forall U, measurable U -> + \int[gauss01 integral_gauss_pdf01]_(y in U) (f1 y)%:E = + \int[mu]_(x0 in U) (gauss_pdf01 x0 * f1 x0)%:E.*) Let mf1 : measurable_fun setT f1. Proof. apply: (measurable_comp (F := [set r : R | r != 0%R])) => //. - exact: open_measurable. -- by move=> /= r [t _ <-]; rewrite gt_eqF// gauss_density_gt0. +- by move=> /= r [t _ <-]; rewrite gt_eqF// gauss_pdf_gt0. - apply: open_continuous_measurable_fun => //. by apply/in_setP => x /= x0; exact: inv_continuous. -- exact: measurable_fun_gauss_density. +- exact: measurable_gauss_pdf. Qed. -Variable mu : {measure set R -> \bar R}. - Definition staton_lebesgue : R.-sfker T ~> _ := - letin (sample_cst (gauss01 integral_gauss01_density : pprobability _ _)) + letin (sample_cst (gauss01 integral_gauss_pdf01 : pprobability _ _)) (letin (score (measurableT_comp mf1 macc1of2)) (ret macc1of3)). @@ -137,23 +213,22 @@ Lemma staton_lebesgueE x U : measurable U -> Proof. move=> mU; rewrite [in LHS]/staton_lebesgue/=. rewrite [in LHS]letinE /=. -transitivity (\int[@mgauss01 R]_(y in U) (f1 y)%:E). +transitivity (\int[gauss01 integral_gauss_pdf01]_(y in U) (f1 y)%:E). rewrite -[in RHS](setTI U) integral_mkcondr/=. apply: eq_integral => //= r _. rewrite letinE/= ge0_integral_mscale//= ger0_norm//; last first. - by rewrite invr_ge0// gauss_density_ge0. - rewrite integral_dirac// diracT mul1e/= diracE epatch_indic/=. + by rewrite invr_ge0// gauss_pdf_ge0. + rewrite integral_dirac// diracT mul1e/= diracE epatch_indic/=. by rewrite indicE. rewrite integral_mgauss01//. transitivity (\int[lebesgue_measure]_(x in U) (\1_U x)%:E). apply: eq_integral => /= y yU. - by rewrite /f1 divrr ?indicE ?yU// unitfE gt_eqF// gauss_density_gt0. + by rewrite /f1 divrr ?indicE ?yU// unitfE gt_eqF// gauss_pdf_gt0. by rewrite integral_indic//= setIid. Qed. End gauss_lebesgue. -(* TODO: move this elsewhere *) (* assuming x > 0 *) Definition Gamma {R : realType} (x : R) : \bar R := \int[lebesgue_measure]_(t in `[0%R, +oo[%classic) (expR (- t) * powR t (x - 1))%:E. @@ -163,20 +238,21 @@ Definition Rfact {R : realType} (x : R) := Gamma (x + 1)%R. Section poisson. Variable R : realType. Local Open Scope ring_scope. +Notation mu := (@lebesgue_measure R). Hypothesis integral_poisson_density : forall k, - (\int[lebesgue_measure]_x (@poisson R k x)%:E = 1%E)%E. + (\int[mu]_x (@poisson_pdf R k x)%:E = 1%E)%E. (* density function for poisson *) -Definition poisson1 := @poisson R 1%N. +Definition poisson1 := @poisson_pdf R 1%N. Lemma poisson1_ge0 (x : R) : 0 <= poisson1 x. -Proof. exact: poisson_ge0. Qed. +Proof. exact: poisson_pdf_ge0. Qed. Definition mpoisson1 (V : set R) : \bar R := (\int[lebesgue_measure]_(x in V) (poisson1 x)%:E)%E. Lemma measurable_fun_poisson1 : measurable_fun setT poisson1. -Proof. exact: measurable_poisson. Qed. +Proof. exact: measurable_poisson_pdf. Qed. Let mpoisson10 : mpoisson1 set0 = 0%E. Proof. by rewrite /mpoisson1 integral_set0. Qed. @@ -215,7 +291,8 @@ rewrite /poisson1. by rewrite integral_poisson_density. Qed. -HB.instance Definition _ := @Measure_isProbability.Build _ _ R mpoisson1 mpoisson1_setT. +HB.instance Definition _ := @Measure_isProbability.Build _ _ R + mpoisson1 mpoisson1_setT. Definition poisson' := [the probability _ _ of mpoisson1]. @@ -226,21 +303,22 @@ End poisson. Section staton_counting. Context d (X : measurableType d). Variable R : realType. +Notation mu := (@lebesgue_measure R). Import Notations. Hypothesis integral_poisson_density : forall k, - (\int[lebesgue_measure]_x (@poisson R k x)%:E = 1%E)%E. + (\int[mu]_x (@poisson_pdf R k x)%:E = 1%E)%E. Let f1 x := ((poisson1 (x : R)) ^-1)%R. Let mf1 : measurable_fun setT f1. -rewrite /f1 /poisson1 /poisson. +rewrite /f1 /poisson1 /poisson_pdf. apply: (measurable_comp (F := [set r : R | r != 0%R])) => //. - exact: open_measurable. - move=> /= r [t ? <-]. by case: ifPn => // t0; rewrite gt_eqF ?mulr_gt0 ?expR_gt0//= invrK ltr0n. - apply: open_continuous_measurable_fun => //. by apply/in_setP => x /= x0; exact: inv_continuous. -- exact: measurable_poisson. +- exact: measurable_poisson_pdf. Qed. Definition staton_counting : R.-sfker X ~> _ := From 8c327a2599f6d7f9d9dcc01a2b4b7c4373c3a960 Mon Sep 17 00:00:00 2001 From: IshiguroYoshihiro Date: Sun, 25 Aug 2024 16:29:03 +0900 Subject: [PATCH 22/48] beta_bdf_uniq_ae, integral_indicator_function --- theories/lang_syntax.v | 178 +++++++++++++++++++++++++++--- theories/lang_syntax_examples.v | 6 +- theories/lang_syntax_table_game.v | 8 +- 3 files changed, 169 insertions(+), 23 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 089915ef6f..0c4dace8d1 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -75,6 +75,130 @@ Local Open Scope classical_set_scope. Local Open Scope ring_scope. Local Open Scope ereal_scope. +(* In this module, we use our lemma continuous_FTC2 to compute the value of + * integration of the indicator function over the interval [0, 1]. + * we can use our lemma continuous_FTC2 because it requires continuous + * within [0, 1], which the indicator function satisfies. + * we also shows that the indicator function is not continuous in [0, 1], + * required by previous version of lemma continuous_FTC2. This shows that + * our lemma continuous_FTC2 is + * enough weak to be usable in practice. + *) +Module integral_indicator_function. +Section integral_indicator_function. + +Context {R : realType}. +Notation mu := lebesgue_measure. +Local Open Scope ereal_scope. +Implicit Types (f : R -> R) (a b : R). + +Local Import set_interval. + +Let uni := @indic R R `[0%R, 1%R]%classic. + +Let integrable_uni : mu.-integrable setT (EFin \o uni). +Proof. +apply/integrableP; split. + apply: measurableT_comp => //. + exact: measurable_indic. +apply/abse_integralP => //. + apply: measurableT_comp => //. + exact: measurable_indic. +rewrite -ge0_fin_numE; last exact: abse_ge0. +rewrite abse_fin_num integral_indic// setIT. +by rewrite /= lebesgue_measure_itv ifT. +Qed. + +Let cuni_within : {within `[0%R, 1%R], continuous uni}. +Proof. +apply/continuous_within_itvP => //; split => //. +- move=> x x01. + apply: (@near_cst_continuous R R 1%R). + near=> z. + rewrite /uni indic_restrict patchE ifT//. + rewrite inE/=. + apply: subset_itv_oo_cc. + near: z. + exact: near_in_itvoo. +- rewrite (_: uni 0 = 1%R); last first. + rewrite /uni indic_restrict patchE ifT//. + by rewrite inE/= boundl_in_itv bnd_simp/=. + apply: cvg_near_cst. + near=> z. + rewrite /uni indic_restrict patchE ifT// inE/= in_itv/=; apply/andP; split => //. + near: z. + exact: nbhs_right_le. +- rewrite (_:uni 1 = 1%R); last first. + rewrite /uni indic_restrict patchE ifT//. + by rewrite inE/= boundr_in_itv bnd_simp/=. + apply: cvg_near_cst. + near=> z. + rewrite /uni indic_restrict patchE ifT// inE/= in_itv/=; apply/andP; split => //. + near: z. + exact: nbhs_left_ge. +Unshelve. all: end_near. Qed. + +Example cuni : ~ {in `[0%R, 1%R], continuous uni}. +Proof. +rewrite -existsNE/=. +exists 0%R. +rewrite not_implyE; split; first by rewrite boundl_in_itv/= bnd_simp. +move/left_right_continuousP. +apply/not_andP; left. +move/(@cvgrPdist_le _ R^o). +apply/existsNP. +exists (2%:R^-1). +rewrite not_implyE; split; first by rewrite invr_gt0. +move=> [e /= e0]. +move/(_ (-(e / 2))%R). +apply/not_implyP; split. + rewrite /= sub0r opprK ger0_norm; last by rewrite divr_ge0// ltW. + rewrite -{1}(add0r e). + exact: (midf_lt e0).2. +apply/not_implyP; split. + rewrite oppr_lt0. + exact: divr_gt0. +apply/negP; rewrite -ltNge. +rewrite /uni !indic_restrict !patchE. +rewrite ifT; last by rewrite inE/= boundl_in_itv/= bnd_simp. +rewrite ifF; last first. + apply: negbTE; apply/negP. + rewrite inE/= in_itv/=. + apply/negP; rewrite negb_and; apply/orP; left. + by rewrite -ltNge oppr_lt0 divr_gt0. +rewrite /point/= {2}/1%R/= subr0. +rewrite ger0_norm//. +rewrite invf_lt1//. +rewrite {1}(_:1%R = 1%:R)//; apply: ltr_nat. +Qed. + +Let dintuni : derivable_oo_continuous_bnd (@id R^o) 0 1. +Proof. +split. +- move=> x _. + exact: derivable_id. +- exact: cvg_at_right_filter. +- exact: cvg_at_left_filter. +Qed. + +Let intuni'uni : {in `]0%R, 1%R[, (@id R^o)^`() =1 uni}. +Proof. +move=> x x01. +rewrite derive1E derive_id. +rewrite /uni indic_restrict patchE ifT// inE/=. +exact: subset_itv_oo_cc. +Qed. + +Lemma intuni1 : (\int[mu]_(x in `[0, 1]) uni x)%R = 1%R. +Proof. +rewrite [RHS](_:1%R = fine (1%:E))//; congr (fine _). +rewrite (continuous_FTC2 ltr01 cuni_within dintuni intuni'uni). +by rewrite sube0. +Qed. + +End integral_indicator_function. +End integral_indicator_function. + Lemma RintegralZl d {T : measurableType d} {R : realType} {mu : measure T R} {D : set T} : d.-measurable D -> forall f : T -> R, @@ -254,13 +378,11 @@ Qed. Lemma Ronem_change (G : R -> R) (r : R) : (0 < r <= 1)%R -> - locally_integrable [set: R] G -> {within `[0%R, r], continuous G} -> - (forall r, lebesgue_measure.-integrable `[0%R, r] (EFin \o G)) -> (\int[mu]_(x in `[0%R, r]) (G x) = \int[mu]_(x in `[(1 - r)%R, 1%R]) (G (1 - x)))%R. Proof. -move=> r01 locG cG iG. +move=> r01 cG. rewrite [in LHS]/Rintegral. by rewrite onem_change. Qed. @@ -708,19 +830,10 @@ Qed. Lemma betafun_sym (a b : nat) : betafun a b = betafun b a :> R. Proof. rewrite -[LHS]Rintegral_mkcond. -rewrite Ronem_change//=; last 4 first. +rewrite Ronem_change//=; last 2 first. by rewrite ltr01 lexx. - split. - - exact: measurable_fun_XMonemX. - - exact: openT. - - move=> K _ cK. - by have /integrableP[] := @continuous_compact_integrable R _ K cK - (@within_continuous_XMonemX R a.-1 b.-1 K). apply: continuous_subspaceT. by move=> x x01; exact: continuous_XMonemX. - move=> r. - exact: (@continuous_compact_integrable R _ _ - (@segment_compact _ _ _) (@within_continuous_XMonemX R a.-1 b.-1 `[0%R, r])). rewrite subrr. rewrite -[RHS]Rintegral_mkcond. apply: eq_Rintegral => x x01. @@ -1213,7 +1326,39 @@ apply: ge0_le_integral => //=; first exact: measurableI. - by move=> x _; rewrite lee_fin XMonemX01_le1. Qed. -Lemma integral_beta_prob a b f U : measurable U -> measurable_fun U f -> +Section beta_pdf_Beta. + +Local Open Scope charge_scope. + +(* beta_pdf is almost density function of Beta *) +Lemma beta_pdf_uniq_ae (a b : nat) : + ae_eq mu `[0%R, 1%R]%classic + ('d ((charge_of_finite_measure (@beta_prob R a b))) '/d mu) + (EFin \o (beta_pdf a b)). +Proof. +apply: integral_ae_eq => //. +- apply: integrableS (Radon_Nikodym_integrable _) => //. + exact: beta_prob_dom. +- apply/measurable_funTS/measurableT_comp => //. + exact: measurable_beta_pdf. +- move=> E E01 mE. + rewrite integral_beta_pdf//. + apply/esym. + rewrite -Radon_Nikodym_integral//=. + exact: beta_prob_dom. +Qed. + +(* need to add lemma about radon-nikodym derivative of + lebesgue_stieltjes measure w.r.t. continuous density function *) +(*Lemma beta_pdf_uniq (a b : nat) : + {in `[0%R, 1%R]%classic, + ('d ((charge_of_finite_measure (@Beta R a b))) '/d mu) =1 + (EFin \o (beta_pdf a b))}. +Proof. Abort.*) + +End beta_pdf_Beta. + +Lemma integral_Beta a b f U : measurable U -> measurable_fun U f -> \int[beta_prob a b]_(x in U) `|f x| < +oo -> \int[beta_prob a b]_(x in U) f x = \int[mu]_(x in U) (f x * (beta_pdf a b x)%:E) :> \bar R. Proof. @@ -1230,6 +1375,7 @@ apply: ae_eq_integral => //. rewrite Radon_NikodymE//=; first exact: beta_prob_dom. move=> ?. case: cid => /= h [h1 h2 h3]. +(* uniqueness of Radon-Nikodym derivertive up to equal on non null sets of mu *) apply: integral_ae_eq => //. + apply: integrableS h2 => //. (* integrableST? *) apply/measurable_funTS/measurableT_comp => //. @@ -1303,7 +1449,7 @@ rewrite integralD//=; last 2 first. congr (_ + _). rewrite integralZr//=; last exact: beta_prob_integrable. congr (_ * _)%E. - rewrite integral_beta_prob//; last 2 first. + rewrite integral_Beta//; last 2 first. by apply/measurableT_comp => //; exact: measurable_fun_XMonemX. by have /integrableP[_] := @beta_prob_integrable R a b c d. rewrite /beta_pdf. @@ -1330,7 +1476,7 @@ under eq_integral do rewrite muleC. rewrite /=. rewrite integralZl//=; last exact: beta_prob_integrable_onem. rewrite muleC; congr (_ * _)%E. -rewrite integral_beta_prob//=; last 2 first. +rewrite integral_Beta//=; last 2 first. apply/measurableT_comp => //=. by apply/measurable_funB => //; exact: measurable_fun_XMonemX. by have /integrableP[] := @beta_prob_integrable_onem R a b c d. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index 57335e3176..5fee74717b 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -612,12 +612,12 @@ rewrite execP_letin !execP_sample execD_beta !execD_bernoulli/=. rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. transitivity (beta_prob_bernoulli 6 4 1 0 U : \bar R). rewrite /beta_prob_bernoulli !letin'E/=. - rewrite integral_beta_prob//=; last 2 first. + rewrite integral_Beta//=; last 2 first. exact: measurable_bernoulli2. exact: integral_beta_prob_bernoulli_lty. - rewrite integral_beta_prob//=; last 2 first. + rewrite integral_Beta//=; last 2 first. by apply: measurable_funTS => /=; exact: measurable_bernoulli_XMonemX01. - rewrite integral_beta_prob//=. + rewrite integral_Beta//=. + suff: mu.-integrable `[0%R, 1%R] (fun x => bernoulli (XMonemX01 2 1 x) U * (beta_pdf 6 4 x)%:E)%E. move=> /integrableP[_]. diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v index cef5fe1cae..fb9c1f4f37 100644 --- a/theories/lang_syntax_table_game.v +++ b/theories/lang_syntax_table_game.v @@ -486,14 +486,14 @@ under [in RHS]eq_integral => y _. rewrite /=. rewrite [RHS]ge0_integral_mscale//=; last first. by move=> _ _; rewrite integral_ge0. -rewrite integral_beta_prob//=; last 2 first. +rewrite integral_Beta//=; last 2 first. - apply: emeasurable_funM => //=. exact: measurable_bernoulli_onemXn. apply/measurable_EFinP; apply: measurableT_comp => //. by apply: measurable_funM => //; exact: measurable_fun_XMonemX. - by have /integrableP[] := @integrable_bernoulli_XMonemX R U. rewrite ger0_norm// integral_dirac// diracT mul1e. -rewrite integral_beta_prob/=; [|by []|exact: measurable_bernoulli_onemXn +rewrite integral_Beta/=; [|by []|exact: measurable_bernoulli_onemXn |exact: integral_beta_prob_bernoulli_onem_lty]. rewrite -integralZl//=; last exact: integrable_bernoulli_beta_pdf. apply: eq_integral => y _. @@ -518,11 +518,11 @@ Lemma int_beta_prob01 {R : realType} (f : R -> R) a b U : \int[beta_prob a b]_(y in `[0%R, 1%R] : set R) bernoulli (f y) U. Proof. move=> mf f01. -rewrite [LHS]integral_beta_prob//=; last 2 first. +rewrite [LHS]integral_Beta//=; last 2 first. apply: measurable_funTS. by apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. exact: integral_beta_prob_bernoulli_lty. -rewrite [RHS]integral_beta_prob//; last 2 first. +rewrite [RHS]integral_Beta//; last 2 first. apply/measurable_funTS => //=. by apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. apply: (le_lt_trans _ (lang_syntax.integral_beta_prob_bernoulli_lty a b U mf f01)). From e533c3c5fa5563d20eb732570f74d8869f13320e Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 18 Sep 2024 15:18:29 +0900 Subject: [PATCH 23/48] fix --- theories/lang_syntax.v | 2 +- theories/prob_lang.v | 108 ----------------------------------------- 2 files changed, 1 insertion(+), 109 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 0c4dace8d1..9ec3cdce72 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -111,7 +111,7 @@ Qed. Let cuni_within : {within `[0%R, 1%R], continuous uni}. Proof. -apply/continuous_within_itvP => //; split => //. +apply/continuous_within_itvP => //; split. - move=> x x01. apply: (@near_cst_continuous R R 1%R). near=> z. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 07bc09b58a..8a1365a74d 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -80,35 +80,6 @@ Definition dep_uncurry (A : Type) (B : A -> Type) (C : Type) : (forall a : A, B a -> C) -> {a : A & B a} -> C := fun f p => let (a, Ba) := p in f a Ba. -Section binomial_total. -Local Open Scope ring_scope. -Variables (R : realType) (n : nat). -Implicit Type p : R. - -Lemma measurable_binomial_prob : - measurable_fun setT (binomial_prob n : R -> pprobability _ _). -Proof. -apply: (@measurability _ _ _ _ _ _ - (@pset _ _ _ : set (set (pprobability _ R)))) => //. -move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. -rewrite /binomial_prob/=. -set f := (X in measurable_fun _ X). -apply: measurable_fun_if => //=. - by apply: measurable_and => //; exact: measurable_fun_ler. -apply: (eq_measurable_fun (fun t => - \sum_(k x /set_mem[_/= x01]. - rewrite nneseries_esum// -1?[in RHS](set_mem_set Ys)// => k kYs. - by rewrite lee_fin binomial_pmf_ge0. -apply: ge0_emeasurable_sum. - by move=> k x/= [_ x01] _; rewrite lee_fin binomial_pmf_ge0. -move=> k Ysk; apply/measurableT_comp => //. -exact: measurable_binomial_pmf. -Qed. - -End binomial_total. -Arguments measurable_binomial_prob {R}. - Section poisson_pdf. Variable R : realType. Local Open Scope ring_scope. @@ -255,85 +226,6 @@ Arguments p14 {R}. (*Arguments p27 {R}.*) Arguments p1S {R}. -Section uniform_probability. -Context (R : realType) (a b : R) (ab0 : (0 < b - a)%R). - -Definition uniform_probability (* : set _ -> \bar R *) := - @mscale _ _ R (invr_nonneg (NngNum (ltW ab0))) - (mrestr lebesgue_measure (measurable_itv `[a, b])). - -(* NB: set R -> \bar R を書くとMeasure.onが通らない *) -HB.instance Definition _ := Measure.on uniform_probability. - -(* Let uniform0 : uniform_probability set0 = 0. -Proof. exact: measure0. Qed. - -Let uniform_ge0 U : 0 <= uniform_probability U. -Proof. exact: measure_ge0. Qed. - -Let uniform_sigma_additive : semi_sigma_additive uniform_probability. -Proof. move=> /= F mF tF mUF; exact: measure_semi_sigma_additive. Qed. - -HB.instance Definition _ := isMeasure.Build _ _ _ uniform_probability - uniform0 uniform_ge0 uniform_sigma_additive. *) - -Let uniform_probability_setT : uniform_probability [set: _] = 1%:E. -Proof. -rewrite /uniform_probability /mscale/= /mrestr/=. -rewrite setTI lebesgue_measure_itv/= lte_fin. -by rewrite -subr_gt0 ab0 -EFinD -EFinM mulVf// gt_eqF// subr_gt0. -Qed. - -HB.instance Definition _ := @Measure_isProbability.Build _ _ R - uniform_probability uniform_probability_setT. - -End uniform_probability. - -Section integral_uniform. -Context {R : realType}. - -Let integral_uniform_indic (a b : R) (ab0 : (0 < b - a)%R) E : - measurable E -> let m := uniform_probability ab0 in - \int[m]_x (\1_E x)%:E = - (b - a)^-1%:E * \int[lebesgue_measure]_(x in `[a, b]) (\1_E x)%:E. -Proof. -move=> mE m. -by rewrite !integral_indic//= /uniform_probability/= /mscale/= /mrestr setIT. -Qed. - -Import HBNNSimple. - -Let integral_uniform_nnsfun (f : {nnsfun measurableTypeR R >-> R}) - (a b : R) (ab0 : (0 < b - a)%R) : let m := uniform_probability ab0 in - \int[m]_x (f x)%:E = - (b - a)^-1%:E * \int[lebesgue_measure]_(x in `[a, b]) (f x)%:E. -Proof. -move=> m. -under [LHS]eq_integral do rewrite fimfunE -fsumEFin//. -rewrite [LHS]ge0_integral_fsum//; last 2 first. - - by move=> r; exact/measurable_EFinP/measurableT_comp. - - by move=> n x _; rewrite EFinM nnfun_muleindic_ge0. -rewrite -[RHS]ge0_integralZl//; last 3 first. - - exact/measurable_EFinP/measurable_funTS. - - by move=> x _; rewrite lee_fin. - - by rewrite lee_fin invr_ge0// ltW. -under [RHS]eq_integral. - move=> x xD; rewrite fimfunE -fsumEFin// ge0_mule_fsumr; last first. - by move=> r; rewrite EFinM nnfun_muleindic_ge0. - over. -rewrite [RHS]ge0_integral_fsum//; last 2 first. - - by move=> r; apply/measurable_EFinP; do 2 apply/measurableT_comp => //. - - move=> n x _; rewrite EFinM mule_ge0//; last by rewrite nnfun_muleindic_ge0. - by rewrite lee_fin invr_ge0// ltW. -apply: eq_fsbigr => r _; rewrite ge0_integralZl//. -- by rewrite !integralZl_indic_nnsfun//= integral_uniform_indic// muleCA. -- exact/measurable_EFinP/measurableT_comp. -- by move=> t _; rewrite nnfun_muleindic_ge0. -- by rewrite lee_fin invr_ge0// ltW. -Qed. - -End integral_uniform. - Section mscore. Context d (T : measurableType d) (R : realType). Variable f : T -> R. From 58e74782a8acda2f5a417a6e1d84ee0a4e093896 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 24 Oct 2024 12:02:23 +0900 Subject: [PATCH 24/48] generalize case_sum --- theories/lang_syntax.v | 6 +- theories/lang_syntax_examples.v | 2 +- theories/lang_syntax_table_game.v | 8 +- theories/lang_syntax_toy.v | 2 +- theories/lang_syntax_util.v | 2 +- theories/prob_lang.v | 178 ++++++++++++++++-------------- 6 files changed, 104 insertions(+), 94 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 9ec3cdce72..ef3857bfad 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -3,7 +3,7 @@ From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. From mathcomp Require Import interval_inference. From mathcomp Require Import unstable mathcomp_extra boolp classical_sets. -From mathcomp Require Import functions cardinality fsbigop interval_inference. +From mathcomp Require Import functions cardinality fsbigop. From mathcomp Require Import reals ereal topology normedtype sequences. From mathcomp Require Import esum measure lebesgue_measure numfun derive realfun. From mathcomp Require Import lebesgue_integral probability ftc kernel charge. @@ -147,7 +147,7 @@ move/left_right_continuousP. apply/not_andP; left. move/(@cvgrPdist_le _ R^o). apply/existsNP. -exists (2%:R^-1). +exists (2%:R^-1)%R. rewrite not_implyE; split; first by rewrite invr_gt0. move=> [e /= e0]. move/(_ (-(e / 2))%R). @@ -303,7 +303,7 @@ move=> x y x01 y01. by rewrite le_eqVlt => /predU1P[->//|/dF] => /(_ x01 y01)/ltW. Qed. -Lemma derive1_onem {R: realType} : (fun x0 : R => (1 - x0)%R : R^o)^`() = (cst (-1)%R). +Lemma derive1_onem {R : realType} : (fun x0 : R => (1 - x0)%R : R^o)^`() = (cst (-1)%R). Proof. apply/funext => x. by rewrite derive1E deriveB// derive_id derive_cst sub0r. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index 5fee74717b..fd9af1750d 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -1,4 +1,4 @@ -Require Import String. +From Coq Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. From mathcomp Require Import interval_inference. diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v index fb9c1f4f37..aa22d40357 100644 --- a/theories/lang_syntax_table_game.v +++ b/theories/lang_syntax_table_game.v @@ -1,13 +1,13 @@ Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. -From mathcomp Require Import unstable mathcomp_extra boolp classical_sets functions. -From mathcomp Require Import cardinality fsbigop interval_inference. +From mathcomp Require Import ring lra. +From mathcomp Require Import unstable mathcomp_extra boolp classical_sets. +From mathcomp Require Import functions cardinality fsbigop interval_inference. From mathcomp Require Import reals ereal topology normedtype sequences. From mathcomp Require Import esum measure charge lebesgue_measure numfun. From mathcomp Require Import lebesgue_integral probability kernel prob_lang. From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. -From mathcomp Require Import ring lra. (**md**************************************************************************) (* # Eddy's table game example *) @@ -101,7 +101,7 @@ rewrite ler_norml (@le_trans _ _ 0%R)//=. by rewrite onem_ge0. Unshelve. all: by end_near. Qed. -Lemma bounded_cst_01 (x : R) : [bounded x : R^o | _ in `[0%R, 1%R]%classic : set R]. +Lemma bounded_cst_01 (x : R^o) : [bounded x | _ in `[0%R, 1%R]%classic : set R]. Proof. exists `|x|%R; split. by rewrite num_real. diff --git a/theories/lang_syntax_toy.v b/theories/lang_syntax_toy.v index e96777393f..d861a0d2e4 100644 --- a/theories/lang_syntax_toy.v +++ b/theories/lang_syntax_toy.v @@ -1,4 +1,4 @@ -Require Import String Classical. +From Coq Require Import String Classical. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg interval_inference. From mathcomp Require Import mathcomp_extra boolp. diff --git a/theories/lang_syntax_util.v b/theories/lang_syntax_util.v index 91096193f0..2e5b82542e 100644 --- a/theories/lang_syntax_util.v +++ b/theories/lang_syntax_util.v @@ -1,4 +1,4 @@ -Require Import String. +From Coq Require Import String. From HB Require Import structures. Require Import Classical_Prop. (* NB: to compile with Coq 8.17 *) From mathcomp Require Import all_ssreflect. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 8a1365a74d..7f8bd30bde 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -1,13 +1,12 @@ (* mathcomp analysis (c) 2022 Inria and AIST. License: CeCILL-C. *) From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. -From mathcomp Require Import rat archimedean. +From mathcomp Require Import rat archimedean ring lra. From mathcomp Require Import unstable mathcomp_extra boolp classical_sets. From mathcomp Require Import functions cardinality fsbigop interval_inference. From mathcomp Require Import reals ereal topology normedtype sequences. From mathcomp Require Import esum measure lebesgue_measure numfun. From mathcomp Require Import lebesgue_integral probability exp kernel charge. -From mathcomp Require Import ring lra. (**md**************************************************************************) (* # Semantics of a probabilistic programming language using s-finite kernels *) @@ -1005,7 +1004,7 @@ Definition case_nat_ m (xn : X * nat) : {measure set Y -> \bar R} := Let measurable_fun_case_nat_ m U : measurable U -> measurable_fun setT (case_nat_ m ^~ U). Proof. -move=> mU; rewrite /case_nat_ (_ : (fun _ => _) = +move=> mU/=; rewrite /case_nat_ (_ : (fun _ => _) = (fun x => if x.2 == m then k x.1 U else mzero U)) /=; last first. by apply/funext => -[t b]/=; case: ifPn. apply: (@measurable_fun_if_pair_nat _ _ _ _ (k ^~ U) (fun=> mzero U)) => //. @@ -1111,14 +1110,41 @@ Notation "p .-sum.-measurable" := ((p.-sum).-measurable : set (set (_ + _))) : classical_set_scope. +(* TODO: move *) +Lemma bigmaxEFin {R : realDomainType} {I : Type} (s : seq I) (P : I -> bool) + (F : I -> R) (x : R) : + (\big[Order.max/x%:E]_(i <- s | P i) (F i)%:E)%R = + (\big[Order.max/x]_(i <- s | P i) F i)%:E. +Proof. by rewrite (big_morph _ EFin_max erefl). Qed. + +#[short(type="measurableCountType")] +HB.structure Definition MeasurableCountable d := + {T of Measurable d T & Countable T }. + +#[short(type="measurableFinType")] +HB.structure Definition MeasurableFinite d := + {T of Measurable d T & Finite T }. + +Definition measurableTypeUnit := unit. + +HB.instance Definition _ := Pointed.on measurableTypeUnit. +HB.instance Definition _ := Finite.on measurableTypeUnit. +HB.instance Definition _ := Measurable.on measurableTypeUnit. +HB.instance Definition _ := MeasurableFinite.on measurableTypeUnit. + +Definition measurableTypeBool := bool. + +HB.instance Definition _ := Pointed.on measurableTypeBool. +HB.instance Definition _ := Finite.on measurableTypeBool. +HB.instance Definition _ := Measurable.on measurableTypeBool. + Module CASE_SUM. Section case_sum'. Section kcase_sum'. Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Let A : measurableType _ := unit. -Let B : measurableType _ := bool. +Context dA (A : measurableCountType dA) dB (B : measurableCountType dB). Variables (k1 : A -> R.-sfker X ~> Y) (k2 : B -> R.-sfker X ~> Y). Definition case_sum' : X * (A + B) -> {measure set Y -> \bar R} := @@ -1134,66 +1160,45 @@ rewrite /= => mU. apply: (measurability _ (ErealGenInftyO.measurableE R)) => //. move=> /= _ [_ [x ->] <-]; apply: measurableI => //. rewrite /case_sum'/= (_ : _ @^-1` _ = - ([set x1 | k1 tt x1 U < x%:E] `*` inl @` [set tt]) `|` - ([set x1 | k2 false x1 U < x%:E] `*` inr @` [set false]) `|` - ([set x1 | k2 true x1 U < x%:E] `*` inr @` [set true])); last first. + (\bigcup_a ([set x1 | k1 a x1 U < x%:E] `*` inl @` [set a])) `|` + (\bigcup_b ([set x1 | k2 b x1 U < x%:E] `*` inr @` [set b]))); last first. apply/seteqP; split. - - move=> z /=; rewrite in_itv/=; move: z => [z [[]|[|]]]//= ?. - + by do 2 left; split => //; exists tt. - + by right; split => //; exists true. - + by left; right; split => //; exists false. - - move=> z /=; rewrite in_itv/=; move: z => [z [[]|[|]]]//=. - - move=> [[[]//|]|]. - + by move=> [_ []]. - + by move=> [_ []]. - - move=> [[|]|[]//]. - + by move=> [_ []]. - + by move=> [_ [] [|]]. - - move=> [[|[]//]|]. - + by move=> [_ []]. - + by move=> [_ [] [|]]. -pose h1 := [set xub : X * (unit + bool) | k1 tt xub.1 U < x%:E]. -have mh1 : measurable h1. - rewrite -[X in measurable X]setTI; apply: emeasurable_fun_infty_o => //=. - have H : measurable_fun [set: X] (fun x => k1 tt x U) by exact/measurable_kernel. - move=> _ /= C mC; rewrite setTI. - have := H measurableT _ mC; rewrite setTI => {}H. - rewrite [X in measurable X](_ : _ = ((fun x => k1 tt x U) @^-1` C) `*` setT)//. - exact: measurableX. - by apply/seteqP; split => [z//=| z/= []]. -set h2 := [set xub : X * (unit + bool)| k2 false xub.1 U < x%:E]. -have mh2 : measurable h2. - rewrite -[X in measurable X]setTI. - apply: emeasurable_fun_infty_o => //=. - have H : measurable_fun [set: X] (fun x => k2 false x U) by exact/measurable_kernel. - move=> _ /= C mC; rewrite setTI. - have := H measurableT _ mC; rewrite setTI => {}H. - rewrite [X in measurable X](_ : _ = ((fun x => k2 false x U) @^-1` C) `*` setT)//. - exact: measurableX. - by apply/seteqP; split => [z //=|z/= []]. -set h3 := [set xub : X * (unit + bool)| k2 true xub.1 U < x%:E]. -have mh3 : measurable h3. - rewrite -[X in measurable X]setTI. - apply: emeasurable_fun_infty_o => //=. - have H : measurable_fun [set: X] (fun x => k2 (true) x U) by exact/measurable_kernel. - move=> _ /= C mC; rewrite setTI. - have := H measurableT _ mC; rewrite setTI => {}H. - rewrite [X in measurable X](_ : _ = ((fun x => k2 (true) x U) @^-1` C) `*` setT)//. - exact: measurableX. - by apply/seteqP; split=> [z//=|z/= []]. + - move=> z/=; rewrite in_itv/=. + move: z => [z [a|b]]/= ?. + + by left; exists a => //; split => //=; exists a. + + by right; exists b => //; split => //=; exists b. + - move=> z/=; rewrite in_itv/=. + move: z => [z [a|b]]/= [|]. + + by case => a' _ /= [] /[swap] [] [_ ->] [->]. + + by case => b' _ /= [] b'x [_ ->]. + + by case => b' _ /= [] b'x [_ ->]. + + by case => b' _ /= [] /[swap] [] [_ ->] [->]. apply: measurableU. -- apply: measurableU. - + apply: measurableX => //. - rewrite [X in measurable X](_ : _ = ysection h1 (inl tt))//. - * by apply: measurable_ysection. - * by apply/seteqP; split => z /=; rewrite /ysection /= inE. - + apply: measurableX => //. - rewrite [X in measurable X](_ : _ = ysection h2 (inr false))//. - * by apply: measurable_ysection. - * by apply/seteqP; split => z /=; rewrite /ysection /= inE. -- apply: measurableX => //. - rewrite [X in measurable X](_ : _ = ysection h3 (inr true))//. - + by apply: measurable_ysection. +- pose h1 a := [set xub : X * (A + B) | k1 a xub.1 U < x%:E]. + apply: countable_bigcupT_measurable; first exact: countableP. + move=> a; apply: measurableX => //. + rewrite [X in measurable X](_ : _ = ysection (h1 a) (inl a)). + + apply: measurable_ysection. + rewrite -[X in measurable X]setTI. + apply: emeasurable_fun_infty_o => //= => _ /= C mC; rewrite setTI. + have : measurable_fun setT (fun x => k1 a x U) by exact/measurable_kernel. + move=> /(_ measurableT _ mC); rewrite setTI => H. + rewrite [X in measurable X](_ : _ = ((fun x => k1 a x U) @^-1` C) `*` setT)//. + exact: measurableX. + by apply/seteqP; split => [z//=| z/= []]. + + by apply/seteqP; split => z /=; rewrite /ysection /= inE. +- pose h2 a := [set xub : X * (A + B)| k2 a xub.1 U < x%:E]. + apply: countable_bigcupT_measurable; first exact: countableP. + move=> b; apply: measurableX => //. + rewrite [X in measurable X](_ : _ = ysection (h2 b) (inr b))//. + + apply: measurable_ysection. + rewrite -[X in measurable X]setTI. + apply: emeasurable_fun_infty_o => //= _ /= C mC; rewrite setTI. + have : measurable_fun setT (fun x => k2 b x U) by exact/measurable_kernel. + move=> /(_ measurableT _ mC); rewrite setTI => H. + rewrite [X in measurable X](_ : _ = ((fun x => k2 b x U) @^-1` C) `*` setT)//. + exact: measurableX. + by apply/seteqP; split => [z //=|z/= []]. + by apply/seteqP; split => z /=; rewrite /ysection /= inE. Qed. @@ -1204,8 +1209,7 @@ End kcase_sum'. Section sfkcase_sum'. Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Let A : measurableType _ := unit. -Let B : measurableType _ := bool. +Context dA (A : measurableFinType dA) dB (B : measurableFinType dB). Variables (k1 : A -> R.-sfker X ~> Y) (k2 : B-> R.-sfker X ~> Y). Let sfinite_case_sum' : exists2 k_ : (R.-ker _ ~> _)^nat, @@ -1223,17 +1227,24 @@ set Hf2 := fun ab : B => svalP (cid (sfinite_kernel (k2 ab))). rewrite /= in Hf2. exists (fun n => case_sum' (f1 ^~ n) (f2 ^~ n)). move=> n /=. - have [rtt Hrtt] := measure_uub (f1 tt n). - have [rfalse Hrfalse] := measure_uub (f2 false n). - have [rtrue Hrtrue] := measure_uub (f2 true n). - exists (maxr rtt (maxr rfalse rtrue)) => //= -[x [[]|[|]]] /=. - by rewrite 2!EFin_max lt_max Hrtt. - by rewrite 2!EFin_max 2!lt_max Hrtrue 2!orbT. - by rewrite 2!EFin_max 2!lt_max Hrfalse orbT. -move=> [x [[]|[|]]] U mU/=-. -by rewrite (Hf1 tt x _ mU). -by rewrite (Hf2 true x _ mU). -by rewrite (Hf2 false x _ mU). + pose f1' a := sval (cid (measure_uub (f1 a n))). + pose f2' b := sval (cid (measure_uub (f2 b n))). + red. + exists (maxr (\big[Order.max/0%R]_a f1' a) (\big[Order.max/0%R]_b (f2' b)))%R. + move=> /= [x [a|b]]. + - have [bnd Hbnd] := measure_uub (f1 a n). + rewrite EFin_max lt_max; apply/orP; left. + rewrite /case_sum' -bigmaxEFin. + apply: lt_le_trans; last exact: le_bigmax_cond. + by rewrite /f1'; case: cid => /=. + - have [bnd Hbnd] := measure_uub (f2 b n). + rewrite EFin_max lt_max; apply/orP; right. + rewrite /case_sum' -bigmaxEFin. + apply: lt_le_trans; last exact: le_bigmax_cond. + by rewrite /f2'; case: cid => /=. +move=> [x [a|b]] U mU/=-. +- by rewrite (Hf1 a x _ mU). +- by rewrite (Hf2 b x _ mU). Qed. #[export] @@ -1245,10 +1256,8 @@ End case_sum'. Section case_sum. Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). -Let A : measurableType _ := unit. -Let B : measurableType _ := bool. +Context dA (A : measurableFinType dA) dB (B : measurableFinType dB). -(* case analysis on the datatype unit + bool *) Definition case_sum (f : R.-sfker X ~> (A + B)%type) (k1 : A -> R.-sfker X ~> Y) (k2 : B -> R.-sfker X ~> Y) : R.-sfker X ~> Y := f \; case_sum' k1 k2. @@ -1286,8 +1295,7 @@ End kcounting. (* formalization of the iterate construct of Staton ESOP 2017, Sect. 4.2 *) Section iterate. Context d {G : measurableType d} {R : realType}. -Let A : measurableType _ := unit. -Let B : measurableType _ := bool. +Context dA (A : measurableFinType dA) dB (B : measurableFinType dB). Import CASE_SUM. @@ -1306,7 +1314,7 @@ Fixpoint iterate_ n : R.-sfker G ~> B := (fun v => fail) end. -(* formalization of iterate (A = unit, B = bool) +(* formalization of iterate Gamma, x : A |-p t : A + B Gamma |-d u : A ----------------------------------------------- Gamma |-p iterate t from x = u : B *) @@ -1350,7 +1358,7 @@ Section von_neumann_trick. Context d {T : measurableType d} {R : realType}. Definition minltt {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} := - @measurable_cst _ _ T1 _ setT (@inl unit T2 tt). + @measurable_cst _ _ T1 _ setT (@inl _ T2 tt). Definition finrb d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) : T1 * bool -> T2 + bool := fun t1b => inr t1b.2. @@ -1369,8 +1377,10 @@ Definition trick : R.-sfker (T * unit) ~> (unit + bool)%type := (ite macc3of4 (letin (ret macc1of4) (ret minrb)) (ret minltt)))). +(* trick : R.-sfker T * unit ~> (unit + bool)%type *) -Definition von_neumann_trick : R.-sfker T ~> bool := iterate trick ktt. +Definition von_neumann_trick : R.-sfker T ~> measurableTypeBool := + (@iterate _ T R _ measurableTypeUnit _ measurableTypeBool trick _ ktt). End von_neumann_trick. From afcaad2827092e23d1e73cf393609f98e47e9c34 Mon Sep 17 00:00:00 2001 From: Cyril Cohen Date: Tue, 24 Dec 2024 03:29:16 +0100 Subject: [PATCH 25/48] Von Neumann Trick Proof complete --- theories/lang_syntax.v | 6 +- theories/lang_syntax_examples.v | 4 - theories/lang_syntax_table_game.v | 2 + theories/prob_lang.v | 158 +++++++++++++++++++++++++++++- 4 files changed, 158 insertions(+), 12 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index ef3857bfad..3b2b376ff7 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -641,7 +641,7 @@ Qed. Local Open Scope ereal_scope. -Lemma integral_exprn {R : realType} (n : nat) : +Lemma integral_exprn {R : realType} n : fine (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (x ^+ n)%:E) = n.+1%:R^-1%R :> R. Proof. pose F (x : R) : R^o := (n.+1%:R^-1 * x ^+ n.+1)%R. @@ -695,7 +695,7 @@ rewrite derive1E deriveX_idfun derive1E deriveB//. by rewrite -derive1E derive1_cst derive_id sub0r mulrN1 [in RHS]mulNr. Qed. -Lemma integral_onemXn {R : realType} (n : nat) : +Lemma integral_onemXn {R : realType} n : fine (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (`1-x ^+ n)%:E) = n.+1%:R^-1%R :> R. Proof. rewrite (@continuous_FTC2 _ _ (fun x : R => ((1 - x) ^+ n.+1 / - n.+1%:R))%R)//=. @@ -718,7 +718,7 @@ apply: continuous_in_subspaceT => x _. exact: continuous_XMonemX. Qed. -Lemma Rintegral_onemXn {R : realType} (n : nat) : +Lemma Rintegral_onemXn {R : realType} n : (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (`1-x ^+ n))%R = n.+1%:R^-1%R :> R. Proof. rewrite /Rintegral. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index fd9af1750d..de25619715 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -64,10 +64,6 @@ Local Open Scope ring_scope. Local Open Scope ereal_scope. Local Open Scope string_scope. -Local Open Scope lang_scope. - -Local Close Scope lang_scope. - (* simple tests to check bidirectional hints *) Module bidi_tests. Section bidi_tests. diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v index aa22d40357..27da2a4ef1 100644 --- a/theories/lang_syntax_table_game.v +++ b/theories/lang_syntax_table_game.v @@ -33,6 +33,8 @@ Local Open Scope ring_scope. Local Open Scope string_scope. Local Open Scope ereal_scope. +Local Open Scope string_scope. + Lemma letin'_sample_uniform {R : realType} d d' (T : measurableType d) (T' : measurableType d') (a b : R) (ab : (a < b)%R) (u : R.-sfker [the measurableType _ of (_ * T)%type] ~> T') x y : diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 7f8bd30bde..bf1099f7fe 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -1186,7 +1186,7 @@ apply: measurableU. rewrite [X in measurable X](_ : _ = ((fun x => k1 a x U) @^-1` C) `*` setT)//. exact: measurableX. by apply/seteqP; split => [z//=| z/= []]. - + by apply/seteqP; split => z /=; rewrite /ysection /= inE. + + by rewrite ysectionE. - pose h2 a := [set xub : X * (A + B)| k2 a xub.1 U < x%:E]. apply: countable_bigcupT_measurable; first exact: countableP. move=> b; apply: measurableX => //. @@ -1199,7 +1199,7 @@ apply: measurableU. rewrite [X in measurable X](_ : _ = ((fun x => k2 b x U) @^-1` C) `*` setT)//. exact: measurableX. by apply/seteqP; split => [z //=|z/= []]. - + by apply/seteqP; split => z /=; rewrite /ysection /= inE. + + by rewrite ysectionE. Qed. #[export] @@ -1369,6 +1369,8 @@ Proof. exact: measurableT_comp. Qed. (* biased coin *) Variable (D : pprobability bool R). +Let unit := measurableTypeUnit. +Let bool := measurableTypeBool. Definition trick : R.-sfker (T * unit) ~> (unit + bool)%type := letin (sample_cst D) @@ -1377,10 +1379,156 @@ Definition trick : R.-sfker (T * unit) ~> (unit + bool)%type := (ite macc3of4 (letin (ret macc1of4) (ret minrb)) (ret minltt)))). -(* trick : R.-sfker T * unit ~> (unit + bool)%type *) -Definition von_neumann_trick : R.-sfker T ~> measurableTypeBool := - (@iterate _ T R _ measurableTypeUnit _ measurableTypeBool trick _ ktt). +Definition von_neumann_trick_kernel : _ -> _ := + (@iterate _ _ R _ unit _ bool trick _ ktt). +Definition von_neumann_trick x : _ -> _ := von_neumann_trick_kernel x. + +HB.instance Definition _ := SFiniteKernel.on von_neumann_trick_kernel. +HB.instance Definition _ x := Measure.on (von_neumann_trick x). + +Section von_neumann_trick_proof. + +Hypothesis D_nontrivial : 0 < D [set true] < 1. + +Let p : R := fine (D [set true]). +Let q : R := p * (1 - p). +Let r : R := p ^+ 2 + (1 - p) ^+ 2. + +Let Dtrue : D [set true] = p%:E. +Proof. by rewrite fineK//= fin_num_measure. Qed. + +Let Dfalse : D [set false] = 1 - (p%:E). +Proof. +rewrite -Dtrue; have <- : D [set: bool] = 1 := probability_setT _. +rewrite setT_bool measureU//=; last first. + by rewrite disjoints_subset => -[]//. +by rewrite addeAC subee ?fin_num_measure ?add0e. +Qed. + +Let p_ge0 : (0 <= p)%R. +Proof. by rewrite fine_ge0 ?measure_ge0. Qed. + +Let p'_ge0 : (1 - p >= 0)%R. +Proof. by rewrite -lee_fin EFinB -Dfalse measure_ge0. Qed. + +Let q_ge0 : (q >= 0)%R. Proof. by rewrite mulr_ge0 ?p_ge0 ?p'_ge0. Qed. + +Let r_ge0 : (r >= 0)%R. +Proof. by rewrite addr_ge0// exprn_ge0 ?p_ge0 ?p'_ge0. Qed. + +Let intDE f : (forall x, 0 <= f x) -> + \int[D]_x (f x) = (1 - (p%:E)) * f false + p%:E * f true. +Proof. +move=> f_ge0. +pose M := measure_add (mscale (NngNum p'_ge0) \d_false) + (mscale (NngNum p_ge0) \d_true). +rewrite (eq_measure_integral M) => [|A Ameas _]. + rewrite ge0_integral_measure_add//. + by rewrite !ge0_integral_mscale ?integral_dirac ?diracT ?mul1e. +rewrite /M/= /msum/= !big_ord_recl/= big_ord0 addr0 /mscale/=. +by case: (set_bool A) => /eqP->/=; + rewrite ?measure0 ?probability_setT ?Dtrue ?Dfalse//= + ?diracE/= ?in_set1 ?inE/= ?(mule0, mul0e, adde0, add0e, mule1)// + -EFinD addrNK. +Qed. + +Let p_gt0 : (0 < p)%R. +Proof. by rewrite -lte_fin -Dtrue; case/andP : D_nontrivial. Qed. + +Let p_lt1 : (p < 1)%R. +Proof. by rewrite -lte_fin -Dtrue; case/andP : D_nontrivial. Qed. + +Let p'_gt0 : (0 < 1 - p)%R. Proof. by rewrite subr_gt0. Qed. + +Let r_lt1 : (r < 1)%R. +Proof. +rewrite /r -subr_gt0 [ltRHS](_ : _ = 2 * p * (1 - p))%R; last by ring. +by rewrite !mulr_gt0. +Qed. + +Lemma iterate_trick gamma n (b : bool) : + @iterate_ _ _ R _ unit _ bool trick _ ktt n gamma [set b] + = (geometric q r n)%:E. +Proof. +elim: n => [|n IHn] //=; rewrite /kcomp; rewrite integral_kcomp//=; + rewrite integral_dirac//= ?diracT ?mul1e integral_kcomp//=; + under eq_integral do [rewrite integral_kcomp//=; + under eq_integral do rewrite integral_kcomp//= + ?integral_dirac//= ?diracT ?mul1e]; + do ![rewrite intDE//=; last by + [move=> b'; do !rewrite integral_ge0//= => *]]; + rewrite !iteE//= ?integral_kcomp//= + ?integral_dirac//= ?diracT ?mul1e/=; + rewrite /acc1of4/= /kcomp ?ge0_integral_mscale //= + ?diracE/= ?in_set1 ?inE/= + ?(mule0, mul0e, adde0, add0e, mule1, normr0)//=; + rewrite -?(EFinB, EFinN, EFinM, EFinD) ?lee_fin ?expr0 ?mulr1//. + by case: b => //=; rewrite ?(mulr0, mulr1, add0r, addr0)// ?[(_ * p)%R]mulrC. +rewrite IHn -?(EFinB, EFinN, EFinM, EFinD) ?lee_fin ?expr0 ?mulr1//. +rewrite [geometric _ _ _]lock !mulrA -mulrDl addrC. +by rewrite /geometric/= -/r -lock mulrCA exprS. +Qed. + +Lemma iterate_trickT gamma n : + @iterate_ _ _ R _ unit _ bool trick _ ktt n gamma [set: bool] = + (2 * geometric q r n)%:E. +Proof. +rewrite setT_bool measureU//=; last first. + by rewrite disjoints_subset => -[]//. +by rewrite !iterate_trick -EFinD -mulr2n mulr_natl. +Qed. + +Lemma von_neumann_trick_prob_kernel gamma b : + von_neumann_trick_kernel gamma [set b] = 2^-1%:E. +Proof. +rewrite /= /kcomp/= /case_nat_/= /mseries. +under eq_integral => n _. + under (@congr_lim _ _ _ \o @eq_fun _ _ _ _) => k. + under eq_bigr do rewrite fun_if/= (fun_if (@^~ _))/mzero eq_sym. + rewrite -big_mkcond/= big_nat1_eq iterate_trick. + over. +over. +rewrite /= (eq_integral (EFin \o geometric q r))//=; last first. + by move=> k _; apply/lim_near_cst => //; near do rewrite ifT//. +have cvgg: series (geometric q r) x @[x --> \oo] --> (q / (1 - r))%R. + by apply/cvg_geometric_series; rewrite ger0_norm ?r_lt1. +have limgg := cvg_lim (@Rhausdorff R) cvgg. +have sumgE : \big[+%R/0%R]_(0 <= k *; field. + rewrite [X in X != _](_ : _ = 2 * (p * (1 - p)))%R; last by ring. + by rewrite mulf_eq0 ?pnatr_eq0/= mulf_neq0// gt_eqF ?p_gt0 ?p'_gt0//. +suff summableg : summable setT (EFin \o geometric q r) + by rewrite integral_count. +rewrite /summable /=. +rewrite (_ : [set: nat] = [set x | x \in predT]); last first. + by apply/eq_set => x; rewrite inE trueE. +rewrite -(@nneseries_esum _ _ predT)//=. +under eq_eseriesr do rewrite ger0_norm// ?geometric_ge0//. +by rewrite sumgE ltey. +Unshelve. all: end_near. Qed. + +Lemma von_neumann_trick_prob_kernelT gamma : + von_neumann_trick gamma [set: bool] = 1. +Proof. +rewrite setT_bool measureU//=; last first. + by rewrite disjoints_subset => -[]. +rewrite !von_neumann_trick_prob_kernel -EFinD. +by have := splitr (1 : R); rewrite mul1r => <-. +Qed. + +HB.instance Definition _ gamma := Measure.on (von_neumann_trick gamma). +HB.instance Definition _ gamma := Measure_isProbability.Build _ _ _ + (von_neumann_trick gamma) (von_neumann_trick_prob_kernelT gamma). +HB.instance Definition _ := Kernel_isProbability.Build _ _ _ _ _ + von_neumann_trick_kernel von_neumann_trick_prob_kernelT. + +Theorem von_neumann_trickP gamma : von_neumann_trick gamma =1 bernoulli 2^-1. +Proof. by apply: eq_bernoulli; rewrite von_neumann_trick_prob_kernel. Qed. + +End von_neumann_trick_proof. End von_neumann_trick. From 1ec0d97dde0f30e5678582958e450edf721a603c Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Fri, 27 Dec 2024 14:10:03 +0900 Subject: [PATCH 26/48] probability_fin can be avoided --- theories/prob_lang.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/prob_lang.v b/theories/prob_lang.v index bf1099f7fe..8365f4d8f5 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -314,7 +314,7 @@ rewrite (_ : (fun x => _) = (fun x => x * apply/funext => x; case: ifPn => ix; first by rewrite indicE/= mem_set ?mule1. by rewrite indicE/= memNset ?mule0// /= in_itv/=; exact/negP. apply: emeasurable_funM => //=; apply/measurable_EFinP. -by rewrite (_ : \1__ = mindic R (emeasurable_itv `[(i%:R)%:E, (i.+1%:R)%:E[)). +by rewrite (_ : \1__ = mindic R (emeasurable_itv `[i%:R%:E, i.+1%:R%:E[)). Qed. Definition mk i t := [the measure _ _ of k mf i t]. @@ -1511,7 +1511,7 @@ by rewrite sumgE ltey. Unshelve. all: end_near. Qed. Lemma von_neumann_trick_prob_kernelT gamma : - von_neumann_trick gamma [set: bool] = 1. + von_neumann_trick gamma [set: bool] = 1. Proof. rewrite setT_bool measureU//=; last first. by rewrite disjoints_subset => -[]. From 1d199c2706b7fd0f3b2cb83e25a9ff618ceecdc7 Mon Sep 17 00:00:00 2001 From: Cyril Cohen Date: Sun, 29 Dec 2024 21:33:45 +0100 Subject: [PATCH 27/48] Better Von Neumann trick + generalization --- theories/lang_syntax_examples_wip.v | 696 ---------------------------- theories/prob_lang.v | 227 ++++++--- 2 files changed, 161 insertions(+), 762 deletions(-) delete mode 100644 theories/lang_syntax_examples_wip.v diff --git a/theories/lang_syntax_examples_wip.v b/theories/lang_syntax_examples_wip.v deleted file mode 100644 index eabf4b4fab..0000000000 --- a/theories/lang_syntax_examples_wip.v +++ /dev/null @@ -1,696 +0,0 @@ -Require Import String. -From HB Require Import structures. -From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. -From mathcomp Require Import ring lra. -From mathcomp Require Import mathcomp_extra boolp classical_sets functions. -From mathcomp Require Import cardinality fsbigop interval_inference reals. -From mathcomp Require Import ereal topology normedtype sequences esum. -From mathcomp Require Import measure lebesgue_measure numfun lebesgue_integral. -From mathcomp Require Import kernel prob_lang lang_syntax_util lang_syntax. - -(******************************************************************************) -(* Casino example *) -(* some steps *) - -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -Import Order.TTheory GRing.Theory Num.Def Num.Theory. -Import numFieldTopology.Exports. - -Local Open Scope classical_set_scope. -Local Open Scope ring_scope. -(* Local Open Scope ereal_scope. *) -Local Open Scope string_scope. - -Section beta_example. -Open Scope ring_scope. -Open Scope lang_scope. -Context {R : realType}. - -Lemma beta_bern610 : - @execP R [::] _ - [let "p" := - Sample {exp_beta 6 4} in Sample {exp_bernoulli_trunc [#{"p"}]}] = - execP [Sample {exp_bernoulli_trunc [{6 / 10}:R]}]. -Proof. -rewrite execP_letin !execP_sample !execD_beta_nat !execD_bernoulli_trunc/=. -rewrite execD_real exp_var'E !(execD_var_erefl "p")/=. -apply: eq_sfkernel=> x U. -rewrite letin'E/=. -transitivity ((\int[beta_nat 6 4]_(y in `[0%R, 1%R]) bernoulli_trunc y U)%E : \bar R). - rewrite [in RHS]integral_mkcond /=. - apply: eq_integral => y _. - rewrite patchE. - case: ifPn => //. - (* Unset Printing Notations. *) - simpl in *. - rewrite mem_setE /= in_itv /= negb_and -!ltNge => /orP[y0|y1]; - rewrite /bernoulli_trunc/=. - case: sumbool_ler. - move=> a. - by rewrite ltNge a in y0. - rewrite /prob_lang.bernoulli0 /bernoulli => _. - rewrite [LHS]measure_addE/= /mscale/=. - (* match default value *) - admit. - admit. - rewrite integral_beta_nat. -under eq_integral => y y01. - rewrite bernoulli_truncE. - rewrite muleC muleDr// !muleA muleC [X in _ + X]muleC. - over. - admit. -rewrite /= bernoulli_truncE. -rewrite integralD//. -congr (_ + _)%E. -rewrite integralZl//. -rewrite muleC. -congr (_ * _)%E. -(* rewrite /= setTI. *) -rewrite /beta_nat_pdf. -rewrite /ubeta_nat_pdf/ubeta_nat_pdf'/=. -transitivity (\int[@lebesgue_measure R]_(x0 in `[0%R, 1%R]) - ((ubeta_nat_pdf 7 4 x0 / beta_nat_norm 6 4)%:E))%E. - apply: eq_integral => p _. - rewrite muleC -EFinM. - rewrite -[X in X * _]divr1 mulf_div. - congr (_ / _)%:E; last by rewrite mul1r. - by rewrite /ubeta_nat_pdf/= /ubeta_nat_pdf' mulrA -exprS. -under eq_integral do rewrite mulrC EFinM. -rewrite integralZl//=. -rewrite -beta_nat_normE /beta_nat_norm/= !factE. -rewrite -!EFinM/=. -congr _%:E; lra. - admit. - admit. -rewrite integralZl//. -rewrite muleC. -congr (_ * _)%E. -rewrite /beta_nat_pdf. -rewrite /ubeta_nat_pdf/ubeta_nat_pdf'/=. -transitivity (\int[@lebesgue_measure R]_(x0 in `[0%R, 1%R]) - ((ubeta_nat_pdf 6 5 x0 / beta_nat_norm 6 4)%:E))%E. - apply: eq_integral => p _. - by rewrite mulrC -EFinM -2!mulrA -exprSr mulrC. -under eq_integral do rewrite mulrC EFinM. -rewrite integralZl//=. -rewrite -beta_nat_normE /beta_nat_norm/= !factE. -rewrite -!EFinM/onem/=. -congr _%:E; lra. - admit. - admit. - admit. - admit. - lra. -Admitted. - -End beta_example. - -Section casino_example. -Open Scope ring_scope. -Open Scope lang_scope. -Context (R : realType). -Lemma a01 : 0 < 1 - 0 :> R. Proof. by []. Qed. - -(* guard test *) -Definition test_guard : @exp R _ [::] _ := [ - let "p" := Sample {exp_bernoulli (1 / 3)%:nng (p1S 2)} in - let "_" := if #{"p"} then return TT else Score {0}:R in - return #{"p"} -]. - -Lemma exec_guard t U : execP test_guard t U = ((1 / 3)%:E * \d_true U)%E. -Proof. -rewrite /test_guard 2!execP_letin execP_sample execD_bernoulli execP_if/=. -rewrite !execP_return !exp_var'E !(execD_var_erefl "p") execD_unit execP_score execD_real/=. -rewrite letin'E ge0_integral_measure_sum//. -rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//= !integral_dirac//. -rewrite !letin'E !iteE/= integral_dirac// ge0_integral_mscale//=. -by rewrite normr0 mul0e !mule0 !adde0 !diracT !mul1e. -Qed. - -Lemma binomial_over1 n p U : - 0 <= p <= 1 -> - (\int[binomial_probability_trunc n p]_y0 \d_(0 < y0)%N U = - bernoulli_trunc (1 - `1-p ^+ n) U :> \bar R)%E. -Proof. -move=> /andP[p0 p1]. -rewrite bernoulli_truncE; last first. - apply/andP; split. - apply/onemX_ge0; rewrite /onem; lra. - apply/onem_le1/exprn_ge0; rewrite /onem; lra. -rewrite (@integral_binomial_probabilty_trunc _ n p _ _ (fun y => \d_(1 <= y)%N U))//; last first. -rewrite !big_ord_recl/=. -rewrite /bump. -under eq_bigr => i _. - rewrite /=. - have -> : (0 < 1 + i)%N => //. - over. -rewrite addeC -ge0_sume_distrl. - congr (_ + _)%E; congr (_ * _)%E. - have -> : (\sum_(i < n) (p ^+ (1 + i) * `1-p ^+ (n - (1 + i)) *+ 'C(n, 1 + i))%:E)%E = - (\sum_(i < n.+1) (p ^+ i * `1-p ^+ (n - i) *+ 'C(n, i))%:E - (`1-p ^+ n)%:E)%E. - rewrite big_ord_recl/= expr0 subn0 mul1r bin0 mulr1n addeC addeA. - have <- : 0%E = ((- `1-p ^+ n)%:E + (`1-p ^+ n)%:E)%E. - rewrite EFinN. - congr _%:E. - lra. - by rewrite add0e. - congr _%E. - rewrite sumEFin. - rewrite !EFinB EFin_expe. - congr (_ - _)%E. - under eq_bigr do rewrite mulrC. - rewrite -(@exprDn_comm _ `1-p p n); last first. - by rewrite /GRing.comm/onem; lra. - rewrite /onem addrC. - have -> : p + (1 - p) = 1 by lra. - by rewrite expr1n. - rewrite subn0 expr0 bin0 mulr1n. - rewrite /onem. - congr _%:E. - set pn := (1-p) ^+ n. - lra. -move=> i _. -apply/mulrn_wge0/mulr_ge0; apply/exprn_ge0. -exact: p0. -apply/onem_ge0/p1. -Qed. - -Lemma binomial_le1 n p U : - 0 <= p <= 1 -> - (\int[binomial_probability_trunc n p]_y0 \d_(0 < y0)%N U = - \int[bernoulli_trunc (1 - `1-p ^+ n)]_y0 \d_y0 U :> \bar R)%E. -Proof. -move=> /andP[p0 p1]. -rewrite (@integral_bernoulli_trunc _ _ (fun x => \d_x U))//; last first. - apply/andP; split. - apply: onemX_ge0; rewrite /onem; lra. - apply/onem_le1/exprn_ge0; rewrite /onem; lra. -rewrite (@integral_binomial_probabilty_trunc _ n p _ _ (fun y => \d_(1 <= y)%N U))//; last first. -rewrite !big_ord_recl/=. -rewrite /bump. -under eq_bigr => i _. - rewrite /=. - have -> : (0 < 1 + i)%N => //. - over. -rewrite addeC -ge0_sume_distrl. - congr (_ + _)%E; congr (_ * _)%E. - have -> : (\sum_(i < n) (p ^+ (1 + i) * `1-p ^+ (n - (1 + i)) *+ 'C(n, 1 + i))%:E)%E = - (\sum_(i < n.+1) (p ^+ i * `1-p ^+ (n - i) *+ 'C(n, i))%:E - (`1-p ^+ n)%:E)%E. - rewrite big_ord_recl/= expr0 subn0 mul1r bin0 mulr1n addeC addeA. - have <- : 0%E = ((- `1-p ^+ n)%:E + (`1-p ^+ n)%:E)%E. - rewrite EFinN. - congr _%:E. - lra. - by rewrite add0e. - congr _%E. - rewrite sumEFin. - rewrite !EFinB EFin_expe. - congr (_ - _)%E. - under eq_bigr do rewrite mulrC. - rewrite -(@exprDn_comm _ `1-p p n); last first. - by rewrite /GRing.comm/onem; lra. - rewrite /onem addrC. - have -> : p + (1 - p) = 1 by lra. - by rewrite expr1n. - rewrite subn0 expr0 bin0 mulr1n. - rewrite /onem. - congr _%:E. - set pn := (1-p) ^+ n. - lra. -move=> i _. -apply/mulrn_wge0/mulr_ge0; apply/exprn_ge0. -exact: p0. -apply/onem_ge0/p1. -Qed. - -Lemma __ : uniform_probability a01 `[0, (1 / 2)] = (1 / 2)%:E. -Proof. -rewrite /uniform_probability /mscale/= /mrestr. -Abort. - -Lemma letin'_sample_uniform d d' (T : measurableType d) - (T' : measurableType d') (a b : R) (ab0 : (0 < b - a)%R) - (u : R.-sfker [the measurableType _ of (_ * T)%type] ~> T') x y : - measurable y -> - letin' (sample_cst (uniform_probability ab0)) u x y = - ((b - a)^-1%:E * \int[lebesgue_measure]_(x0 in `[a, b]) u (x0, x) y)%E. -Proof. -move=> my; rewrite letin'E/=. -rewrite integral_uniform//= => _ /= Y mY /=. -have /= := measurable_kernel u _ my measurableT _ mY. -move/measurable_ysection => /(_ x) /=. -set A := (X in measurable X). -set B := (X in _ -> measurable X). -suff : A = B by move=> ->. -rewrite {}/A {}/B !setTI /ysection/= (*TODO: lemma?*) /preimage/=. -by apply/seteqP; split => [z|z] /=; rewrite inE/=. -Qed. - -Lemma execP_letin_uniform g t str (s0 s1 : exp P ((str, Real) :: g) t) : - (forall (p : R) x U, 0 <= p <= 1 -> - execP s0 (p, x) U = execP s1 (p, x) U) -> - forall x U, measurable U -> - execP [let str := Sample {@exp_uniform _ g 0 1 a01} in {s0}] x U = - execP [let str := Sample {@exp_uniform _ g 0 1 a01} in {s1}] x U. -Proof. -move=> s01 x U mU. -rewrite !execP_letin execP_sample execD_uniform/=. -rewrite !letin'_sample_uniform//. -congr (_ * _)%E. -apply: eq_integral => p p01. -apply: s01. -by rewrite inE in p01. -Qed. - -(* Lemma measurable_mtyp (t : typ) (U : set (@mtyp R t)) : measurable U. -Proof. -induction t => //. *) - -Lemma congr_letinl g t1 t2 str (e1 e2 : @exp _ _ g t1) -(e : @exp _ _ (_ :: g) t2) x U : - (forall y V, execP e1 y V = execP e2 y V) -> - measurable U -> - @execP R g t2 [let str := e1 in e] x U = - @execP R g t2 [let str := e2 in e] x U. -Proof. -move=> He mU. -apply eq_sfkernel in He. -by rewrite !execP_letin He. -Qed. - -Lemma congr_letin g t1 t2 str (e : @exp _ _ _ t1) (e1 e2 : @exp _ _ (_ :: g) t2) x U : - (forall y V, execP e1 (y, x) V = execP e2 (y, x) V) -> - @execP R g t2 [let str := e in e1] x U = @execP R g t2 [let str := e in e2] x U. -Proof. -move=> He. -rewrite !execP_letin !letin'E. -apply: eq_integral => ? _. -apply: He. -Qed. - -Lemma congr_normalize g t (e1 e2 : @exp R _ g t) : - (forall x U, execP e1 x U = execP e2 x U) -> - execD [Normalize e1] = execD [Normalize e2]. -Proof. -move=> He. -apply: eq_execD. -rewrite !execD_normalize_pt /=. -f_equal. -apply: eq_kernel => y V. -apply: He. -Qed. - -Definition uniform_syntax : @exp R _ [::] _ := - [let "p" := Sample {exp_uniform 0 1 a01} in - return #{"p"}]. - -Lemma exec_uniform_syntax t U : measurable U -> - execP uniform_syntax t U = uniform_probability a01 U. -Proof. -move=> mU. -rewrite /uniform_syntax execP_letin execP_sample execP_return !execD_uniform. -rewrite exp_var'E (execD_var_erefl "p")/=. -rewrite letin'E /=. -rewrite integral_uniform//=; last exact: measurable_fun_dirac. -rewrite subr0 invr1 mul1e. -rewrite {1}/uniform_probability. -rewrite /mscale/= subr0 invr1 mul1e. -by rewrite integral_indic. -Qed. - -Definition binomial_le : @exp R _ [::] Bool := - [let "a2" := Sample {exp_binomial 3 (1 / 2)%:nng (p1S 1)} in - return {1}:N <= #{"a2"}]. - -Lemma exec_binomial_le t U : - execP binomial_le t U = ((7 / 8)%:E * \d_true U + - (1 / 8)%:E * \d_false U)%E. -Proof. -rewrite /binomial_le execP_letin execP_sample execP_return execD_rel execD_nat. -rewrite exp_var'E (execD_var_erefl "a2") execD_binomial. -rewrite letin'E//= /binomial_probability ge0_integral_measure_sum//=. -rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//=. -rewrite !integral_dirac// /bump. -rewrite !binS/= !bin0 bin1 bin2 bin_small// addn0. -rewrite addeC adde0. -congr (_ + _)%:E. - rewrite !indicT !(mul0n,add0n,lt0n,mul1r)/=. - rewrite -!mulrDl; congr (_ * _). - rewrite /onem. - lra. -rewrite !expr0 ltnn indicT/= !(mul1r,mul1e) /onem. -lra. -Qed. - -Definition binomial_guard : @exp R _ [::] Nat := - [let "a1" := Sample {exp_binomial 3 (1 / 2)%:nng (p1S 1)} in - let "_" := if #{"a1"} == {1}:N then return TT else Score {0}:R in - return #{"a1"}]. - -Lemma exec_binomial_guard t U : - execP binomial_guard t U = ((3 / 8)%:E * \d_1%N U(* + - (1 / 8)%:E * \d_0%N U*))%E. -Proof. -rewrite /binomial_guard !execP_letin execP_sample execP_return execP_if. -rewrite !exp_var'E execD_rel !(execD_var_erefl "a1") execP_return. -rewrite execD_unit execD_binomial execD_nat execP_score execD_real. -rewrite !letin'E//= /binomial_probability ge0_integral_measure_sum//=. -rewrite !big_ord_recl big_ord0 !ge0_integral_mscale//=. -rewrite !integral_dirac//. -rewrite /bump/=. -rewrite !binS/= !bin0 bin1 bin2 bin_small//. -rewrite !diracT !addn0 !expr0 !subn0 !mulr1n !mul1r !expr1 !mul1e. -rewrite !letin'E//= !iteE/= !diracE/=. -rewrite !ge0_integral_mscale//=. -rewrite !integral_dirac// !diracT//. -rewrite !(normr0,mul0e,mule0,add0e,add0n,mul1e,adde0). -rewrite /onem. -congr (_%:E * _)%E. -lra. -Qed. - -Lemma exec_beta_a1 U : - @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in - Sample {exp_bernoulli_trunc [#{"p"}]}] tt U = - @execP R [::] _ [Sample {exp_bernoulli_trunc [{3 / 5}:R]}] tt U. -Proof. -rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. -rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. -transitivity (beta_nat_bern 6 4 1 0 U : \bar R). - rewrite /beta_nat_bern !letin'E/= /ubeta_nat_pdf/= /onem. - apply: eq_integral => x _. - rewrite /=. - do 2 f_equal. - by rewrite /ubeta_nat_pdf' expr1 expr0 mulr1. -rewrite beta_nat_bern_bern// !bernoulli_truncE; last 2 first. - by lra. - apply/andP; split. - by apply/Baa'bb'Bab_ge0. - by apply/Baa'bb'Bab_le1. -congr (_ * _ + _ * _)%:E. - rewrite /Baa'bb'Bab /beta_nat_norm/=. - rewrite !factS/= fact0. - by field. -rewrite /onem; rewrite /Baa'bb'Bab /beta_nat_norm/=; -rewrite !factS/= fact0. -by field. -Qed. - -Definition casino0 : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_uniform 0 1 a01} in - let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - let "a2" := Sample {exp_binomial_trunc 3 [#{"p"}]} in - return {1}:N <= #{"a2"}]. - -Definition casino1 : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_uniform 0 1 a01} in - let "a1" := Sample {exp_binomial_trunc 8 [#{"p"}]} in - let "_" := if #{"a1"} == {5}:N then return TT else Score {0}:R in - Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. - -Definition casino2 : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_uniform 0 1 a01} in - let "_" := - Score {[{56}:R * #{"p"} ^+ {5%nat} * {[{1}:R - #{"p"}]} ^+ {3%nat}]} in - Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. - -Definition casino2' : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_beta 1 1} in - let "_" := Score - {[{56}:R * #{"p"} ^+ {5%nat} * {[{1}:R - #{"p"}]} ^+ {3%nat}]} in - Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. - -Definition casino3 : @exp R _ [::] _ := - [Normalize - let "_" := Score {1 / 9}:R in - let "p" := Sample {exp_beta 6 4} in - Sample {exp_bernoulli_trunc [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%nat}]}]. - -Definition casino4 : @exp R _ [::] _ := - [Normalize - let "_" := Score {1 / 9}:R in - Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. - -Definition casino5 : @exp R _ [::] _ := - [Normalize Sample {exp_bernoulli_trunc [{10 / 11}:R]}]. - -Lemma casino01 : - execD casino0 = execD casino1. -Proof. -rewrite /casino0 /casino1. -apply: eq_execD. -f_equal. -apply: congr_normalize => y V. -apply: execP_letin_uniform => //. -move=> p x U r01. -apply: congr_letin => y0 V0. -apply: congr_letin => y1 V1. -rewrite !execP_letin !execP_sample !execD_binomial_trunc /=. -rewrite !execP_return !execD_bernoulli_trunc/=. -rewrite !execD_rel (@execD_bin _ _ binop_minus) execD_pow. -rewrite (@execD_bin _ _ binop_minus) !execD_real/= !execD_nat. -rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a2")/=. -rewrite !letin'E/=. -move: r01 => /andP[r0 r1]. -by apply/binomial_over1/andP. -Qed. - -Lemma casino12 : - execD casino1 = execD casino2. -Proof. -apply: eq_execD. -f_equal. -apply: congr_normalize => y V. -apply: execP_letin_uniform => //. -move=> p x U /andP[p0 p1]. -rewrite !execP_letin !execP_sample execP_if execD_rel/=. -rewrite !execP_score !(@execD_bin _ _ binop_mult). -rewrite !execD_bernoulli_trunc/= !(@execD_bin _ _ binop_minus) !execD_pow. -rewrite !(@execD_bin _ _ binop_minus)/=. -rewrite !execD_real !execD_nat/= execP_return execD_unit. -rewrite !execD_binomial_trunc/=. -rewrite !exp_var'E !(execD_var_erefl "p") !(execD_var_erefl "a1")/=. -rewrite !letin'E/=. -rewrite integral_binomial_probabilty_trunc//=. -rewrite (bigD1 (inord 5))//=. - rewrite big1; last first. - move=> [[|[|[|[|[|[|[|[|[|//]]]]]]]]]]//= Hi Hi5; rewrite letin'E iteE; - rewrite ?ge0_integral_mscale//= ?normr0 ?mul0e ?mule0 ?add0e//. - suff: false by []. - move/negbTE: Hi5 => <-. - by apply/eqP/val_inj => /=; rewrite inordK. -rewrite letin'E iteE ge0_integral_mscale//= inordK//= adde0 /onem. -congr (_ * _)%E. -rewrite ger0_norm. - by rewrite -mulrA mulr_natl. -apply/mulr_ge0. - exact/mulr_ge0/exprn_ge0. -apply/exprn_ge0. -by rewrite subr_ge0. -Qed. - -Lemma casino22' : - execD casino2 = execD casino2'. -Proof. -apply: eq_execD. -f_equal. -apply: congr_normalize => //= x U. -apply: congr_letinl => //= y V. -rewrite !execP_sample execD_uniform execD_beta_nat/=. -rewrite beta11_uniform//. -rewrite /=. -apply: sub_sigma_algebra. -rewrite //. -admit. -Admitted. - -Lemma casino23 : - execD casino2' = execD casino3. -Proof. -apply: eq_execD. -f_equal. -apply: congr_normalize => x U. -rewrite !execP_letin !execP_sample !execP_score !execD_beta_nat. -rewrite !execD_bernoulli_trunc/= !(@execD_bin _ _ binop_mult). -rewrite !execD_pow !(@execD_bin _ _ binop_minus) !execD_real/=. -rewrite !execD_pow !(@execD_bin _ _ binop_minus) !execD_real/=. -rewrite !exp_var'E !(execD_var_erefl "p")/=. -rewrite !letin'E/= ![in RHS]ge0_integral_mscale//=. -rewrite /=. -under eq_integral => y _. - rewrite letin'E/= /mscale/=. - over. -rewrite /=. -(* set f := letin' _ _. *) -(* oops... -transitivity (\int[beta_nat 1 1]_(y in `[0%R, 1%R]) f (y, x) U)%E. - rewrite [in RHS]integral_mkcond /=. - apply: eq_integral => y _. - rewrite patchE. - case: ifPn => //. - simpl in *. - rewrite mem_setE /= in_itv /= negb_and -!ltNge => /orP[y0|y1]; - rewrite /bernoulli_trunc/=. - case: sumbool_ler. - move=> a. - by rewrite ltNge a in y0. - rewrite /prob_lang.bernoulli0 /bernoulli => _. - rewrite [LHS]measure_addE/= /mscale/=. - (* match default value *) - admit. - admit. - rewrite integral_beta_nat. - admit. -rewrite (integral_beta_nat 6 4). - rewrite ger0_norm// integral_dirac// diracT mul1e letin'E/=. - transitivity (((1 / 9)%:E * \int[beta_nat 6 4]_(y in `[0%R, 1%R]) - bernoulli_trunc (1 - (1 - y) ^+ 3) U)%E : \bar R); last first. - admit. -rewrite (integral_beta_nat 6 4)//=. -rewrite -integralZl//=. - apply: eq_integral => y y01. - rewrite /f letin'E /= !ge0_integral_mscale//= integral_dirac// diracT mul1e. - rewrite -muleAC muleC -[in RHS]muleCA. - congr (_ * _)%E. - rewrite ger0_norm. - rewrite /beta_nat_pdf ubeta_nat_pdf11/= muleC -!EFinM. - rewrite !div1r. - rewrite /beta_nat_norm/= /ubeta_nat_pdf/ubeta_nat_pdf' factE/=/onem. - congr _%:E; lra. - rewrite inE/= in_itv/= in y01. - move: y01 => /andP[y0 y1]. - apply/mulr_ge0/exprn_ge0 => //. - apply/mulr_ge0/exprn_ge0 => //. - lra. -admit. -admit. -admit. -apply: (@measurableT_comp _ _ _ _ _ _ (fun x0 => f x0 U)). - apply: measurable_kernel. - done. -apply: measurable_pair2. -admit.*) -Admitted. - -Lemma casino34' U : - @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in - Sample {exp_bernoulli_trunc [{[{1}:R - #{"p"}]} ^+ {3%nat}]}] tt U = - @execP R [::] _ [Sample {exp_bernoulli_trunc [{1 / 11}:R]}] tt U. -Proof. -rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. -rewrite execD_pow/= (@execD_bin _ _ binop_minus) !execD_real/=. -rewrite exp_var'E (execD_var_erefl "p")/=. -(* TODO: generalize *) -rewrite letin'E/=. -have := (@beta_nat_bern_bern R 6 4 0 3 U). -rewrite /beta_nat_bern /ubeta_nat_pdf/ubeta_nat_pdf'/=. -under eq_integral do rewrite expr0 mul1r. -move=> ->//. -rewrite /Baa'bb'Bab addn0 /beta_nat_norm/= !factE/= !factE. -by congr (bernoulli_trunc _ _); field. -Qed. - -Lemma bern_onem (f : _ -> R) U p : - (forall x, 0 <= f x <= 1) -> - (\int[beta_nat 6 4]_y bernoulli_trunc (f y) U = - p%:E * \d_true U + `1-p%:E * \d_false U)%E -> - (\int[beta_nat 6 4]_y bernoulli_trunc (1 - f y) U - = `1-p%:E * \d_true U + p%:E * \d_false U)%E. -Proof. -move=> f01. -under eq_integral => x _. - rewrite bernoulli_truncE. - over. -done. -move=> h1. -rewrite /= in h1. -rewrite /bernoulli_trunc. -(* /beta_nat/mscale/= /beta_nat_norm/= /ubeta_nat/ubeta_nat_pdf/=. *) -Admitted. - -Lemma casino34 : - execD casino3 = execD casino4. -Proof. -apply: eq_execD. -f_equal. -apply: congr_normalize => y V. -apply: congr_letin => x U. -rewrite execP_letin !execP_sample execD_beta_nat !execD_bernoulli_trunc/=. -rewrite (@execD_bin _ _ binop_minus) execD_pow/= (@execD_bin _ _ binop_minus). -rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. -transitivity (\int[beta_nat 6 4]_y bernoulli_trunc (1 - (1 - y) ^+ 3) U : \bar R)%E. - by rewrite /beta_nat_bern !letin'E/= /onem. -rewrite bernoulli_truncE; last by lra. -have -> := (@bern_onem (fun x => (1 - x) ^+ 3) U (1 / 11) _). - congr (_ * _ + _ * _)%E; congr _%:E; rewrite /onem; lra. - admit. -transitivity (beta_nat_bern 6 4 0 3 U : \bar R). - rewrite /beta_nat_bern /ubeta_nat_pdf/= /onem. - apply: eq_integral => y0 _. - do 2 f_equal. - rewrite /ubeta_nat_pdf'/=. - rewrite expr0. - by rewrite mul1r. -rewrite beta_nat_bern_bern//= bernoulli_truncE; last first. - apply/andP; split. - apply/Baa'bb'Bab_ge0. - apply/Baa'bb'Bab_le1. -congr (_ * _ + _ * _)%:E; rewrite /onem. - rewrite /Baa'bb'Bab /beta_nat_norm/=. - by rewrite !factS/= fact0; field. -rewrite /Baa'bb'Bab /beta_nat_norm/=. -by rewrite !factS/= fact0; field. -Admitted. - -Lemma norm_score_bern g p1 p2 (p10 : p1 != 0) (p1_ge0 : 0 <= p1) -(p201 : 0 <= p2 <= 1) : - @execD R g _ - [Normalize let "_" := Score {p1}:R in - Sample {exp_bernoulli_trunc [{p2}:R]}] = - execD [Normalize Sample {exp_bernoulli_trunc [{p2}:R]}]. -Proof. -apply: eq_execD. -rewrite !execD_normalize_pt/= !execP_letin !execP_score. -rewrite !execP_sample !execD_bernoulli_trunc !execD_real/=. -apply: funext=> x. -apply: eq_probability=> /= y. -rewrite /normalize_pt !normalizeE/=. -rewrite !bernoulli_truncE; last lra; last lra. -rewrite !diracT !mule1 /onem -EFinD addrCA subrr addr0. -rewrite !letin'E. -under eq_integral. - move=> x0 _ /=. - rewrite !bernoulli_truncE; last lra. - rewrite !diracT !mule1 /onem -EFinD addrCA subrr addr0. - over. -rewrite !ge0_integral_mscale//= ger0_norm//. -rewrite integral_dirac// diracT !mule1. -rewrite !ifF; last first. - rewrite eqe. - apply/negbTE/negP => /orP[/eqP|//]. - by rewrite /onem; lra. - rewrite eqe. - apply/negbTE/negP => /orP[/eqP|//]. - by rewrite /onem; lra. -rewrite !bernoulli_truncE; last lra. -rewrite integral_dirac//= diracT !diracE. -by rewrite muleC muleA -EFinM mulVf// invr1 /onem !(mul1r, mule1). -Qed. - -Lemma casino45 : - execD casino4 = execD casino5. -Proof. -rewrite norm_score_bern//. -lra. -Qed. - -End casino_example. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 8365f4d8f5..21771c0dba 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -1322,6 +1322,130 @@ Definition iterate : R.-sfker G ~> B := case_nat (kcounting R) iterate_. End iterate. +Section iterate_unit. + +Let unit := measurableTypeUnit. +Let bool := measurableTypeBool. +Context d {G : measurableType d} {R : realType}. +Context dB (B : measurableFinType dB). + +Section iterate_elim. +Variables (t : R.-sfker (G * unit) ~> (unit + B)%type) + (u : G -> unit) (mu : measurable_fun setT u). +Variables (r : R) (tlE : forall gamma, t (gamma, tt) [set inl tt] = r%:E). + +Variables (gamma : G) (X : set B) (q : R). +Hypothesis trE : t (gamma, tt) [set inr x | x in X] = q%:E. + +Let q_ge0 : (0 <= q)%R. Proof. by rewrite -lee_fin -trE measure_ge0. Qed. +Let r_ge0 : (0 <= r)%R. +Proof. by rewrite -lee_fin -(tlE gamma) measure_ge0. Qed. + +Lemma iterate_E n : iterate_ t mu n gamma X = (geometric q r n)%:E. +Proof. +elim: n => [|n IHn] //=; + rewrite /kcomp; rewrite integral_kcomp//=; + rewrite /= integral_dirac//= ?diracT ?mul1e ?expr0 ?exprS ?mulr1. + rewrite (eq_integral (EFin \o \1_[set inr x | x in X]))//=; last first. + move=> [a' _|b _]//=; last first. + by rewrite diracE indicE/= (mem_image inr_inj). + rewrite /kcomp/= indicE /= ge0_integral_mscale//= normr0 mul0e. + by rewrite [_ \in _](introF idP)// inE /= => -[]. + by rewrite ?unitE integral_indic//= setIT. +pose g : unit + B -> R^o := (geometric q r n \o* \1_[set inl tt])%R. +rewrite (eq_integral (EFin \o g))//=; last first. + move=> [[] _|b _]//=. + by rewrite /g/= indicE//= in_set1 inE eqxx mul1r. + rewrite /kcomp/= ge0_integral_mscale//= normr0 mul0e. + by rewrite /g /= indicE//= in_set1 inE mul0r. +rewrite /g /=; under eq_integral do rewrite EFinM. +rewrite integralZr//=; last first. + apply/integrableP; split=> //. + under eq_integral => x. + rewrite gee0_abs//=. + over. + by rewrite /= integral_indic// setIT [u gamma]unitE tlE ltey. +by rewrite integral_indic//= [u gamma]unitE setIT tlE -EFinM mulrCA. +Qed. + +Hypothesis r_lt1 : (r < 1)%R. + +Lemma iterateE : iterate t mu gamma X = (q / (1 - r))%:E. +Proof. +rewrite /= /kcomp/= /case_nat_/= /mseries. +under eq_integral => n _. + under (@congr_lim _ _ _ \o @eq_fun _ _ _ _) => k. + under eq_bigr do rewrite fun_if/= (fun_if (@^~ _))/mzero eq_sym. + rewrite -big_mkcond/= big_nat1_eq. + over. +over. +rewrite /= (eq_integral (EFin \o geometric q r))//=; last first. + move=> k _; apply/lim_near_cst => //; rewrite iterate_E ?r_ge0 ?r_lt1//. + by near do rewrite ifT//. +have cvgg: series (geometric q r) x @[x --> \oo] --> (q / (1 - r))%R. + by apply/cvg_geometric_series; rewrite ger0_norm ?r_lt1//. +have limgg := cvg_lim (@Rhausdorff R) cvgg. +have sumgE : \big[+%R/0%R]_(0 <= k x; rewrite inE trueE. +rewrite -(@nneseries_esum _ _ predT)//=. +under eq_eseriesr do rewrite ger0_norm// ?geometric_ge0//. +by rewrite sumgE ltey. +Unshelve. all: end_near. Qed. + +End iterate_elim. + +Import CASE_SUM. + +Variables (t : R.-pker (G * unit) ~> (unit + B)%type) + (u : G -> unit) (mu : measurable_fun setT u). +Variables (r : R) (r_lt1 : (r < 1)%R). +Hypothesis (tlE : forall gamma, t (gamma, tt) [set inl tt] = r%:E). + +Let trE gamma X : t (gamma, tt) [set inr x | x in X] \in fin_num. +Proof. +apply/fin_numPlt; rewrite (@lt_le_trans _ _ 0)//=. +rewrite (@le_lt_trans _ _ 1)//= ?ltey//. +rewrite -( @prob_kernel _ _ _ _ _ t (gamma, tt) ). +by apply/le_measure => //=; rewrite inE//=. +Qed. + +Lemma iterate_normalize p : + iterate t mu = knormalize (case_sum (letin (ret mu) t) + (fun u' => fail) + (fun v => ret (measurable_cst v))) p. +Proof. +apply/eq_sfkernel => gamma U. +have /EFin_fin_numP[q trE'] := trE gamma U. +rewrite (iterateE mu tlE trE')//; symmetry. +rewrite /= /mnormalize/= (fun_if (@^~ U))/=. +set m := kcomp _ _ _. +have mE V : m V = t (gamma, tt) [set inr x | x in V]. + rewrite /m/= /kcomp/= integral_kcomp//= integral_dirac//= diracT mul1e. + rewrite (eq_integral (EFin \o \1_[set inr x | x in V])). + by rewrite integral_indic ?setIT ?unitE. + move=> [x|x] xV /=; rewrite indicE. + rewrite ?inl_in_set_inr /kcomp/=. + by rewrite ge0_integral_mscale//= ?normr0 mul0e. + by rewrite inr_in_set_inr// indicE. +rewrite !mE trE'. +suff -> : t (gamma, tt) (range inr) = 1 - t (gamma, tt) [set inl tt]. + by rewrite tlE -EFinB/= orbF eqe subr_eq0 eq_sym lt_eqF. +rewrite -( @prob_kernel _ _ _ _ _ t (gamma, tt) ). +have -> : [set: unit + B] = [set inl tt] `|` (range inr). + symmetry; apply/eq_set => -[[]|b]//=; apply/propT; first by left. + by right; exists b. +rewrite measureU//=; first by rewrite addeAC subee ?add0e// ?tlE//. +by apply/eq_set => -[[]|b]//=; apply/propF; case=> []// _ []. +Qed. + +End iterate_unit. + (* an s-finite kernel to test that two expressions are different *) Section lift_neq. Context {R : realType} d (G : measurableType d). @@ -1380,17 +1504,15 @@ Definition trick : R.-sfker (T * unit) ~> (unit + bool)%type := (letin (ret macc1of4) (ret minrb)) (ret minltt)))). -Definition von_neumann_trick_kernel : _ -> _ := +Definition kvon_neumann_trick : _ -> _ := (@iterate _ _ R _ unit _ bool trick _ ktt). -Definition von_neumann_trick x : _ -> _ := von_neumann_trick_kernel x. +Definition von_neumann_trick x : _ -> _ := kvon_neumann_trick x. -HB.instance Definition _ := SFiniteKernel.on von_neumann_trick_kernel. +HB.instance Definition _ := SFiniteKernel.on kvon_neumann_trick. HB.instance Definition _ x := Measure.on (von_neumann_trick x). Section von_neumann_trick_proof. -Hypothesis D_nontrivial : 0 < D [set true] < 1. - Let p : R := fine (D [set true]). Let q : R := p * (1 - p). Let r : R := p ^+ 2 + (1 - p) ^+ 2. @@ -1433,6 +1555,30 @@ by case: (set_bool A) => /eqP->/=; -EFinD addrNK. Qed. +Lemma trick_bool gamma b : trick gamma [set inr b] = q%:E. +Proof. +rewrite /trick/= /kcomp /= intDE//= /kcomp/= !intDE//= /kcomp. +rewrite !integral_dirac ?diracT//= ?mul1e. +rewrite !iteE//= ?diracE/= /kcomp/=. +rewrite !integral_dirac ?diracT ?diracE ?mul1e//=. +rewrite /finrb ?inl_in_set_inr//= /acc1of4/= ?inr_in_set_inr !in_set1 !inE. +rewrite /q -?(EFinB, EFinN, EFinM, EFinD); congr (_)%:E. +by case: b => //=; ring. +Qed. + +Lemma trick_unit gamma : trick gamma [set inl tt] = r%:E. +Proof. +rewrite /trick/= /kcomp /= intDE//= /kcomp/= !intDE//= /kcomp. +rewrite !integral_dirac ?diracT//= ?mul1e. +rewrite !iteE//= ?diracE/= /kcomp/=. +rewrite !integral_dirac ?diracT ?diracE ?mul1e//=. +rewrite /finrb ?inl_in_set_inr//= /acc1of4/= ?inr_in_set_inr. +rewrite !in_set1 !inE/=. +by rewrite /r -?(EFinB, EFinN, EFinM, EFinD); congr (_)%:E; ring. +Qed. + +Hypothesis D_nontrivial : 0 < D [set true] < 1. + Let p_gt0 : (0 < p)%R. Proof. by rewrite -lte_fin -Dtrue; case/andP : D_nontrivial. Qed. @@ -1447,74 +1593,23 @@ rewrite /r -subr_gt0 [ltRHS](_ : _ = 2 * p * (1 - p))%R; last by ring. by rewrite !mulr_gt0. Qed. -Lemma iterate_trick gamma n (b : bool) : - @iterate_ _ _ R _ unit _ bool trick _ ktt n gamma [set b] - = (geometric q r n)%:E. -Proof. -elim: n => [|n IHn] //=; rewrite /kcomp; rewrite integral_kcomp//=; - rewrite integral_dirac//= ?diracT ?mul1e integral_kcomp//=; - under eq_integral do [rewrite integral_kcomp//=; - under eq_integral do rewrite integral_kcomp//= - ?integral_dirac//= ?diracT ?mul1e]; - do ![rewrite intDE//=; last by - [move=> b'; do !rewrite integral_ge0//= => *]]; - rewrite !iteE//= ?integral_kcomp//= - ?integral_dirac//= ?diracT ?mul1e/=; - rewrite /acc1of4/= /kcomp ?ge0_integral_mscale //= - ?diracE/= ?in_set1 ?inE/= - ?(mule0, mul0e, adde0, add0e, mule1, normr0)//=; - rewrite -?(EFinB, EFinN, EFinM, EFinD) ?lee_fin ?expr0 ?mulr1//. - by case: b => //=; rewrite ?(mulr0, mulr1, add0r, addr0)// ?[(_ * p)%R]mulrC. -rewrite IHn -?(EFinB, EFinN, EFinM, EFinD) ?lee_fin ?expr0 ?mulr1//. -rewrite [geometric _ _ _]lock !mulrA -mulrDl addrC. -by rewrite /geometric/= -/r -lock mulrCA exprS. -Qed. - -Lemma iterate_trickT gamma n : - @iterate_ _ _ R _ unit _ bool trick _ ktt n gamma [set: bool] = - (2 * geometric q r n)%:E. -Proof. -rewrite setT_bool measureU//=; last first. - by rewrite disjoints_subset => -[]//. -by rewrite !iterate_trick -EFinD -mulr2n mulr_natl. -Qed. - Lemma von_neumann_trick_prob_kernel gamma b : - von_neumann_trick_kernel gamma [set b] = 2^-1%:E. + kvon_neumann_trick gamma [set b] = 2^-1%:E. Proof. -rewrite /= /kcomp/= /case_nat_/= /mseries. -under eq_integral => n _. - under (@congr_lim _ _ _ \o @eq_fun _ _ _ _) => k. - under eq_bigr do rewrite fun_if/= (fun_if (@^~ _))/mzero eq_sym. - rewrite -big_mkcond/= big_nat1_eq iterate_trick. - over. -over. -rewrite /= (eq_integral (EFin \o geometric q r))//=; last first. - by move=> k _; apply/lim_near_cst => //; near do rewrite ifT//. -have cvgg: series (geometric q r) x @[x --> \oo] --> (q / (1 - r))%R. - by apply/cvg_geometric_series; rewrite ger0_norm ?r_lt1. -have limgg := cvg_lim (@Rhausdorff R) cvgg. -have sumgE : \big[+%R/0%R]_(0 <= k *; field. rewrite [X in X != _](_ : _ = 2 * (p * (1 - p)))%R; last by ring. - by rewrite mulf_eq0 ?pnatr_eq0/= mulf_neq0// gt_eqF ?p_gt0 ?p'_gt0//. -suff summableg : summable setT (EFin \o geometric q r) - by rewrite integral_count. -rewrite /summable /=. -rewrite (_ : [set: nat] = [set x | x \in predT]); last first. - by apply/eq_set => x; rewrite inE trueE. -rewrite -(@nneseries_esum _ _ predT)//=. -under eq_eseriesr do rewrite ger0_norm// ?geometric_ge0//. -by rewrite sumgE ltey. -Unshelve. all: end_near. Qed. + by rewrite mulf_eq0 ?pnatr_eq0/= mulf_neq0// gt_eqF ?p_gt0 ?p'_gt0. +- by move=> gamma'; exact: trick_unit. +- suff -> : [set inr x | x in [set b]] = [set inr b] by exact: trick_bool. + by move=> A; apply/seteqP; split=> [a [_ ->]|_ ->]//=; exists b. +Qed. Lemma von_neumann_trick_prob_kernelT gamma : von_neumann_trick gamma [set: bool] = 1. Proof. -rewrite setT_bool measureU//=; last first. - by rewrite disjoints_subset => -[]. +rewrite setT_bool measureU//=; last by rewrite disjoints_subset => -[]. rewrite !von_neumann_trick_prob_kernel -EFinD. by have := splitr (1 : R); rewrite mul1r => <-. Qed. @@ -1523,7 +1618,7 @@ HB.instance Definition _ gamma := Measure.on (von_neumann_trick gamma). HB.instance Definition _ gamma := Measure_isProbability.Build _ _ _ (von_neumann_trick gamma) (von_neumann_trick_prob_kernelT gamma). HB.instance Definition _ := Kernel_isProbability.Build _ _ _ _ _ - von_neumann_trick_kernel von_neumann_trick_prob_kernelT. + kvon_neumann_trick von_neumann_trick_prob_kernelT. Theorem von_neumann_trickP gamma : von_neumann_trick gamma =1 bernoulli 2^-1. Proof. by apply: eq_bernoulli; rewrite von_neumann_trick_prob_kernel. Qed. From bf0c323afc78c8ce78efd6d8cea521ff1d12e1fd Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 1 Jan 2025 17:52:33 +0900 Subject: [PATCH 28/48] doc --- theories/prob_lang.v | 485 ++++++++++++++++++--------------------- theories/prob_lang_wip.v | 10 +- 2 files changed, 228 insertions(+), 267 deletions(-) diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 21771c0dba..532a0db19f 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -14,26 +14,62 @@ From mathcomp Require Import lebesgue_integral probability exp kernel charge. (* Reference: *) (* - R. Affeldt, C. Cohen, A. Saito. Semantics of probabilistic programs *) (* using s-finite kernels in Coq. CPP 2023 *) +(* - S. Staton. Commutative Semantics for Probabilistic Programming. *) +(* ESOP 2017 *) (* *) (* ``` *) (* poisson_pdf == Poisson pdf *) (* exponential_pdf == exponential distribution pdf *) +(* measurable_sum X Y == the type X + Y, as a measurable type *) +(* *) +(* mscore f t := mscale `|f t| \d_tt *) +(* kscore f := fun=> mscore f *) +(* This is an s-finite kernel. *) +(* kite k1 k2 mf := kdirac mf \; kadd (kiteT k1) (kiteF k2). *) +(* k1 has type R.-sfker T ~> T'. *) +(* k2 has type R.-sfker T ~> T'. *) +(* mf is a proof that f : T -> bool is measurable. *) +(* KITE.kiteT k1 is k1 \o fst if f returne true *) +(* and zero otherwise. *) +(* KITE.kiteF k2 is k2 \o fst if f returne false *) +(* and zero otherwise. *) (* *) -(* sample mP == sample according to the probability P where mP is *) -(* a proof that P is a measurable function *) -(* sample_cst P == sample according to the probability P *) -(* letin l k == execute l, augment the context, and execute k *) (* ret mf == access the context with f and return the result *) -(* score mf == observe t from d, where f is the density of d and *) -(* t occurs in f *) -(* e.g., score (r e^(-r * t)) = observe t from exp(r) *) -(* normalize k P == normalize the kernel k into a probability kernel, *) +(* mf is a proof that f is measurable. *) +(* This is a probability kernel. *) +(* sample mP == sample according to the probability measure P *) +(* mP is a proof that P is a measurable function. *) +(* sample_cst P == same as sample with a constant probability measure *) +(* normalize k P == normalize the kernel k into a probability kernel *) (* P is a default probability in case normalization *) -(* is not possible *) +(* is not possible. *) +(* normalize_pt k := normalize k point *) (* ite mf k1 k2 == access the context with the boolean function f and *) (* behaves as k1 or k2 according to the result *) -(* case_nat == case analysis on the nat datatype *) +(* letin l k == execute l, augment the context, and execute k *) +(* fail := let _ := score 0 in ret point *) +(* score mf == observe t from d, where f is the density of d and *) +(* t occurs in f *) +(* e.g., score (r e^(-r * t)) = observe t from exp(r) *) +(* acc0of2, acc1of2, etc. == accessor function *) +(* case_nat t u_ == case analysis on natural numbers *) +(* t has type R.-sfker T ~> nat *) +(* u_ has type nat -> R.-sfker T ~> T' *) +(* CASE_SUM.case_sum g k1 k2 == case analysis on a sum type *) +(* g has type R.-sfker X ~> (A + B). *) +(* k1 has type A -> R.-sfker X ~> Y. *) +(* k2 has type B -> R.-sfker X ~> Y. *) +(* kcounting == the counting measure as a kernel *) +(* iterate k mu == iteration *) +(* k has type R.-sfker G * A ~> (A + B). *) +(* mu is a proof that u : G -> A is measurable. *) +(* flift_neq == an s-finite kernel to test that two expressions *) +(* are different *) +(* ``` *) +(* *) +(* Examples: Staton's bus, von Neumann's trick, etc. *) (* *) +(* ``` *) (* mkswap k == given a kernel k : (Y * X) ~> Z, *) (* returns a kernel of type (X * Y) ~> Z *) (* letin' := mkcomp \o mkswap *) @@ -75,14 +111,12 @@ subst p2. by f_equal. Qed. -Definition dep_uncurry (A : Type) (B : A -> Type) (C : Type) : - (forall a : A, B a -> C) -> {a : A & B a} -> C := - fun f p => let (a, Ba) := p in f a Ba. - +(* NB: to be PRed to probability.v *) Section poisson_pdf. Variable R : realType. Local Open Scope ring_scope. +(* density function for Poisson *) Definition poisson_pdf k r : R := if r > 0 then r ^+ k / k`!%:R^-1 * expR (- r) else 1%:R. @@ -106,12 +140,16 @@ by apply: measurable_funM => /=; [exact: measurable_funM|exact: measurableT_comp]. Qed. +Definition poisson3 := poisson_pdf 4 3%:R. (* 0.168 *) +Definition poisson10 := poisson_pdf 4 10%:R. (* 0.019 *) + End poisson_pdf. Section exponential_pdf. Variable R : realType. Local Open Scope ring_scope. +(* density function for exponential *) Definition exponential_pdf x r : R := r * expR (- r * x). Lemma exponential_pdf_gt0 x r : 0 < r -> 0 < exponential_pdf x r. @@ -145,8 +183,8 @@ Let sumU (F : (set (X + Y))^nat) : (forall i, measurable_sum (F i)) -> measurable_sum (\bigcup_i F i). Proof. by []. Qed. -HB.instance Definition _ := @isMeasurable.Build default_measure_display (X + Y)%type - measurable_sum sum0 sumC sumU. +HB.instance Definition _ := @isMeasurable.Build default_measure_display + (X + Y)%type measurable_sum sum0 sumC sumU. End measurable_sum. @@ -177,11 +215,6 @@ move=> mx my; apply: measurable_fun_ifT => //=. - exact: measurableT_comp. Qed. -(* Definition mR (R : realType) : Type := R. -HB.instance Definition _ (R : realType) := Measurable.on (mR R). -(* [the measurableType (R.-ocitv.-measurable).-sigma of - salgebraType (R.-ocitv.-measurable)]. *) *) - Module Notations. Notation munit := (unit : measurableType _). Notation mbool := (bool : measurableType _). @@ -192,6 +225,7 @@ Lemma invr_nonneg_proof (R : numDomainType) (p : {nonneg R}) : (0 <= (p%:num)^-1)%R. Proof. by rewrite invr_ge0. Qed. +(* TODO: move *) Definition invr_nonneg (R : numDomainType) (p : {nonneg R}) := NngNum (invr_nonneg_proof p). @@ -249,7 +283,7 @@ Qed. End mscore. -(* decomposition of score into finite kernels *) +(* decomposition of score into finite kernels [Section 3.2, Staton ESOP 2017] *) Module SCORE. Section score. Context d (T : measurableType d) (R : realType). @@ -389,7 +423,7 @@ HB.instance Definition _ := End kscore. -(* decomposition of ite into s-finite kernels *) +(* decomposition of ite into s-finite kernels [Section 3.2, Staton ESOP 2017] *) Module ITE. Section ite. Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). @@ -536,10 +570,6 @@ HB.instance Definition _ t := isMeasure.Build _ _ _ (mite mf t) Import ITE. -(* -Definition kite : R.-sfker T ~> T' := - kdirac mf \; kadd (kiteT u1) (kiteF u2). -*) Definition kite : R.-sfker T ~> T' := kdirac mf \; kadd (kiteT u1) (kiteF u2). @@ -631,7 +661,7 @@ Section insn3. Context d d' d3 (X : measurableType d) (Y : measurableType d') (Z : measurableType d3) (R : realType). -Definition letin (l : R.-sfker X ~> Y) (k : R.-sfker [the measurableType _ of (X * Y)%type] ~> Z) +Definition letin (l : R.-sfker X ~> Y) (k : R.-sfker (X * Y) ~> Z) : R.-sfker X ~> Z := [the R.-sfker X ~> Z of l \; k]. @@ -677,8 +707,7 @@ End letin_return. Section insn1. Context d (X : measurableType d) (R : realType). -Definition score (f : X -> R) (mf : measurable_fun setT f) - : R.-sfker X ~> _ := +Definition score (f : X -> R) (mf : measurable_fun setT f) : R.-sfker X ~> _ := [the R.-sfker X ~> _ of kscore mf]. End insn1. @@ -704,7 +733,6 @@ Definition k3 : measurable_fun _ _ := kr 3%:R. Definition k10 : measurable_fun _ _ := kr 10%:R. Definition ktt := @measurable_cst _ _ T _ setT tt. Definition kb (b : bool) := @measurable_cst _ _ T _ setT b. - Definition kn (n : nat) := @measurable_cst _ _ T _ setT n. End cst_fun. @@ -837,7 +865,7 @@ Import Notations. Context d0 d1 d2 d3 (T0 : measurableType d0) (T1 : measurableType d1) (T2 : measurableType d2) (T3 : measurableType d3) (R : realType). -Definition T01 : seq {d & measurableType d} := [:: existT _ _ T0; existT _ _ T1]. +Let T01 : seq {d & measurableType d} := [:: existT _ _ T0; existT _ _ T1]. Definition acc0of2 : T0 * T1 -> T0 := acc T01 0 \o pairAr unit tt. @@ -855,7 +883,7 @@ Proof. by apply: measurableT_comp; [exact: (measurable_acc T01 1)|exact: mpairAr]. Qed. -Definition T02 := [:: existT _ _ T0; existT _ _ T1; existT _ _ T2]. +Let T02 := [:: existT _ _ T0; existT _ _ T1; existT _ _ T2]. Definition acc1of3 : T0 * T1 * T2 -> T1 := acc T02 1 \o pairAAr. @@ -897,8 +925,7 @@ Proof. by apply: measurableT_comp; [exact: (measurable_acc T02 2)|exact: mpairAArAi]. Qed. -Definition T03 := - [:: existT _ _ T0; existT _ _ T1; existT _ d2 T2; existT _ d3 T3]. +Let T03 := [:: existT _ _ T0; existT _ _ T1; existT _ d2 T2; existT _ d3 T3]. Definition acc1of4 : T0 * T1 * T2 * T3 -> T1 := acc T03 1 \o pairAAAr. @@ -936,61 +963,6 @@ Arguments macc1of4 {d0 d1 d2 d3 _ _ _ _}. Arguments macc2of4' {d0 d1 d2 d3 _ _ _ _}. Arguments macc3of4 {d0 d1 d2 d3 _ _ _ _}. -(* sample programs *) -Section poisson. -Variable R : realType. -Local Open Scope ring_scope. - -(* density function for Poisson *) -Definition poisson k r : R := - if r > 0 then r ^+ k / k`!%:R^-1 * expR (- r) else 1%:R. - -Lemma poisson_ge0 k r : 0 <= poisson k r. -Proof. -rewrite /poisson; case: ifPn => r0//. -by rewrite mulr_ge0 ?expR_ge0// mulr_ge0// exprn_ge0 ?ltW. -Qed. - -Lemma poisson_gt0 k r : 0 < r -> 0 < poisson k.+1 r. -Proof. -move=> r0; rewrite /poisson r0 mulr_gt0 ?expR_gt0//. -by rewrite divr_gt0// ?exprn_gt0// invr_gt0 ltr0n fact_gt0. -Qed. - -Lemma measurable_poisson k : measurable_fun setT (poisson k). -Proof. -rewrite /poisson; apply: measurable_fun_if => //. - exact: measurable_fun_ltr. -by apply: measurable_funM => /=; - [exact: measurable_funM|exact: measurableT_comp]. -Qed. - -Definition poisson3 := poisson 4 3%:R. (* 0.168 *) -Definition poisson10 := poisson 4 10%:R. (* 0.019 *) - -End poisson. - -Section exponential. -Variable R : realType. -Local Open Scope ring_scope. - -(* density function for exponential *) -Definition exp_density x r : R := r * expR (- r * x). - -Lemma exp_density_gt0 x r : 0 < r -> 0 < exp_density x r. -Proof. by move=> r0; rewrite /exp_density mulr_gt0// expR_gt0. Qed. - -Lemma exp_density_ge0 x r : 0 <= r -> 0 <= exp_density x r. -Proof. by move=> r0; rewrite /exp_density mulr_ge0// expR_ge0. Qed. - -Lemma mexp_density x : measurable_fun setT (exp_density x). -Proof. -apply: measurable_funM => //=; apply: measurableT_comp => //. -exact: measurable_funM. -Qed. - -End exponential. - Module CASE_NAT. Section case_nat. Context d d' (X : measurableType d) (Y : measurableType d') (R : realType). @@ -1071,10 +1043,10 @@ Definition case_nat (t : R.-sfker T ~> nat) (u_ : (R.-sfker T ~> T')^nat) End case_nat. Definition measure_sum_display : - (measure_display * measure_display) -> measure_display. + measure_display * measure_display -> measure_display. Proof. exact. Qed. -Definition image_classes d1 d2 +Definition g_sigma_imageU d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) (T : Type) (f1 : T1 -> T) (f2 : T2 -> T) := < T1 + T2 := @inl T1 T2. Let f2 : T2 -> T1 + T2 := @inr T1 T2. -Lemma sum_salgebra_set0 : image_classes f1 f2 (set0 : set (T1 + T2)). +Lemma sum_salgebra_set0 : g_sigma_imageU f1 f2 (set0 : set (T1 + T2)). Proof. exact: sigma_algebra0. Qed. -Lemma sum_salgebra_setC A : image_classes f1 f2 A -> - image_classes f1 f2 (~` A). +Lemma sum_salgebra_setC A : g_sigma_imageU f1 f2 A -> + g_sigma_imageU f1 f2 (~` A). Proof. exact: sigma_algebraC. Qed. -Lemma sum_salgebra_bigcup (F : _^nat) : (forall i, image_classes f1 f2 (F i)) -> - image_classes f1 f2 (\bigcup_i (F i)). +Lemma sum_salgebra_bigcup (F : _^nat) : (forall i, g_sigma_imageU f1 f2 (F i)) -> + g_sigma_imageU f1 f2 (\bigcup_i (F i)). Proof. exact: sigma_algebra_bigcup. Qed. HB.instance Definition sum_salgebra_mixin := @isMeasurable.Build (measure_sum_display (d1, d2)) - (T1 + T2)%type (image_classes f1 f2) + (T1 + T2)%type (g_sigma_imageU f1 f2) sum_salgebra_set0 sum_salgebra_setC sum_salgebra_bigcup. End sum_salgebra_instance. @@ -1110,13 +1082,6 @@ Notation "p .-sum.-measurable" := ((p.-sum).-measurable : set (set (_ + _))) : classical_set_scope. -(* TODO: move *) -Lemma bigmaxEFin {R : realDomainType} {I : Type} (s : seq I) (P : I -> bool) - (F : I -> R) (x : R) : - (\big[Order.max/x%:E]_(i <- s | P i) (F i)%:E)%R = - (\big[Order.max/x]_(i <- s | P i) F i)%:E. -Proof. by rewrite (big_morph _ EFin_max erefl). Qed. - #[short(type="measurableCountType")] HB.structure Definition MeasurableCountable d := {T of Measurable d T & Countable T }. @@ -1234,14 +1199,14 @@ exists (fun n => case_sum' (f1 ^~ n) (f2 ^~ n)). move=> /= [x [a|b]]. - have [bnd Hbnd] := measure_uub (f1 a n). rewrite EFin_max lt_max; apply/orP; left. - rewrite /case_sum' -bigmaxEFin. + rewrite /case_sum' -EFin_bigmax. apply: lt_le_trans; last exact: le_bigmax_cond. by rewrite /f1'; case: cid => /=. - have [bnd Hbnd] := measure_uub (f2 b n). rewrite EFin_max lt_max; apply/orP; right. - rewrite /case_sum' -bigmaxEFin. + rewrite /case_sum' -EFin_bigmax. apply: lt_le_trans; last exact: le_bigmax_cond. - by rewrite /f2'; case: cid => /=. + by rewrite /f2'; case: cid => /=C. move=> [x [a|b]] U mU/=-. - by rewrite (Hf1 a x _ mU). - by rewrite (Hf2 b x _ mU). @@ -1292,7 +1257,7 @@ HB.instance Definition _ := End kcounting. -(* formalization of the iterate construct of Staton ESOP 2017, Sect. 4.2 *) +(* formalization of the iterate construct [Section 4.2, Staton ESOP 2017] *) Section iterate. Context d {G : measurableType d} {R : realType}. Context dA (A : measurableFinType dA) dB (B : measurableFinType dB). @@ -1446,7 +1411,6 @@ Qed. End iterate_unit. -(* an s-finite kernel to test that two expressions are different *) Section lift_neq. Context {R : realType} d (G : measurableType d). Variables (f : G -> bool) (g : G -> bool). @@ -1478,155 +1442,6 @@ Definition lift_neq : R.-sfker G ~> bool := ret measurable_fun_flift_neq. End lift_neq. -Section von_neumann_trick. -Context d {T : measurableType d} {R : realType}. - -Definition minltt {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} := - @measurable_cst _ _ T1 _ setT (@inl _ T2 tt). - -Definition finrb d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) : - T1 * bool -> T2 + bool := fun t1b => inr t1b.2. - -Lemma minrb {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} : - measurable_fun setT (@finrb _ _ T1 T2). -Proof. exact: measurableT_comp. Qed. - -(* biased coin *) -Variable (D : pprobability bool R). -Let unit := measurableTypeUnit. -Let bool := measurableTypeBool. - -Definition trick : R.-sfker (T * unit) ~> (unit + bool)%type := - letin (sample_cst D) - (letin (sample_cst D) - (letin (lift_neq macc1of3 macc2of3) - (ite macc3of4 - (letin (ret macc1of4) (ret minrb)) - (ret minltt)))). - -Definition kvon_neumann_trick : _ -> _ := - (@iterate _ _ R _ unit _ bool trick _ ktt). -Definition von_neumann_trick x : _ -> _ := kvon_neumann_trick x. - -HB.instance Definition _ := SFiniteKernel.on kvon_neumann_trick. -HB.instance Definition _ x := Measure.on (von_neumann_trick x). - -Section von_neumann_trick_proof. - -Let p : R := fine (D [set true]). -Let q : R := p * (1 - p). -Let r : R := p ^+ 2 + (1 - p) ^+ 2. - -Let Dtrue : D [set true] = p%:E. -Proof. by rewrite fineK//= fin_num_measure. Qed. - -Let Dfalse : D [set false] = 1 - (p%:E). -Proof. -rewrite -Dtrue; have <- : D [set: bool] = 1 := probability_setT _. -rewrite setT_bool measureU//=; last first. - by rewrite disjoints_subset => -[]//. -by rewrite addeAC subee ?fin_num_measure ?add0e. -Qed. - -Let p_ge0 : (0 <= p)%R. -Proof. by rewrite fine_ge0 ?measure_ge0. Qed. - -Let p'_ge0 : (1 - p >= 0)%R. -Proof. by rewrite -lee_fin EFinB -Dfalse measure_ge0. Qed. - -Let q_ge0 : (q >= 0)%R. Proof. by rewrite mulr_ge0 ?p_ge0 ?p'_ge0. Qed. - -Let r_ge0 : (r >= 0)%R. -Proof. by rewrite addr_ge0// exprn_ge0 ?p_ge0 ?p'_ge0. Qed. - -Let intDE f : (forall x, 0 <= f x) -> - \int[D]_x (f x) = (1 - (p%:E)) * f false + p%:E * f true. -Proof. -move=> f_ge0. -pose M := measure_add (mscale (NngNum p'_ge0) \d_false) - (mscale (NngNum p_ge0) \d_true). -rewrite (eq_measure_integral M) => [|A Ameas _]. - rewrite ge0_integral_measure_add//. - by rewrite !ge0_integral_mscale ?integral_dirac ?diracT ?mul1e. -rewrite /M/= /msum/= !big_ord_recl/= big_ord0 addr0 /mscale/=. -by case: (set_bool A) => /eqP->/=; - rewrite ?measure0 ?probability_setT ?Dtrue ?Dfalse//= - ?diracE/= ?in_set1 ?inE/= ?(mule0, mul0e, adde0, add0e, mule1)// - -EFinD addrNK. -Qed. - -Lemma trick_bool gamma b : trick gamma [set inr b] = q%:E. -Proof. -rewrite /trick/= /kcomp /= intDE//= /kcomp/= !intDE//= /kcomp. -rewrite !integral_dirac ?diracT//= ?mul1e. -rewrite !iteE//= ?diracE/= /kcomp/=. -rewrite !integral_dirac ?diracT ?diracE ?mul1e//=. -rewrite /finrb ?inl_in_set_inr//= /acc1of4/= ?inr_in_set_inr !in_set1 !inE. -rewrite /q -?(EFinB, EFinN, EFinM, EFinD); congr (_)%:E. -by case: b => //=; ring. -Qed. - -Lemma trick_unit gamma : trick gamma [set inl tt] = r%:E. -Proof. -rewrite /trick/= /kcomp /= intDE//= /kcomp/= !intDE//= /kcomp. -rewrite !integral_dirac ?diracT//= ?mul1e. -rewrite !iteE//= ?diracE/= /kcomp/=. -rewrite !integral_dirac ?diracT ?diracE ?mul1e//=. -rewrite /finrb ?inl_in_set_inr//= /acc1of4/= ?inr_in_set_inr. -rewrite !in_set1 !inE/=. -by rewrite /r -?(EFinB, EFinN, EFinM, EFinD); congr (_)%:E; ring. -Qed. - -Hypothesis D_nontrivial : 0 < D [set true] < 1. - -Let p_gt0 : (0 < p)%R. -Proof. by rewrite -lte_fin -Dtrue; case/andP : D_nontrivial. Qed. - -Let p_lt1 : (p < 1)%R. -Proof. by rewrite -lte_fin -Dtrue; case/andP : D_nontrivial. Qed. - -Let p'_gt0 : (0 < 1 - p)%R. Proof. by rewrite subr_gt0. Qed. - -Let r_lt1 : (r < 1)%R. -Proof. -rewrite /r -subr_gt0 [ltRHS](_ : _ = 2 * p * (1 - p))%R; last by ring. -by rewrite !mulr_gt0. -Qed. - -Lemma von_neumann_trick_prob_kernel gamma b : - kvon_neumann_trick gamma [set b] = 2^-1%:E. -Proof. -rewrite [LHS](@iterateE _ _ _ _ _ _ _ _ r _ _ _ q)//=. -- rewrite /r /q; congr (_)%:E. - suff: (1 - ((p ^+ 2)%R + ((1 - p) ^+ 2)%R)%E)%R != 0%R by move=> *; field. - rewrite [X in X != _](_ : _ = 2 * (p * (1 - p)))%R; last by ring. - by rewrite mulf_eq0 ?pnatr_eq0/= mulf_neq0// gt_eqF ?p_gt0 ?p'_gt0. -- by move=> gamma'; exact: trick_unit. -- suff -> : [set inr x | x in [set b]] = [set inr b] by exact: trick_bool. - by move=> A; apply/seteqP; split=> [a [_ ->]|_ ->]//=; exists b. -Qed. - -Lemma von_neumann_trick_prob_kernelT gamma : - von_neumann_trick gamma [set: bool] = 1. -Proof. -rewrite setT_bool measureU//=; last by rewrite disjoints_subset => -[]. -rewrite !von_neumann_trick_prob_kernel -EFinD. -by have := splitr (1 : R); rewrite mul1r => <-. -Qed. - -HB.instance Definition _ gamma := Measure.on (von_neumann_trick gamma). -HB.instance Definition _ gamma := Measure_isProbability.Build _ _ _ - (von_neumann_trick gamma) (von_neumann_trick_prob_kernelT gamma). -HB.instance Definition _ := Kernel_isProbability.Build _ _ _ _ _ - kvon_neumann_trick von_neumann_trick_prob_kernelT. - -Theorem von_neumann_trickP gamma : von_neumann_trick gamma =1 bernoulli 2^-1. -Proof. by apply: eq_bernoulli; rewrite von_neumann_trick_prob_kernel. Qed. - -End von_neumann_trick_proof. - -End von_neumann_trick. - Section insn1_lemmas. Import Notations. Context d (T : measurableType d) (R : realType). @@ -1696,6 +1511,7 @@ Qed. End letin_ite. +(* associativity of let [Section 4.2, Staton ESOP 2017] *) Section letinA. Context d d' d1 d2 d3 (X : measurableType d) (Y : measurableType d') (T1 : measurableType d1) (T2 : measurableType d2) (T3 : measurableType d3) @@ -1724,6 +1540,7 @@ Qed. End letinA. +(* commutativity of let [Section 4.2, Staton ESOP 2017] *) Section letinC. Context d d1 d' (X : measurableType d) (Y : measurableType d1) (Z : measurableType d') (R : realType). @@ -1822,7 +1639,6 @@ Context d (T : measurableType d) (R : realType). (* let x = sample (bernoulli (2/7)) in let r = case x of {(1, _) => return (k3()), (2, _) => return (k10())} in return r *) - Definition sample_and_branch : R.-sfker T ~> _ := letin (sample_cst (bernoulli (2 / 7))) (* T -> B *) @@ -1978,6 +1794,155 @@ Qed. End staton_bus_exponential. + +Section von_neumann_trick. +Context d {T : measurableType d} {R : realType}. + +Definition minltt {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} := + @measurable_cst _ _ T1 _ setT (@inl _ T2 tt). + +Definition finrb d1 d2 (T1 : measurableType d1) (T2 : measurableType d2) : + T1 * bool -> T2 + bool := fun t1b => inr t1b.2. + +Lemma minrb {d1 d2} {T1 : measurableType d1} {T2 : measurableType d2} : + measurable_fun setT (@finrb _ _ T1 T2). +Proof. exact: measurableT_comp. Qed. + +Variable (D : pprobability bool R). (* biased coin *) +Let unit := measurableTypeUnit. +Let bool := measurableTypeBool. + +Definition trick : R.-sfker (T * unit) ~> (unit + bool)%type := + letin (sample_cst D) + (letin (sample_cst D) + (letin (lift_neq macc1of3 macc2of3) + (ite macc3of4 + (letin (ret macc1of4) (ret minrb)) + (ret minltt)))). + +Definition kvon_neumann_trick : _ -> _ := + (@iterate _ _ R _ unit _ bool trick _ ktt). +Definition von_neumann_trick x : _ -> _ := kvon_neumann_trick x. + +HB.instance Definition _ := SFiniteKernel.on kvon_neumann_trick. +HB.instance Definition _ x := Measure.on (von_neumann_trick x). + +Section von_neumann_trick_proof. + +Let p : R := fine (D [set true]). +Let q : R := p * (1 - p). +Let r : R := p ^+ 2 + (1 - p) ^+ 2. + +Let Dtrue : D [set true] = p%:E. +Proof. by rewrite fineK//= fin_num_measure. Qed. + +Let Dfalse : D [set false] = 1 - (p%:E). +Proof. +rewrite -Dtrue; have <- : D [set: bool] = 1 := probability_setT _. +rewrite setT_bool measureU//=; last first. + by rewrite disjoints_subset => -[]//. +by rewrite addeAC subee ?fin_num_measure ?add0e. +Qed. + +Let p_ge0 : (0 <= p)%R. +Proof. by rewrite fine_ge0 ?measure_ge0. Qed. + +Let p'_ge0 : (1 - p >= 0)%R. +Proof. by rewrite -lee_fin EFinB -Dfalse measure_ge0. Qed. + +Let q_ge0 : (q >= 0)%R. Proof. by rewrite mulr_ge0 ?p_ge0 ?p'_ge0. Qed. + +Let r_ge0 : (r >= 0)%R. +Proof. by rewrite addr_ge0// exprn_ge0 ?p_ge0 ?p'_ge0. Qed. + +Let intDE f : (forall x, 0 <= f x) -> + \int[D]_x (f x) = (1 - (p%:E)) * f false + p%:E * f true. +Proof. +move=> f_ge0. +pose M := measure_add (mscale (NngNum p'_ge0) \d_false) + (mscale (NngNum p_ge0) \d_true). +rewrite (eq_measure_integral M) => [|A Ameas _]. + rewrite ge0_integral_measure_add//. + by rewrite !ge0_integral_mscale ?integral_dirac ?diracT ?mul1e. +rewrite /M/= /msum/= !big_ord_recl/= big_ord0 addr0 /mscale/=. +by case: (set_bool A) => /eqP->/=; + rewrite ?measure0 ?probability_setT ?Dtrue ?Dfalse//= + ?diracE/= ?in_set1 ?inE/= ?(mule0, mul0e, adde0, add0e, mule1)// + -EFinD addrNK. +Qed. + +Lemma trick_bool gamma b : trick gamma [set inr b] = q%:E. +Proof. +rewrite /trick/= /kcomp /= intDE//= /kcomp/= !intDE//= /kcomp. +rewrite !integral_dirac ?diracT//= ?mul1e. +rewrite !iteE//= ?diracE/= /kcomp/=. +rewrite !integral_dirac ?diracT ?diracE ?mul1e//=. +rewrite /finrb ?inl_in_set_inr//= /acc1of4/= ?inr_in_set_inr !in_set1 !inE. +rewrite /q -?(EFinB, EFinN, EFinM, EFinD); congr (_)%:E. +by case: b => //=; ring. +Qed. + +Lemma trick_unit gamma : trick gamma [set inl tt] = r%:E. +Proof. +rewrite /trick/= /kcomp /= intDE//= /kcomp/= !intDE//= /kcomp. +rewrite !integral_dirac ?diracT//= ?mul1e. +rewrite !iteE//= ?diracE/= /kcomp/=. +rewrite !integral_dirac ?diracT ?diracE ?mul1e//=. +rewrite /finrb ?inl_in_set_inr//= /acc1of4/= ?inr_in_set_inr. +rewrite !in_set1 !inE/=. +by rewrite /r -?(EFinB, EFinN, EFinM, EFinD); congr (_)%:E; ring. +Qed. + +Hypothesis D_nontrivial : 0 < D [set true] < 1. + +Let p_gt0 : (0 < p)%R. +Proof. by rewrite -lte_fin -Dtrue; case/andP : D_nontrivial. Qed. + +Let p_lt1 : (p < 1)%R. +Proof. by rewrite -lte_fin -Dtrue; case/andP : D_nontrivial. Qed. + +Let p'_gt0 : (0 < 1 - p)%R. Proof. by rewrite subr_gt0. Qed. + +Let r_lt1 : (r < 1)%R. +Proof. +rewrite /r -subr_gt0 [ltRHS](_ : _ = 2 * p * (1 - p))%R; last by ring. +by rewrite !mulr_gt0. +Qed. + +Lemma von_neumann_trick_prob_kernel gamma b : + kvon_neumann_trick gamma [set b] = 2^-1%:E. +Proof. +rewrite [LHS](@iterateE _ _ _ _ _ _ _ _ r _ _ _ q)//=. +- rewrite /r /q; congr (_)%:E. + suff: (1 - ((p ^+ 2)%R + ((1 - p) ^+ 2)%R)%E)%R != 0%R by move=> *; field. + rewrite [X in X != _](_ : _ = 2 * (p * (1 - p)))%R; last by ring. + by rewrite mulf_eq0 ?pnatr_eq0/= mulf_neq0// gt_eqF ?p_gt0 ?p'_gt0. +- by move=> gamma'; exact: trick_unit. +- suff -> : [set inr x | x in [set b]] = [set inr b] by exact: trick_bool. + by move=> A; apply/seteqP; split=> [a [_ ->]|_ ->]//=; exists b. +Qed. + +Lemma von_neumann_trick_prob_kernelT gamma : + von_neumann_trick gamma [set: bool] = 1. +Proof. +rewrite setT_bool measureU//=; last by rewrite disjoints_subset => -[]. +rewrite !von_neumann_trick_prob_kernel -EFinD. +by have := splitr (1 : R); rewrite mul1r => <-. +Qed. + +HB.instance Definition _ gamma := Measure.on (von_neumann_trick gamma). +HB.instance Definition _ gamma := Measure_isProbability.Build _ _ _ + (von_neumann_trick gamma) (von_neumann_trick_prob_kernelT gamma). +HB.instance Definition _ := Kernel_isProbability.Build _ _ _ _ _ + kvon_neumann_trick von_neumann_trick_prob_kernelT. + +Theorem von_neumann_trickP gamma : von_neumann_trick gamma =1 bernoulli 2^-1. +Proof. by apply: eq_bernoulli; rewrite von_neumann_trick_prob_kernel. Qed. + +End von_neumann_trick_proof. + +End von_neumann_trick. + (**md letin' variants *) diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index a9b321353a..782f01c20b 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -269,12 +269,12 @@ move=> /= F mF tF mUF. rewrite /mpoisson1/= integral_bigcup//=; last first. apply/integrableP; split. apply/measurable_EFinP. - exact: measurable_funS (measurable_poisson _). + exact: measurable_funS (measurable_poisson_pdf _). rewrite (_ : (fun x => _) = (EFin \o poisson1)); last first. by apply/funext => x; rewrite gee0_abs// lee_fin poisson1_ge0//. apply: le_lt_trans. apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. - by apply/measurable_EFinP; exact: measurable_poisson. + by apply/measurable_EFinP; exact: measurable_poisson_pdf. by move=> ? _; rewrite lee_fin poisson1_ge0//. by rewrite /= integral_poisson_density// ltry. apply: is_cvg_ereal_nneg_natsum_cond => n _ _. @@ -285,11 +285,7 @@ HB.instance Definition _ := isMeasure.Build _ _ _ mpoisson1 mpoisson10 mpoisson1_ge0 mpoisson1_sigma_additive. Let mpoisson1_setT : mpoisson1 [set: _] = 1%E. -Proof. -rewrite /mpoisson1. -rewrite /poisson1. -by rewrite integral_poisson_density. -Qed. +Proof. exact: integral_poisson_density. Qed. HB.instance Definition _ := @Measure_isProbability.Build _ _ R mpoisson1 mpoisson1_setT. From 9b42f824e0ab1996729dc73e4cbaa7f086808475 Mon Sep 17 00:00:00 2001 From: Cyril Cohen Date: Thu, 2 Jan 2025 19:16:03 +0100 Subject: [PATCH 29/48] golfing --- theories/lang_syntax.v | 6 ++-- theories/prob_lang.v | 77 +++++++++++++--------------------------- theories/prob_lang_wip.v | 5 +-- 3 files changed, 32 insertions(+), 56 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 3b2b376ff7..e97dbb849d 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -1337,7 +1337,8 @@ Lemma beta_pdf_uniq_ae (a b : nat) : (EFin \o (beta_pdf a b)). Proof. apply: integral_ae_eq => //. -- apply: integrableS (Radon_Nikodym_integrable _) => //. +- apply: (integrableS _ _ (@subsetT _ _)) => //=. + apply: (Radon_Nikodym_integrable _) => //=. exact: beta_prob_dom. - apply/measurable_funTS/measurableT_comp => //. exact: measurable_beta_pdf. @@ -1367,7 +1368,8 @@ rewrite -(Radon_Nikodym_change_of_variables (beta_prob_dom a b)) //=; last first by apply/integrableP; split. apply: ae_eq_integral => //. - apply: emeasurable_funM => //; apply: measurable_int. - apply: integrableS (Radon_Nikodym_integrable _) => //=. + apply: (integrableS _ _ (@subsetT _ _)) => //=. + apply: (Radon_Nikodym_integrable _) => //=. exact: beta_prob_dom. - apply: emeasurable_funM => //=; apply/measurableT_comp => //=. by apply/measurable_funTS; exact: measurable_beta_pdf. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 532a0db19f..cb9a4842f2 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -1820,6 +1820,9 @@ Definition trick : R.-sfker (T * unit) ~> (unit + bool)%type := (letin (ret macc1of4) (ret minrb)) (ret minltt)))). +HB.instance Definition _ := SFiniteKernel.on trick. +HB.instance Definition _ x := Measure.on (trick x). + Definition kvon_neumann_trick : _ -> _ := (@iterate _ _ R _ unit _ bool trick _ ktt). Definition von_neumann_trick x : _ -> _ := kvon_neumann_trick x. @@ -1836,63 +1839,33 @@ Let r : R := p ^+ 2 + (1 - p) ^+ 2. Let Dtrue : D [set true] = p%:E. Proof. by rewrite fineK//= fin_num_measure. Qed. -Let Dfalse : D [set false] = 1 - (p%:E). -Proof. -rewrite -Dtrue; have <- : D [set: bool] = 1 := probability_setT _. -rewrite setT_bool measureU//=; last first. - by rewrite disjoints_subset => -[]//. -by rewrite addeAC subee ?fin_num_measure ?add0e. -Qed. - -Let p_ge0 : (0 <= p)%R. -Proof. by rewrite fine_ge0 ?measure_ge0. Qed. - -Let p'_ge0 : (1 - p >= 0)%R. -Proof. by rewrite -lee_fin EFinB -Dfalse measure_ge0. Qed. - -Let q_ge0 : (q >= 0)%R. Proof. by rewrite mulr_ge0 ?p_ge0 ?p'_ge0. Qed. - -Let r_ge0 : (r >= 0)%R. -Proof. by rewrite addr_ge0// exprn_ge0 ?p_ge0 ?p'_ge0. Qed. - -Let intDE f : (forall x, 0 <= f x) -> - \int[D]_x (f x) = (1 - (p%:E)) * f false + p%:E * f true. +Lemma trickE gamma X : trick gamma X = + (r *+ (inl tt \in X) + + q *+ ((inr true \in X) + (inr false \in X)))%:E. Proof. -move=> f_ge0. -pose M := measure_add (mscale (NngNum p'_ge0) \d_false) - (mscale (NngNum p_ge0) \d_true). -rewrite (eq_measure_integral M) => [|A Ameas _]. - rewrite ge0_integral_measure_add//. - by rewrite !ge0_integral_mscale ?integral_dirac ?diracT ?mul1e. -rewrite /M/= /msum/= !big_ord_recl/= big_ord0 addr0 /mscale/=. -by case: (set_bool A) => /eqP->/=; - rewrite ?measure0 ?probability_setT ?Dtrue ?Dfalse//= - ?diracE/= ?in_set1 ?inE/= ?(mule0, mul0e, adde0, add0e, mule1)// - -EFinD addrNK. -Qed. - -Lemma trick_bool gamma b : trick gamma [set inr b] = q%:E. -Proof. -rewrite /trick/= /kcomp /= intDE//= /kcomp/= !intDE//= /kcomp. +have Dbernoulli : D =1 bernoulli p by exact/eq_bernoulli/Dtrue. +have p_itv01 : (0 <= p <= 1)%R. + by rewrite -2!lee_fin -Dtrue?measure_ge0 ?probability_le1. +pose eqbern := eq_measure_integral _ (fun x _ _ => Dbernoulli x). +rewrite /trick/= /kcomp. +do 2?rewrite ?eqbern ?integral_bernoulli//= /kcomp/=. rewrite !integral_dirac ?diracT//= ?mul1e. rewrite !iteE//= ?diracE/= /kcomp/=. -rewrite !integral_dirac ?diracT ?diracE ?mul1e//=. -rewrite /finrb ?inl_in_set_inr//= /acc1of4/= ?inr_in_set_inr !in_set1 !inE. -rewrite /q -?(EFinB, EFinN, EFinM, EFinD); congr (_)%:E. -by case: b => //=; ring. +rewrite !integral_dirac /acc1of4/= ?diracT ?diracE ?mul1e//. +rewrite /finrb /acc1of4/= -?(EFinB, EFinN, EFinM, EFinD) /q /r /onem. +by congr (_)%:E; do 3!move: (_ \in _) => ? /=; ring. Qed. -Lemma trick_unit gamma : trick gamma [set inl tt] = r%:E. +Lemma trick_prob_kernelT gamma : trick gamma setT = 1. Proof. -rewrite /trick/= /kcomp /= intDE//= /kcomp/= !intDE//= /kcomp. -rewrite !integral_dirac ?diracT//= ?mul1e. -rewrite !iteE//= ?diracE/= /kcomp/=. -rewrite !integral_dirac ?diracT ?diracE ?mul1e//=. -rewrite /finrb ?inl_in_set_inr//= /acc1of4/= ?inr_in_set_inr. -rewrite !in_set1 !inE/=. -by rewrite /r -?(EFinB, EFinN, EFinM, EFinD); congr (_)%:E; ring. +by rewrite trickE !mem_setT mulr2n mulr1n /r /q; congr (_)%:E; ring. Qed. +HB.instance Definition _ gamma := Measure_isProbability.Build _ _ _ + (trick gamma) (trick_prob_kernelT gamma). +HB.instance Definition _ := Kernel_isProbability.Build _ _ _ _ _ + trick trick_prob_kernelT. + Hypothesis D_nontrivial : 0 < D [set true] < 1. Let p_gt0 : (0 < p)%R. @@ -1917,9 +1890,9 @@ rewrite [LHS](@iterateE _ _ _ _ _ _ _ _ r _ _ _ q)//=. suff: (1 - ((p ^+ 2)%R + ((1 - p) ^+ 2)%R)%E)%R != 0%R by move=> *; field. rewrite [X in X != _](_ : _ = 2 * (p * (1 - p)))%R; last by ring. by rewrite mulf_eq0 ?pnatr_eq0/= mulf_neq0// gt_eqF ?p_gt0 ?p'_gt0. -- by move=> gamma'; exact: trick_unit. -- suff -> : [set inr x | x in [set b]] = [set inr b] by exact: trick_bool. - by move=> A; apply/seteqP; split=> [a [_ ->]|_ ->]//=; exists b. +- by move=> gamma'; rewrite trickE//= ?in_set1 ?inE//= addr0. +- rewrite trickE/= ?inl_in_set_inr ?inr_in_set_inr// add0r !in_set1 !inE. + by case: b. Qed. Lemma von_neumann_trick_prob_kernelT gamma : diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 782f01c20b..0b6b743c9c 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -168,8 +168,9 @@ apply: ae_eq_integral => //=. - apply: emeasurable_funM => //. apply/measurable_funTS/measurableT_comp => //. exact: measurable_fun_f1. - apply: measurable_int. - apply: integrableS (Radon_Nikodym_integrable _) => //=. + apply: (measurable_int mu). + apply: (integrableS _ _ (@subsetT _ _)) => //=. + apply: (Radon_Nikodym_integrable _) => //=. exact: gauss01_dom. - apply: emeasurable_funM => //. apply/measurable_funTS/measurableT_comp => //. From fad7c1d55195186c1708d6943f285718bd91ccd8 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Fri, 3 Jan 2025 11:08:40 +0900 Subject: [PATCH 30/48] use gauss_integral to complete another example --- theories/lang_syntax.v | 13 +- theories/lang_syntax_examples.v | 1 + theories/lang_syntax_table_game.v | 1 + theories/lang_syntax_toy.v | 7 +- theories/lang_syntax_util.v | 5 +- theories/prob_lang.v | 143 ++++++++++++++++++- theories/prob_lang_wip.v | 223 +----------------------------- 7 files changed, 161 insertions(+), 232 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index e97dbb849d..02bc16d627 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -18,10 +18,13 @@ From mathcomp Require Import ring lra. (* Probabilistic Programming Language in Coq using s-finite kernels in Coq. *) (* APLAS 2023 *) (* *) -(* beta distribution specialized to nat *) +(* beta distribution *) +(* ``` *) (* beta_pdf == probability density function for beta *) (* beta_prob == beta probability measure *) +(* ``` *) (* *) +(* ``` *) (* typ == syntax for types of data structures *) (* measurable_of_typ t == the measurable type corresponding to type t *) (* It is of type {d & measurableType d} *) @@ -58,6 +61,7 @@ From mathcomp Require Import ring lra. (* measurable *) (* execP e == a s-finite kernel corresponding to the evaluation *) (* of the probabilistic expression e *) +(* ``` *) (* *) (******************************************************************************) @@ -1337,8 +1341,8 @@ Lemma beta_pdf_uniq_ae (a b : nat) : (EFin \o (beta_pdf a b)). Proof. apply: integral_ae_eq => //. -- apply: (integrableS _ _ (@subsetT _ _)) => //=. - apply: (Radon_Nikodym_integrable _) => //=. +- apply: (@integrableS _ _ _ _ setT) => //=. + apply: Radon_Nikodym_integrable => //=. exact: beta_prob_dom. - apply/measurable_funTS/measurableT_comp => //. exact: measurable_beta_pdf. @@ -1369,8 +1373,7 @@ rewrite -(Radon_Nikodym_change_of_variables (beta_prob_dom a b)) //=; last first apply: ae_eq_integral => //. - apply: emeasurable_funM => //; apply: measurable_int. apply: (integrableS _ _ (@subsetT _ _)) => //=. - apply: (Radon_Nikodym_integrable _) => //=. - exact: beta_prob_dom. + by apply: (Radon_Nikodym_integrable _) => //=; exact: beta_prob_dom. - apply: emeasurable_funM => //=; apply/measurableT_comp => //=. by apply/measurable_funTS; exact: measurable_beta_pdf. - apply: ae_eqe_mul2l => /=. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index de25619715..c72101c828 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -1,3 +1,4 @@ +(* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *) From Coq Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v index 27da2a4ef1..26d61be498 100644 --- a/theories/lang_syntax_table_game.v +++ b/theories/lang_syntax_table_game.v @@ -1,3 +1,4 @@ +(* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *) Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. diff --git a/theories/lang_syntax_toy.v b/theories/lang_syntax_toy.v index d861a0d2e4..7a4ca222f0 100644 --- a/theories/lang_syntax_toy.v +++ b/theories/lang_syntax_toy.v @@ -1,3 +1,4 @@ +(* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *) From Coq Require Import String Classical. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg interval_inference. @@ -5,8 +6,8 @@ From mathcomp Require Import mathcomp_extra boolp. From mathcomp Require Import reals topology normedtype. From mathcomp Require Import lang_syntax_util. -(******************************************************************************) -(* Intrinsically-typed concrete syntax for a toy language *) +(**md**************************************************************************) +(* # Intrinsically-typed concrete syntax for a toy language *) (* *) (* The main module provided by this file is "lang_intrinsic_tysc" which *) (* provides an example of intrinsically-typed concrete syntax for a toy *) @@ -14,10 +15,12 @@ From mathcomp Require Import lang_syntax_util. (* lang_syntax.v). Other modules provide even more simplified language for *) (* pedagogical purposes. *) (* *) +(* ``` *) (* lang_extrinsic == non-intrinsic definition of expression *) (* lang_intrinsic_ty == intrinsically-typed syntax *) (* lang_intrinsic_sc == intrinsically-scoped syntax *) (* lang_intrinsic_tysc == intrinsically-typed/scoped syntax *) +(* ``` *) (* *) (******************************************************************************) diff --git a/theories/lang_syntax_util.v b/theories/lang_syntax_util.v index 2e5b82542e..a7b19a603a 100644 --- a/theories/lang_syntax_util.v +++ b/theories/lang_syntax_util.v @@ -1,11 +1,12 @@ +(* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *) From Coq Require Import String. From HB Require Import structures. Require Import Classical_Prop. (* NB: to compile with Coq 8.17 *) From mathcomp Require Import all_ssreflect. From mathcomp Require Import interval_inference. -(******************************************************************************) -(* Shared by lang_syntax_*.v files *) +(**md**************************************************************************) +(* Shared by lang_syntax_*.v files *) (******************************************************************************) HB.instance Definition _ := hasDecEq.Build string eqb_spec. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index cb9a4842f2..0647e16482 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -1,12 +1,12 @@ -(* mathcomp analysis (c) 2022 Inria and AIST. License: CeCILL-C. *) +(* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *) From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval finmap. From mathcomp Require Import rat archimedean ring lra. From mathcomp Require Import unstable mathcomp_extra boolp classical_sets. From mathcomp Require Import functions cardinality fsbigop interval_inference. From mathcomp Require Import reals ereal topology normedtype sequences. -From mathcomp Require Import esum measure lebesgue_measure numfun. -From mathcomp Require Import lebesgue_integral probability exp kernel charge. +From mathcomp Require Import esum measure lebesgue_measure numfun exp. +From mathcomp Require Import lebesgue_integral trigo probability kernel charge. (**md**************************************************************************) (* # Semantics of a probabilistic programming language using s-finite kernels *) @@ -21,7 +21,9 @@ From mathcomp Require Import lebesgue_integral probability exp kernel charge. (* poisson_pdf == Poisson pdf *) (* exponential_pdf == exponential distribution pdf *) (* measurable_sum X Y == the type X + Y, as a measurable type *) +(* ``` *) (* *) +(* ``` *) (* mscore f t := mscale `|f t| \d_tt *) (* kscore f := fun=> mscore f *) (* This is an s-finite kernel. *) @@ -1466,8 +1468,8 @@ by rewrite /score/= /mscale/= ger0_norm//= f0. Qed. Lemma score_score (f : R -> R) (g : R * unit -> R) - (mf : measurable_fun setT f) - (mg : measurable_fun setT g) : + (mf : measurable_fun [set: R] f) + (mg : measurable_fun [set: R * unit] g) : letin (score mf) (score mg) = score (measurable_funM mf (measurableT_comp mg (pair2_measurable tt))). Proof. @@ -2181,3 +2183,134 @@ rewrite letin'E/= /sample; unlock. rewrite integral_bernoulli ?r0//=. by rewrite /mscale/= iteE//= iteE//= fail'E mule0 adde0 ger0_norm. Qed. + +(* TODO: move to probability.v? *) +Section gauss. +Variable R : realType. +Local Open Scope ring_scope. + +Definition gauss_pdf := @normal_pdf R 0 1. + +Lemma normal_pdf_gt0 m s x : 0 < s -> 0 < normal_pdf m s x :> R. +Proof. +move=> s0; rewrite /normal_pdf gt_eqF// mulr_gt0 ?expR_gt0// invr_gt0. +by rewrite sqrtr_gt0 pmulrn_rgt0// mulr_gt0 ?pi_gt0 ?exprn_gt0. +Qed. + +Lemma gauss_pdf_gt0 x : 0 < gauss_pdf x. +Proof. exact: normal_pdf_gt0. Qed. + +Definition gauss_prob := @normal_prob R 0 1. + +HB.instance Definition _ := Probability.on gauss_prob. + +Lemma gauss_prob_dominates : gauss_prob `<< lebesgue_measure. +Proof. exact: normal_prob_dominates. Qed. + +Lemma continuous_gauss_pdf x : {for x, continuous gauss_pdf}. +Proof. exact: continuous_normal_pdf. Qed. + +End gauss. + +(* the Lebesgue measure is definable in Staton's language + [equation (10), Section 4.1, Staton ESOP 2017] *) +Section gauss_lebesgue. +Context d (T : measurableType d) (R : realType). +Notation mu := (@lebesgue_measure R). + +Let f1 (x : g_sigma_algebraType (R.-ocitv.-measurable)) := (gauss_pdf x)^-1%R. + +Let f1E (x : R) : f1 x = (Num.sqrt (pi *+ 2) * expR (- (- x ^+ 2 / 2)))%R. +Proof. +rewrite /f1 /gauss_pdf /normal_pdf oner_eq0. +rewrite /normal_peak expr1n mul1r. +by rewrite /normal_fun subr0 expr1n invfM invrK expRN. +Qed. + +Let f1_gt0 (x : R) : (0 < f1 x)%R. +Proof. by rewrite f1E mulr_gt0 ?expR_gt0// sqrtr_gt0 mulrn_wgt0// pi_gt0. Qed. + +Lemma measurable_fun_f1 : measurable_fun setT f1. +Proof. +apply: continuous_measurable_fun => x. +apply: (@continuousV _ _ (@gauss_pdf R)). + by rewrite gt_eqF// gauss_pdf_gt0. +exact: continuous_gauss_pdf. +Qed. + +Lemma integral_mgauss01 : forall U, measurable U -> + \int[(@gauss_prob R)]_(y in U) (f1 y)%:E = + \int[mu]_(x0 in U) (gauss_pdf x0 * f1 x0)%:E. +Proof. +move=> U mU. +under [in RHS]eq_integral do rewrite EFinM/= muleC. +rewrite /=. +rewrite -(@Radon_Nikodym_SigmaFinite.change_of_variables + _ _ _ _ (@lebesgue_measure R))//=; last 3 first. + exact: gauss_prob_dominates. + by move=> /= x; rewrite lee_fin ltW. + apply/measurable_EFinP. + apply: measurable_funTS. + exact: measurable_fun_f1. +apply: ae_eq_integral => //=. +- apply: emeasurable_funM => //. + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_fun_f1. + apply: (measurable_int mu). + apply: (integrableS _ _ (@subsetT _ _)) => //=. + apply: Radon_Nikodym_SigmaFinite.f_integrable => /=. + exact: gauss_prob_dominates. +- apply: emeasurable_funM => //. + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_fun_f1. + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_normal_pdf. +- apply: ae_eqe_mul2l => /=. + rewrite /Radon_Nikodym_SigmaFinite.f/=. + case: pselect => [gauss_prob_dom|]; last first. + by move=> /(_ (@gauss_prob_dominates R)). + case: cid => //= h [h1 h2 h3] gauss_probE. + apply: integral_ae_eq => //=. + + exact: integrableS h3. + + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_normal_pdf. + + move=> E EU mE. + by rewrite -gauss_probE. +Qed. + +Let mf1 : measurable_fun setT f1. +Proof. +apply: (measurable_comp (F := [set r : R | r != 0%R])) => //. +- exact: open_measurable. +- by move=> /= r [t _ <-]; rewrite gt_eqF// gauss_pdf_gt0. +- apply: open_continuous_measurable_fun => //. + by apply/in_setP => x /= x0; exact: inv_continuous. +- exact: measurable_normal_pdf. +Qed. + +Definition staton_lebesgue : R.-sfker T ~> _ := + letin (sample_cst (@gauss_prob R : pprobability _ _)) + (letin + (score (measurableT_comp mf1 macc1of2)) + (ret macc1of3)). + +Lemma staton_lebesgueE x U : measurable U -> + staton_lebesgue x U = lebesgue_measure U. +Proof. +move=> mU; rewrite [in LHS]/staton_lebesgue/=. +rewrite [in LHS]letinE /=. +transitivity (\int[(@gauss_prob R)]_(y in U) (f1 y)%:E). + rewrite -[in RHS](setTI U) integral_mkcondr/=. + apply: eq_integral => //= r _. + rewrite letinE/= ge0_integral_mscale//= ger0_norm//; last first. + by rewrite invr_ge0// normal_pdf_ge0. + rewrite integral_dirac// diracT mul1e/= diracE epatch_indic/=. + by rewrite indicE. +rewrite integral_mgauss01//. +transitivity (\int[lebesgue_measure]_(x in U) (\1_U x)%:E). + apply: eq_integral => /= y yU. + by rewrite /f1 divrr ?indicE ?yU// unitfE gt_eqF// gauss_pdf_gt0. +by rewrite integral_indic//= setIid. +Qed. + +End gauss_lebesgue. diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 0b6b743c9c..537236ed40 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -6,11 +6,11 @@ From mathcomp Require Import functions cardinality fsbigop. From mathcomp Require Import interval_inference reals ereal topology normedtype. From mathcomp Require Import sequences esum measure lebesgue_measure numfun. From mathcomp Require Import lebesgue_integral exp kernel trigo prob_lang. -From mathcomp Require Import realfun charge. +From mathcomp Require Import realfun charge probability derive ftc. +From mathcomp Require Import gauss_integral. -(******************************************************************************) -(* Semantics of a probabilistic programming language using s-finite kernels *) -(* (wip about the definition of Lebesgue measure) *) +(**md**************************************************************************) +(* wip waiting for the Poisson distribution *) (******************************************************************************) Set Implicit Arguments. @@ -23,219 +23,6 @@ Local Open Scope classical_set_scope. Local Open Scope ring_scope. Local Open Scope ereal_scope. -Section gauss_pdf. -Context {R : realType}. -Local Open Scope ring_scope. - -Definition gauss_pdf m s x : R := - (s * sqrtr (pi *+ 2))^-1 * expR (- ((x - m) / s) ^+ 2 / 2%:R). - -Lemma gauss_pdf_ge0 m s x : 0 <= s -> 0 <= gauss_pdf m s x. -Proof. by move=> s0; rewrite mulr_ge0 ?expR_ge0// invr_ge0 mulr_ge0. Qed. - -Lemma gauss_pdf_gt0 m s x : 0 < s -> 0 < gauss_pdf m s x. -Proof. -move=> s0; rewrite mulr_gt0 ?expR_gt0// invr_gt0 mulr_gt0//. -by rewrite sqrtr_gt0 pmulrn_rgt0// pi_gt0. -Qed. - -Lemma measurable_gauss_pdf m s : measurable_fun setT (gauss_pdf m s). -Proof. -apply: measurable_funM => //=; apply: measurableT_comp => //=. -apply: measurable_funM => //=; apply: measurableT_comp => //=. -apply: measurableT_comp (exprn_measurable _) _ => /=. -by apply: measurable_funM => //=; exact: measurable_funD. -Qed. - -Definition gauss_pdf01 : R -> R := gauss_pdf 0 1. - -Lemma gauss_pdf01E x : - gauss_pdf01 x = (sqrtr (pi *+ 2))^-1 * expR (- (x ^+ 2) / 2%:R). -Proof. by rewrite /gauss_pdf01 /gauss_pdf mul1r subr0 divr1. Qed. - -Lemma gauss_pdf01_ub x : gauss_pdf01 x <= (Num.sqrt (pi *+ 2))^-1. -Proof. -rewrite -[leRHS]mulr1. -rewrite /gauss_pdf01 /gauss_pdf; last first. -rewrite mul1r subr0 ler_pM2l ?invr_gt0// ?sqrtr_gt0; last by rewrite mulrn_wgt0// pi_gt0. -by rewrite -[leRHS]expR0 ler_expR mulNr oppr_le0 mulr_ge0// sqr_ge0. -Qed. - -Lemma continuous_gauss_pdf1 x : {for x, continuous gauss_pdf01}. -Proof. -apply: continuousM => //=; first exact: cvg_cst. -apply: continuous_comp => /=; last exact: continuous_expR. -apply: continuousM => //=; last exact: cvg_cst. -apply: continuous_comp => //=; last exact: (@continuousN _ R^o). -apply: (@continuous_comp _ _ _ _ (fun x : R => x ^+ 2)%R); last exact: exprn_continuous. -apply: continuousM => //=; last exact: cvg_cst. -by apply: (@continuousD _ R^o) => //=; exact: cvg_cst. -Qed. - -End gauss_pdf. - -Definition gauss01 {R : realType} - of \int[@lebesgue_measure R]_x (gauss_pdf01 x)%:E = 1%E : set _ -> \bar R := - fun V => (\int[lebesgue_measure]_(x in V) (gauss_pdf01 x)%:E)%E. - -Section gauss. -Variable R : realType. -Local Open Scope ring_scope. - -Hypothesis integral_gauss_pdf01 : - (\int[@lebesgue_measure R]_x (gauss_pdf01 x)%:E = 1%E)%E. - -Local Notation gauss01 := (gauss01 integral_gauss_pdf01). - -Let gauss010 : gauss01 set0 = 0%E. -Proof. by rewrite /gauss01 integral_set0. Qed. - -Let gauss01_ge0 A : (0 <= gauss01 A)%E. -Proof. -by rewrite /gauss01 integral_ge0//= => x _; rewrite lee_fin gauss_pdf_ge0. -Qed. - -Let gauss01_sigma_additive : semi_sigma_additive gauss01. -Proof. -move=> /= F mF tF mUF. -rewrite /gauss01/= integral_bigcup//=; last first. - apply/integrableP; split. - apply/measurable_funTS/measurableT_comp => //. - exact: measurable_gauss_pdf. - rewrite (_ : (fun x => _) = EFin \o gauss_pdf01); last first. - by apply/funext => x; rewrite gee0_abs// lee_fin gauss_pdf_ge0. - apply: le_lt_trans. - apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. - by apply/measurable_EFinP; exact: measurable_gauss_pdf. - by move=> ? _; rewrite lee_fin gauss_pdf_ge0. - by rewrite integral_gauss_pdf01 // ltey. -apply: is_cvg_ereal_nneg_natsum_cond => n _ _. -by apply: integral_ge0 => /= x ?; rewrite lee_fin gauss_pdf_ge0. -Qed. - -HB.instance Definition _ := isMeasure.Build _ _ _ - gauss01 gauss010 gauss01_ge0 gauss01_sigma_additive. - -Let gauss01_setT : gauss01 [set: _] = 1%E. -Proof. by rewrite /gauss01 integral_gauss_pdf01. Qed. - -HB.instance Definition _ := @Measure_isProbability.Build _ _ R gauss01 gauss01_setT. - -End gauss. - -Section gauss_lebesgue. -Context d (T : measurableType d) (R : realType). -Notation mu := (@lebesgue_measure R). -Hypothesis integral_gauss_pdf01 : \int[mu]_x (gauss_pdf01 x)%:E = 1%E. - -Lemma gauss01_dom : gauss01 integral_gauss_pdf01 `<< mu. -Proof. -move=> A mA muA0; rewrite /gauss01. -apply/eqP; rewrite eq_le; apply/andP; split; last first. - by apply: integral_ge0 => x _; rewrite lee_fin gauss_pdf_ge0. -apply: (@le_trans _ _ (\int[mu]_(x in A) (Num.sqrt (pi *+ 2))^-1%:E))%E; last first. - by rewrite integral_cst//= muA0 mule0. -apply: ge0_le_integral => //=. -- by move=> x _; rewrite lee_fin gauss_pdf_ge0. -- apply/measurable_funTS/measurableT_comp => //. - exact: measurable_gauss_pdf. -- by move=> x _; rewrite lee_fin gauss_pdf01_ub. -Qed. - -Let f1 (x : g_sigma_algebraType (R.-ocitv.-measurable)) := ((gauss_pdf01 x) ^-1)%R. - -Lemma measurable_fun_f1 : measurable_fun setT f1. -Proof. -apply: continuous_measurable_fun => x. -apply: (@continuousV _ _ gauss_pdf01). - by rewrite gt_eqF// gauss_pdf_gt0. -exact: continuous_gauss_pdf1. -Qed. - -Lemma integrable_f1 U : measurable U -> - (gauss01 integral_gauss_pdf01).-integrable U (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (f1 x)%:E). -Proof. -Admitted. - -Lemma integral_mgauss01 : forall U, measurable U -> - \int[gauss01 integral_gauss_pdf01]_(y in U) (f1 y)%:E = - \int[mu]_(x0 in U) (gauss_pdf01 x0 * f1 x0)%:E. -Proof. -move=> U mU. -under [in RHS]eq_integral do rewrite EFinM/= muleC. -rewrite -(Radon_Nikodym_change_of_variables gauss01_dom _ (integrable_f1 mU))//=. -apply: ae_eq_integral => //=. -- apply: emeasurable_funM => //. - apply/measurable_funTS/measurableT_comp => //. - exact: measurable_fun_f1. - apply: (measurable_int mu). - apply: (integrableS _ _ (@subsetT _ _)) => //=. - apply: (Radon_Nikodym_integrable _) => //=. - exact: gauss01_dom. -- apply: emeasurable_funM => //. - apply/measurable_funTS/measurableT_comp => //. - exact: measurable_fun_f1. - apply/measurable_funTS/measurableT_comp => //. - exact: measurable_gauss_pdf. -- apply: ae_eqe_mul2l => /=. - rewrite Radon_NikodymE//=. - exact: gauss01_dom. - move=> gauss01_dom'. - case: cid => //= h [h1 h2 h3]. - apply: integral_ae_eq => //=. - + exact: integrableS h2. - + apply/measurable_funTS/measurableT_comp => //. - exact: measurable_gauss_pdf. - + by move=> E EU mE; rewrite -(h3 _ mE). -Qed. - -(*Hypothesis integral_mgauss01 : forall U, measurable U -> - \int[gauss01 integral_gauss_pdf01]_(y in U) (f1 y)%:E = - \int[mu]_(x0 in U) (gauss_pdf01 x0 * f1 x0)%:E.*) - -Let mf1 : measurable_fun setT f1. -Proof. -apply: (measurable_comp (F := [set r : R | r != 0%R])) => //. -- exact: open_measurable. -- by move=> /= r [t _ <-]; rewrite gt_eqF// gauss_pdf_gt0. -- apply: open_continuous_measurable_fun => //. - by apply/in_setP => x /= x0; exact: inv_continuous. -- exact: measurable_gauss_pdf. -Qed. - -Definition staton_lebesgue : R.-sfker T ~> _ := - letin (sample_cst (gauss01 integral_gauss_pdf01 : pprobability _ _)) - (letin - (score (measurableT_comp mf1 macc1of2)) - (ret macc1of3)). - -Lemma staton_lebesgueE x U : measurable U -> - staton_lebesgue x U = lebesgue_measure U. -Proof. -move=> mU; rewrite [in LHS]/staton_lebesgue/=. -rewrite [in LHS]letinE /=. -transitivity (\int[gauss01 integral_gauss_pdf01]_(y in U) (f1 y)%:E). - rewrite -[in RHS](setTI U) integral_mkcondr/=. - apply: eq_integral => //= r _. - rewrite letinE/= ge0_integral_mscale//= ger0_norm//; last first. - by rewrite invr_ge0// gauss_pdf_ge0. - rewrite integral_dirac// diracT mul1e/= diracE epatch_indic/=. - by rewrite indicE. -rewrite integral_mgauss01//. -transitivity (\int[lebesgue_measure]_(x in U) (\1_U x)%:E). - apply: eq_integral => /= y yU. - by rewrite /f1 divrr ?indicE ?yU// unitfE gt_eqF// gauss_pdf_gt0. -by rewrite integral_indic//= setIid. -Qed. - -End gauss_lebesgue. - -(* assuming x > 0 *) -Definition Gamma {R : realType} (x : R) : \bar R := - \int[lebesgue_measure]_(t in `[0%R, +oo[%classic) (expR (- t) * powR t (x - 1))%:E. - -Definition Rfact {R : realType} (x : R) := Gamma (x + 1)%R. - Section poisson. Variable R : realType. Local Open Scope ring_scope. @@ -296,7 +83,7 @@ Definition poisson' := [the probability _ _ of mpoisson1]. End poisson. (* Staton's definition of the counting measure - Staton ESOP 2017, Sect. 4.2, equation (13) *) + [equation (13), Sect. 4.2, Staton ESOP 2017] *) Section staton_counting. Context d (X : measurableType d). Variable R : realType. From 7cc8245cff2f3274cbd91d88a5bbc548c307dc6c Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Fri, 6 Jun 2025 07:41:52 +0900 Subject: [PATCH 31/48] fix after rebase --- theories/lang_syntax.v | 1192 +---------------------------- theories/lang_syntax_examples.v | 103 +-- theories/lang_syntax_table_game.v | 100 +-- theories/prob_lang.v | 121 ++- theories/prob_lang_wip.v | 22 +- 5 files changed, 179 insertions(+), 1359 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 02bc16d627..ae334c2698 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -4,7 +4,7 @@ From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. From mathcomp Require Import interval_inference. From mathcomp Require Import unstable mathcomp_extra boolp classical_sets. From mathcomp Require Import functions cardinality fsbigop. -From mathcomp Require Import reals ereal topology normedtype sequences. +From mathcomp Require Import reals ereal topology normedtype sequences exp. From mathcomp Require Import esum measure lebesgue_measure numfun derive realfun. From mathcomp Require Import lebesgue_integral probability ftc kernel charge. From mathcomp Require Import prob_lang lang_syntax_util. @@ -151,7 +151,7 @@ move/left_right_continuousP. apply/not_andP; left. move/(@cvgrPdist_le _ R^o). apply/existsNP. -exists (2%:R^-1)%R. +exists 2%:R^-1%R. rewrite not_implyE; split; first by rewrite invr_gt0. move=> [e /= e0]. move/(_ (-(e / 2))%R). @@ -176,7 +176,7 @@ rewrite invf_lt1//. rewrite {1}(_:1%R = 1%:R)//; apply: ltr_nat. Qed. -Let dintuni : derivable_oo_continuous_bnd (@id R^o) 0 1. +Let dintuni : derivable_oo_LRcontinuous (@id R^o) 0 1. Proof. split. - move=> x _. @@ -203,19 +203,6 @@ Qed. End integral_indicator_function. End integral_indicator_function. -Lemma RintegralZl d {T : measurableType d} {R : realType} - {mu : measure T R} {D : set T} : d.-measurable D -> - forall f : T -> R, - mu.-integrable D (EFin \o f) -> - forall r : R, (\int[mu]_(x in D) (r * f x) = r * \int[mu]_(x in D) f x)%R. -Proof. -move=> mD f intf r. -rewrite /Rintegral. -under eq_integral do rewrite EFinM. -rewrite integralZl// fineM//=. -by apply: integral_fune_fin_num. -Qed. - (* TODO: naming *) Lemma cvg_atNP {T : topologicalType} {R : numFieldType} (f : R -> T) (a : R) (l : T) : f x @[x --> a] --> l <-> (f \o -%R) x @[x --> (- a)%R] --> l. @@ -226,15 +213,15 @@ by apply/seteqP; split=> A; move=> [/= e e0 H]; exists e => //= B /H/=; rewrite Qed. Lemma derivable_oo_bnd_id {R : numFieldType} (a b : R) : - derivable_oo_continuous_bnd (@id R^o) a b. + derivable_oo_LRcontinuous (@id R^o) a b. Proof. by split => //; [exact/cvg_at_right_filter/cvg_id|exact/cvg_at_left_filter/cvg_id]. Qed. Lemma derivable_oo_bndN {R : realFieldType} (f : R -> R^o) a b : - derivable_oo_continuous_bnd f (- a) (- b) -> - derivable_oo_continuous_bnd (f \o -%R) b a. + derivable_oo_LRcontinuous f (- a) (- b) -> + derivable_oo_LRcontinuous (f \o -%R) b a. Proof. move=> [dF cFa cFb]. have oppK : (-%R \o -%R) = @id R by apply/funext => x/=; rewrite opprK. @@ -307,12 +294,6 @@ move=> x y x01 y01. by rewrite le_eqVlt => /predU1P[->//|/dF] => /(_ x01 y01)/ltW. Qed. -Lemma derive1_onem {R : realType} : (fun x0 : R => (1 - x0)%R : R^o)^`() = (cst (-1)%R). -Proof. -apply/funext => x. -by rewrite derive1E deriveB// derive_id derive_cst sub0r. -Qed. - Local Close Scope ereal_scope. Lemma cvg_comp_filter {R : realType} (f g : R -> R) (r l : R) : continuous f -> @@ -341,310 +322,17 @@ by rewrite (lt_le_trans grze)// ge_min lexx. Qed. Local Open Scope ereal_scope. -Section change_of_variables_onem. -Context {R : realType}. -Let mu := (@lebesgue_measure R). - -Lemma onem_change (G : R -> R) (r : R) : - (0 < r <= 1)%R -> - {within `[0%R, r], continuous G} -> - (\int[mu]_(x in `[0%R, r]) (G x)%:E = - \int[mu]_(x in `[(1 - r)%R, 1%R]) (G (1 - x))%:E). -Proof. -move=> r01 cG. -have := @integration_by_substitution_decreasing R (fun x => 1 - x)%R G (1 - r) 1%R. -rewrite opprB subrr addrCA subrr addr0. -move=> ->//. -- apply: eq_integral => x xr. - rewrite !fctE. - by rewrite derive1_onem opprK mulr1. -- rewrite ltrBlDl ltrDr. - by case/andP : r01. -- by move=> x y _ _ xy; rewrite ler_ltB. -- by rewrite derive1_onem; move=> ? ?; apply: cvg_cst. -- by rewrite derive1_onem; exact: is_cvg_cst. -- by rewrite derive1_onem; exact: is_cvg_cst. -- split => /=. - + move=> x xr1. - by apply: derivableB => //. - + apply: cvg_at_right_filter. - rewrite opprB addrCA addrA addrK. - apply: (@cvg_comp_filter _ _ (fun x => 1 - x)%R)=> //=. - move=> x. - apply: (@continuousB _ R^o) => //. - exact: cvg_cst. - under eq_fun do rewrite opprD addrA subrr add0r opprK. - apply: cvg_id. - apply: cvg_at_left_filter. - apply: (@cvgB _ R^o) => //. - exact: cvg_cst. -Qed. - -Lemma Ronem_change (G : R -> R) (r : R) : - (0 < r <= 1)%R -> - {within `[0%R, r], continuous G} -> - (\int[mu]_(x in `[0%R, r]) (G x) = - \int[mu]_(x in `[(1 - r)%R, 1%R]) (G (1 - x)))%R. -Proof. -move=> r01 cG. -rewrite [in LHS]/Rintegral. -by rewrite onem_change. -Qed. - -End change_of_variables_onem. - -Section factD. - -Let factD' n m : (n`! * m`! <= (n + m).+1`!)%N. -Proof. -elim: n m => /= [m|n ih m]. - by rewrite fact0 mul1n add0n factS leq_pmull. -rewrite 2!factS [in X in (_ <= _ * X)%N]addSn -mulnA leq_mul//. -by rewrite ltnS addSnnS leq_addr. -Qed. - -Lemma factD n m : (n`! * m.-1`! <= (n + m)`!)%N. -Proof. -case: m => //= [|m]. - by rewrite fact0 muln1 addn0. -by rewrite addnS factD'. -Qed. - -End factD. - -Lemma leq_prod2 (x y n m : nat) : (n <= x)%N -> (m <= y)%N -> - (\prod_(m <= i < y) i * \prod_(n <= i < x) i <= \prod_(n + m <= i < x + y) i)%N. -Proof. -move=> nx my; rewrite big_addn -addnBA//. -rewrite [in leqRHS]/index_iota -addnBAC// iotaD big_cat/=. -rewrite mulnC leq_mul//. - by apply: leq_prod; move=> i _; rewrite leq_addr. -rewrite subnKC//. -rewrite -[in leqLHS](add0n m) big_addn. -rewrite [in leqRHS](_ : y - m = ((y - m + x) - x))%N; last first. - by rewrite -addnBA// subnn addn0. -rewrite -[X in iota X _](add0n x) big_addn -addnBA// subnn addn0. -by apply: leq_prod => i _; rewrite leq_add2r leq_addr. -Qed. - -Lemma leq_fact2 (x y n m : nat) : (n <= x) %N -> (m <= y)%N -> - (x`! * y`! * ((n + m).+1)`! <= n`! * m`! * ((x + y).+1)`!)%N. -Proof. -move=> nx my. -rewrite (fact_split nx) -!mulnA leq_mul2l; apply/orP; right. -rewrite (fact_split my) mulnCA -!mulnA leq_mul2l; apply/orP; right. -rewrite [leqRHS](_ : _ = (n + m).+1`! * \prod_((n + m).+2 <= i < (x + y).+2) i)%N; last first. - by rewrite -fact_split// ltnS leq_add. -rewrite mulnA mulnC leq_mul2l; apply/orP; right. -do 2 rewrite -addSn -addnS. -exact: leq_prod2. -Qed. - -Lemma bounded_norm_expn_onem {R : realType} (a b : nat) : - [bounded `|x ^+ a * (1 - x) ^+ b|%R : R^o | x in (`[0%R, 1%R]%classic : set R)]. -Proof. -exists 1%R; split; [by rewrite num_real|move=> x x1 /= y]. -rewrite in_itv/= => /andP[y0 y1]. -rewrite ger0_norm// ger0_norm; last first. - by rewrite mulr_ge0 ?exprn_ge0// subr_ge0. -rewrite (le_trans _ (ltW x1))// mulr_ile1 ?exprn_ge0//. -- by rewrite subr_ge0. -- by rewrite exprn_ile1. -- rewrite exprn_ile1 ?subr_ge0//. - by rewrite lerBlDl addrC -lerBlDl subrr. -Qed. - -Lemma measurable_fun_expn_onem {R : realType} a b : - measurable_fun setT (fun x : R => x ^+ a * `1-x ^+ b)%R. -Proof. -apply/measurable_funM => //; apply/measurable_funX => //. -exact: measurable_funB. -Qed. - -Section ubeta_nat_pdf. -Local Open Scope ring_scope. -Context {R : realType}. -Variables a b : nat. - -(* unnormalized pdf *) -(*Definition ubeta_nat_pdf (t : R) := - if (0 <= t <= 1)%R then (t ^+ a.-1 * (`1-t) ^+ b.-1)%R else 0%R. - -Lemma ubeta_nat_pdf_ge0 t : 0 <= ubeta_nat_pdf t. -Proof. -rewrite /ubeta_nat_pdf; case: ifPn => // /andP[t0 t1]. -by rewrite mulr_ge0// exprn_ge0// onem_ge0. -Qed. - -Lemma ubeta_nat_pdf_le1 t : ubeta_nat_pdf t <= 1. -Proof. -rewrite /ubeta_nat_pdf; case: ifPn => // /andP[t0 t1]. -by rewrite mulr_ile1// ?(exprn_ge0,onem_ge0,exprn_ile1,onem_le1). -Qed. - -Lemma measurable_ubeta_nat_pdf : measurable_fun setT ubeta_nat_pdf. -Proof. -rewrite /ubeta_nat_pdf /=; apply: measurable_fun_if => //=; last first. - by rewrite setTI; apply: measurable_funTS; exact: measurable_fun_expn_onem. -by apply: measurable_and => /=; exact: measurable_fun_ler. -Qed. - -Local Notation mu := lebesgue_measure. - -Lemma integral_ubeta_nat_pdf U : - (\int[mu]_(x in U) (ubeta_nat_pdf x)%:E = - \int[mu]_(x in U `&` `[0%R, 1%R]) (ubeta_nat_pdf x)%:E)%E. -Proof. -rewrite [RHS]integral_mkcondr/=; apply: eq_integral => x xU. -rewrite patchE; case: ifPn => //. -rewrite notin_setE/= in_itv/= => /negP. -rewrite negb_and -!ltNge => /orP[x0|x1]. - by rewrite /ubeta_nat_pdf leNgt x0/=. -by rewrite /ubeta_nat_pdf !leNgt x1/= andbF. -Qed. - -Lemma integral_ubeta_nat_pdfT : - (\int[mu]_x (ubeta_nat_pdf x)%:E = - \int[mu]_(x in `[0%R, 1%R]) (ubeta_nat_pdf x)%:E)%E. -Proof. by rewrite integral_ubeta_nat_pdf/= setTI. Qed.*) - -End ubeta_nat_pdf. - -(*Lemma ubeta_nat_pdf11 {R : realType} (x : R) : (0 <= x <= 1)%R -> - ubeta_nat_pdf 1 1 x = 1%R. -Proof. -move=> x01. -by rewrite /ubeta_nat_pdf !expr0 mulr1 x01. -Qed. - -(* normalization constant *) -Definition beta_nat_norm {R : realType} (a b : nat) : R := - fine (\int[@lebesgue_measure R]_x (ubeta_nat_pdf a b x)%:E).*) - -Section beta_nat_Gamma. -Context {R : realType}. -(* -Let mu := @lebesgue_measure R. - -Let B (a b : nat) : \bar R := - \int[mu]_(x in `[0%R, 1%R]%classic) (ubeta_nat_pdf a b x)%:E. -*) -End beta_nat_Gamma. - -(* -Lemma integral_beta_nat_normTE {R : realType} (a b : nat) : - beta_nat_norm a b = - fine (\int[lebesgue_measure]_(t in `[0%R, 1%R]) (t^+a.-1 * (`1-t)^+b.-1)%:E) :> R. -Proof. -rewrite /beta_nat_norm /ubeta_nat_pdf [in RHS]integral_mkcond/=; congr fine. -by apply: eq_integral => /= x _; rewrite patchE mem_setE in_itv/=; case: ifPn. -Qed.*) - -Lemma onemXn_derivable {R : realType} n (x : R) : - derivable (fun y : R^o => `1-y ^+ n : R^o)%R x 1. -Proof. -have := @derivableX R R^o (@onem R) n x 1%R. -rewrite fctE. -apply. -exact: derivableB. -Qed. - Lemma deriveX_idfun {R : realType} n x : 'D_1 (@GRing.exp R^o ^~ n.+1) x = n.+1%:R *: (x ^+ n)%R. Proof. by rewrite exp_derive /GRing.scale/= mulr1. Qed. -Lemma derive1Mr [R : realFieldType] [f : R^o -> R^o] [x r : R^o] : - derivable f x 1 -> ((fun x => f x * r)^`() x = (r * f^`() x)%R :> R)%R. -Proof. -move=> fx1. -rewrite derive1E (deriveM fx1); last by []. -by rewrite -derive1E derive1_cst scaler0 add0r derive1E. -Qed. - -Lemma derive1Ml [R : realFieldType] [f : R^o -> R^o] [x r : R^o] : - derivable f x 1 -> ((fun x => r * f x)^`() x = (r * f^`() x)%R :> R)%R. -Proof. -under eq_fun do rewrite mulrC. -exact: derive1Mr. -Qed. - Lemma decreasing_onem {R : numDomainType} : {homo (fun x : R => (1 - x)%R) : x y /~ (x < y)%R}. Proof. move=> b a ab. by rewrite -ltrN2 !opprB ltr_leB. Qed. -Lemma continuous_onemXn {R : realType} (n : nat) x : - {for x, continuous (fun y : R => `1-y ^+ n)%R}. -Proof. -apply: (@continuous_comp _ _ _ (@onem R) (fun x => GRing.exp x n)). - by apply: (@cvgB _ R^o); [exact: cvg_cst|exact: cvg_id]. -exact: exprn_continuous. -Qed. - -Local Close Scope ereal_scope. - -(* we define a function to help formalizing the beta distribution *) -Section XMonemX. -Context {R : numDomainType}. - -Definition XMonemX a b := fun x : R => x ^+ a * `1-x ^+ b. - -Lemma XMonemX_ge0 a b x : x \in `[0%R, 1%R] -> 0 <= XMonemX a b x :> R. -Proof. -rewrite in_itv/= => /andP[x0 x1]. -by rewrite /XMonemX mulr_ge0// exprn_ge0// subr_ge0. -Qed. - -Lemma XMonemX0 n x : XMonemX 0 n x = `1-x ^+ n :> R. -Proof. by rewrite /XMonemX/= expr0 mul1r. Qed. - -Lemma XMonemX0' n x : XMonemX n 0 x = x ^+ n :> R. -Proof. by rewrite /XMonemX/= expr0 mulr1. Qed. - -Lemma XMonemX00 x : XMonemX 0 0 x = 1 :> R. -Proof. by rewrite XMonemX0 expr0. Qed. - -Lemma XMonemXC a b x : XMonemX a b (1 - x) = XMonemX b a x :> R. -Proof. -by rewrite /XMonemX [in LHS]/onem opprB addrCA subrr addr0 mulrC. -Qed. - -End XMonemX. - -Lemma continuous_XMonemX {R : realType} a b : - continuous (XMonemX a b : R -> R). -Proof. -by move=> x; apply: cvgM; [exact: exprn_continuous|exact: continuous_onemXn]. -Qed. - -Lemma within_continuous_XMonemX {R : realType} a b (A : set R) : - {within A, continuous (XMonemX a b : R -> R)}. -Proof. -by apply: continuous_in_subspaceT => x _; exact: continuous_XMonemX. -Qed. - -Lemma bounded_XMonemX {R : realType} (a b : nat) : - [bounded XMonemX a b x : R^o | x in (`[0, 1]%classic : set R)]. -Proof. -exists 1%R; split; [by rewrite num_real|move=> x x1 /= y y01]. -rewrite ger0_norm//; last by rewrite XMonemX_ge0. -move: y01; rewrite in_itv/= => /andP[? ?]. -rewrite (le_trans _ (ltW x1))// mulr_ile1 ?exprn_ge0//. -- by rewrite subr_ge0. -- by rewrite exprn_ile1. -- by rewrite exprn_ile1 ?subr_ge0// lerBlDl addrC -lerBlDl subrr. -Qed. - -Lemma measurable_fun_XMonemX {R : realType} a b (A : set R) : - measurable_fun A (XMonemX a b). -Proof. -apply/measurable_funM => //; apply/measurable_funX => //. -exact: measurable_funB. -Qed. - -Local Open Scope ereal_scope. - +(* TODO: move *) Lemma integral_exprn {R : realType} n : fine (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (x ^+ n)%:E) = n.+1%:R^-1%R :> R. Proof. @@ -657,7 +345,7 @@ have cF0 : {for 0%R, continuous F}. have cF1 : {for 1%R, continuous F}. apply: continuousM; first exact: cvg_cst. by apply: cX; rewrite /= in_itv/= lexx ler01. -have dcF : derivable_oo_continuous_bnd F 0 1. +have dcF : derivable_oo_LRcontinuous F 0 1. split. - by move=> x x01; apply: derivableM => //; exact: exprn_derivable. - apply: continuous_cvg; first exact: mulrl_continuous. @@ -673,851 +361,9 @@ rewrite (@continuous_FTC2 _ (fun x : R => x ^+ n)%R F)//. by apply: continuous_subspaceT; exact: exprn_continuous. Qed. -Lemma derivable_oo_continuous_bnd_onemXnMr {R : realType} (n : nat) (r : R) : - derivable_oo_continuous_bnd (fun x : R => `1-x ^+ n.+1 * r : R^o)%R 0 1. -Proof. -split. -- by move=> x x01; apply: derivableM => //=; exact: onemXn_derivable. -- apply: cvgM; last exact: cvg_cst. - apply: cvg_at_right_filter. - apply: (@cvg_comp _ _ _ (fun x => 1 - x)%R (fun x => GRing.exp x n.+1)%R). - by apply: (@cvgB _ R^o); [exact: cvg_cst|exact: cvg_id]. - exact: exprn_continuous. -- apply: cvg_at_left_filter. - apply: cvgM; last exact: cvg_cst. - apply: (@cvg_comp _ _ _ (fun x => 1 - x)%R (fun x => GRing.exp x n.+1)%R). - by apply: (@cvgB _ R^o); [exact: cvg_cst|exact: cvg_id]. - exact: exprn_continuous. -Qed. - -Lemma derive_onemXn {R : realType} (n : nat) x : - ((fun y : R => `1-y ^+ n.+1 : R^o)^`() x = - n.+1%:R * `1-x ^+ n)%R. -Proof. -rewrite (@derive1_comp _ (@onem R) (fun x => GRing.exp x n.+1))//; last first. - exact: exprn_derivable. -rewrite derive1E deriveX_idfun derive1E deriveB//. -by rewrite -derive1E derive1_cst derive_id sub0r mulrN1 [in RHS]mulNr. -Qed. - -Lemma integral_onemXn {R : realType} n : - fine (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (`1-x ^+ n)%:E) = n.+1%:R^-1%R :> R. -Proof. -rewrite (@continuous_FTC2 _ _ (fun x : R => ((1 - x) ^+ n.+1 / - n.+1%:R))%R)//=. -- rewrite subrr subr0 expr0n/= mul0r expr1n mul1r sub0r. - by rewrite -invrN -2!mulNrn opprK. -- apply: continuous_in_subspaceT. - by move=> x x01; exact: continuous_onemXn. -- exact: derivable_oo_continuous_bnd_onemXnMr. -- move=> x x01. - rewrite derive1Mr//; last exact: onemXn_derivable. - by rewrite derive_onemXn mulrA mulVf// mul1r. -Qed. - -Lemma integrable_XMonemX {R : realType} (a b : nat) : - lebesgue_measure.-integrable `[0%R, 1%R] (fun x : R => (XMonemX a b x)%:E). -Proof. -apply: continuous_compact_integrable => //. - exact: segment_compact. -apply: continuous_in_subspaceT => x _. -exact: continuous_XMonemX. -Qed. - -Lemma Rintegral_onemXn {R : realType} n : - (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (`1-x ^+ n))%R = n.+1%:R^-1%R :> R. -Proof. -rewrite /Rintegral. -rewrite (@continuous_FTC2 _ _ (fun x : R => ((1 - x) ^+ n.+1 / - n.+1%:R))%R)//=. -- rewrite subrr subr0 expr0n/= mul0r expr1n mul1r sub0r. - by rewrite -invrN -2!mulNrn opprK. -- apply: continuous_in_subspaceT. -- by move=> x x01; exact: continuous_onemXn. -- exact: derivable_oo_continuous_bnd_onemXnMr. -- move=> x x01. - rewrite derive1Mr//; last exact: onemXn_derivable. - by rewrite derive_onemXn mulrA mulVf// mul1r. -Qed. - -(* TODO: move *) -Lemma normr_onem {R : realType} (x : R) : (0 <= x <= 1 -> `| `1-x | <= 1)%R. -Proof. -move=> /andP[x0 x1]; rewrite ler_norml; apply/andP; split. - by rewrite lerBrDl lerBlDr (le_trans x1)// lerDl. -by rewrite lerBlDr lerDl. -Qed. - Local Open Scope ereal_scope. - -Local Open Scope ring_scope. - -Section XMonemX01. -Local Open Scope ring_scope. -Context {R : realType}. -Variables a b : nat. - -Definition XMonemX01 := (@XMonemX R a.-1 b.-1) \_ `[0, 1]. - -Lemma XMonemX01_ge0 t : 0 <= XMonemX01 t. -Proof. -rewrite /XMonemX01 patchE ; case: ifPn => //. -rewrite inE/= in_itv/= => /andP[t0 t1]. -by rewrite mulr_ge0// exprn_ge0// onem_ge0. -Qed. - -Lemma XMonemX01_le1 t : XMonemX01 t <= 1. -Proof. -rewrite /XMonemX01 patchE ; case: ifPn => //. -rewrite inE/= in_itv/= => /andP[t0 t1]. -by rewrite mulr_ile1// ?(exprn_ge0,onem_ge0,exprn_ile1,onem_le1). -Qed. - -Lemma measurable_XMonemX01 : measurable_fun [set: R] XMonemX01. -Proof. -rewrite /XMonemX01 /=; apply/(measurable_restrictT _ _).1 => //. -exact: measurable_fun_XMonemX. -Qed. - -Local Notation mu := lebesgue_measure. - -(* TODO: maybe not that useful *) -Lemma integral_XMonemX01 U : - (\int[mu]_(x in U) (XMonemX01 x)%:E = - \int[mu]_(x in U `&` `[0%R, 1%R]) (XMonemX01 x)%:E)%E. -Proof. -rewrite [RHS]integral_mkcondr /=; apply: eq_integral => x xU /=. -by rewrite /XMonemX01/= restrict_EFin -patch_setI setIid. -Qed. - -End XMonemX01. - -Lemma XMonemX_XMonemX01 {R : realType} a b a' b' (x : R) : (0 < a)%N -> (0 < b)%N -> - x \in `[0%R, 1%R]%classic -> - (XMonemX a' b' x * XMonemX01 a b x = XMonemX01 (a + a') (b + b') x :> R)%R. -Proof. -move=> a0 b0 x01; rewrite /XMonemX01 /= !patchE x01. -rewrite mulrCA -mulrA -exprD mulrA -exprD. -congr (_ ^+ _ * _ ^+ _)%R. - by rewrite addnC -!subn1 subDnCA. -by rewrite -!subn1 subDnCA. -Qed. - -Lemma XMonemX01_11 {R : realType} (x : R) : (0 <= x <= 1)%R -> - XMonemX01 1 1 x = 1%R. -Proof. -by move=> x01; rewrite /XMonemX01 patchE mem_setE in_itv/= x01/= XMonemX00. -Qed. - -(* normalization constant *) -Section betafun. -Context {R : realType}. -Notation mu := (@lebesgue_measure R). -Local Open Scope ring_scope. - -Definition betafun (a b : nat) : R := (\int[mu]_x (XMonemX01 a b x))%R. - -Lemma betafun0 (b : nat) : (0 < b)%N -> betafun 0 b = b%:R ^-1:> R. -Proof. -move=> b0. -rewrite -[LHS]Rintegral_mkcond. -under eq_Rintegral do rewrite XMonemX0. -by rewrite Rintegral_onemXn// prednK. -Qed. - -Lemma betafun00 : betafun 0 0 = 1%R :> R. -Proof. -rewrite -[LHS]Rintegral_mkcond. -under eq_Rintegral do rewrite XMonemX00. -rewrite /Rintegral. -rewrite integral_cst/= ?mul1e; last by exact: measurable_itv. -by rewrite lebesgue_measure_itv/= lte_fin ltr01 -EFinB subr0. -Qed. - -Lemma betafun_sym (a b : nat) : betafun a b = betafun b a :> R. -Proof. -rewrite -[LHS]Rintegral_mkcond. -rewrite Ronem_change//=; last 2 first. - by rewrite ltr01 lexx. - apply: continuous_subspaceT. - by move=> x x01; exact: continuous_XMonemX. -rewrite subrr. -rewrite -[RHS]Rintegral_mkcond. -apply: eq_Rintegral => x x01. -by rewrite XMonemXC. -Qed. - -Lemma betafunS (a b : nat) : - (betafun a.+2 b.+1 = a.+1%:R / b.+1%:R * betafun a.+1 b.+2 :> R)%R. -Proof. -rewrite -[LHS]Rintegral_mkcond. -rewrite (@Rintegration_by_parts _ _ - (fun x => `1-x ^+ b.+1 / - b.+1%:R)%R (fun x => a.+1%:R * x ^+ a)%R); last 7 first. - exact: ltr01. - apply/continuous_subspaceT. - move=> x. - apply: cvgM; [exact: cvg_cst|]. - exact: exprn_continuous. - split. - by move=> x x01; exact: exprn_derivable. - by apply: cvg_at_right_filter; exact: exprn_continuous. - by apply: cvg_at_left_filter; exact: exprn_continuous. - by move=> x x01; rewrite derive1E deriveX_idfun. - apply/continuous_subspaceT. - by move=> x x01; exact: continuous_onemXn. - exact: derivable_oo_continuous_bnd_onemXnMr. - move=> x x01. - rewrite derive1Mr; last exact: onemXn_derivable. - by rewrite derive_onemXn mulrA mulVf// mul1r. -rewrite {1}/onem !(expr1n,mul1r,expr0n,subr0,subrr,mul0r,oppr0)/=. -rewrite sub0r. -transitivity (a.+1%:R / b.+1%:R * (\int[lebesgue_measure]_(x in `[0, 1]) (XMonemX a b.+1 x)) : R)%R. - under [in LHS]eq_Rintegral. - move=> x x01. - rewrite mulrA mulrC mulrA (mulrA _ a.+1%:R) -(mulrA (_ * _)%R). - over. - rewrite /=. - rewrite RintegralZl//=; last exact: integrable_XMonemX. - by rewrite -mulNrn -2!mulNr -invrN -mulNrn opprK (mulrC _ a.+1%:R)//=. -by rewrite Rintegral_mkcond. -Qed. - -Lemma betafunSS (a b : nat) : - (betafun a.+1 b.+1 = - a`!%:R / (\prod_(b.+1 <= i < (a + b).+1) i)%:R * betafun 1 (a + b).+1 :> R)%R. -Proof. -elim: a b => [b|a ih b]. - by rewrite fact0 mul1r add0n /index_iota subnn big_nil invr1 mul1r. -rewrite betafunS. -rewrite ih. -rewrite !mulrA. -congr *%R; last by rewrite addSnnS. -rewrite -mulrA. -rewrite mulrCA. -rewrite 2!mulrA. -rewrite -natrM. -rewrite (mulnC a`!). -rewrite -factS. -rewrite -mulrA. -rewrite -invfM. -rewrite big_add1. -rewrite [in RHS]big_nat_recl/=; last by rewrite addSn ltnS leq_addl. -by rewrite -natrM addSnnS. -Qed. - -Lemma betafun1S (n : nat) : (betafun 1 n.+1 = n.+1%:R^-1 :> R)%R. -Proof. -rewrite /betafun -Rintegral_mkcond. -under eq_Rintegral do rewrite XMonemX0. -by rewrite Rintegral_onemXn. -Qed. - -Lemma betafun_fact (a b : nat) : - (betafun a.+1 b.+1 = (a`! * b`!)%:R / (a + b).+1`!%:R :> R)%R. -Proof. -rewrite betafunSS betafun1S. -rewrite natrM -!mulrA; congr *%R. -(* (b+1 b+2 ... b+1 b+a)^-1 / (a+b+1) = b! / (a+b+1)! *) -rewrite factS. -rewrite [in RHS]mulnC. -rewrite natrM. -rewrite invfM. -rewrite mulrA; congr (_ / _). -rewrite -(@invrK _ b`!%:R) -invfM; congr (_^-1). -apply: (@mulfI _ b`!%:R). - by rewrite gt_eqF// ltr0n fact_gt0. -rewrite mulrA divff// ?gt_eqF// ?ltr0n ?fact_gt0//. -rewrite mul1r. -rewrite [in RHS]fact_prod. -rewrite -natrM; congr (_%:R). -rewrite fact_prod -big_cat/=. -by rewrite /index_iota subn1 -iotaD subn1/= subSS addnK addnC. -Qed. - -Lemma betafunE (a b : nat) : betafun a b = - if (a == 0)%N && (0 < b)%N then - b%:R^-1 - else if (b == 0)%N && (0 < a)%N then - a%:R^-1 - else - a.-1`!%:R * b.-1`!%:R / (a + b).-1`!%:R :> R. -Proof. -case: a => [|a]. - rewrite eqxx/=; case: ifPn => [|]. - by case: b => [|b _] //; rewrite betafun0. - rewrite -leqNgt leqn0 => /eqP ->. - by rewrite betafun00 eqxx ltnn/= fact0 mul1r divr1. -case: b => [|b]. - by rewrite betafun_sym betafun0// fact0 addn0/= mulr1 divff. -by rewrite betafun_fact/= natrM// addnS. -Qed. - -Lemma betafun_gt0 (a b : nat) : (0 < betafun a b :> R)%R. -Proof. -rewrite betafunE. -case: ifPn => [/andP[_ b0]|]; first by rewrite invr_gt0 ltr0n. -rewrite negb_and => /orP[a0|]. - case: ifPn => [/andP[_]|]; first by rewrite invr_gt0// ltr0n. - rewrite negb_and => /orP[b0|]. - by rewrite divr_gt0// ?mulr_gt0 ?ltr0n ?fact_gt0. - by rewrite -leqNgt leqn0 (negbTE a0). -rewrite -leqNgt leqn0 => /eqP ->. -rewrite eqxx/=. -case: ifPn; first by rewrite invr_gt0 ltr0n. -rewrite -leqNgt leqn0 => /eqP ->. -by rewrite fact0 mul1r divr1. -Qed. - -Lemma betafun_ge0 (a b : nat) : (0 <= betafun a b :> R)%R. -Proof. exact/ltW/betafun_gt0. Qed. - -Lemma betafun11 : betafun 1 1 = 1%R :> R. -Proof. by rewrite (betafun1S O) invr1. Qed. - -(* NB: this is not exactly betafun because EFin *) -Definition betafunEFin a b : \bar R := \int[mu]_x (XMonemX01 a b x)%:E. - -(* TODO: rev eq *) -Lemma betafunEFinT a b : - (betafunEFin a b = \int[mu]_(x in `[0%R, 1%R]) (XMonemX01 a b x)%:E)%E. -Proof. by rewrite /betafunEFin integral_XMonemX01/= setTI. Qed. - -Lemma betafunEFin_lty a b : (betafunEFin a b < +oo)%E. -Proof. -have := betafun_gt0 a b; rewrite /betafun /Rintegral /betafunEFin. -by case: (integral _ _ _) => [r _| |//]; rewrite ?ltxx ?ltry. -Qed. - -Lemma betafunEFin_fin_num a b : betafunEFin a b \is a fin_num. -Proof. -rewrite ge0_fin_numE ?betafunEFin_lty//. -by apply: integral_ge0 => //= x _; rewrite lee_fin XMonemX01_ge0. -Qed. - -Lemma betafunEFinE a b : (betafunEFin a b = (betafun a b)%:E :> \bar R)%E. -Proof. by rewrite -[LHS]fineK ?betafunEFin_fin_num. Qed. - -Lemma integrable_XMonemX01 a b : mu.-integrable setT (EFin \o XMonemX01 a b). -Proof. -apply/integrableP; split. - by apply/measurable_EFinP; exact: measurable_XMonemX01. -under eq_integral. - move=> /= x _. - rewrite ger0_norm//; last by rewrite XMonemX01_ge0. - over. -exact: betafunEFin_lty. -Qed. - -End betafun. - -(* normalized pdf for the beta distribution *) -Section beta_pdf. -Local Open Scope ring_scope. -Context {R : realType}. -Variables a b : nat. - -Definition beta_pdf t : R := XMonemX01 a b t / betafun a b. - -Lemma measurable_beta_pdf : measurable_fun setT beta_pdf. -Proof. by apply: measurable_funM => //; exact: measurable_XMonemX01. Qed. - -Lemma beta_pdf_ge0 t : 0 <= beta_pdf t. -Proof. -by rewrite /beta_pdf divr_ge0//; [exact: XMonemX01_ge0|exact: betafun_ge0]. -Qed. - -Lemma beta_pdf_le_betafunV x : beta_pdf x <= (betafun a b)^-1. -Proof. -rewrite /beta_pdf ler_pdivrMr ?betafun_gt0// mulVf ?gt_eqF ?betafun_gt0//. -exact: XMonemX01_le1. -Qed. - -Local Notation mu := lebesgue_measure. - -(* TODO: really useful? *) -Lemma int_beta_pdf01 : - (\int[mu]_(x in `[0%R, 1%R]) (beta_pdf x)%:E = - \int[mu]_x (beta_pdf x)%:E :> \bar R)%E. -Proof. -rewrite /beta_pdf. -under eq_integral do rewrite EFinM. -rewrite /=. -rewrite ge0_integralZr//=; last 3 first. - apply: measurable_funTS => /=; apply/measurable_EFinP => //. - exact: measurable_XMonemX01. - by move=> x _; rewrite lee_fin XMonemX01_ge0. - by rewrite lee_fin invr_ge0// betafun_ge0. -rewrite -betafunEFinT -ge0_integralZr//=. -- by apply/measurableT_comp => //; exact: measurable_XMonemX01. -- by move=> x _; rewrite lee_fin XMonemX01_ge0. -- by rewrite lee_fin invr_ge0// betafun_ge0. -Qed. - -Lemma integrable_beta_pdf : mu.-integrable setT (EFin \o beta_pdf). -Proof. -apply/integrableP; split. - by apply/measurable_EFinP; exact: measurable_beta_pdf. -under eq_integral. - move=> /= x _. - rewrite ger0_norm//; last by rewrite beta_pdf_ge0. - over. -rewrite -int_beta_pdf01. -apply: (@le_lt_trans _ _ (\int[mu]_(x in `[0%R, 1%R]) (betafun a b)^-1%:E)%E). - apply: ge0_le_integral => //=. - - by move=> x _; rewrite lee_fin beta_pdf_ge0. - - by apply/measurable_funTS/measurable_EFinP => /=; exact: measurable_beta_pdf. - - by move=> x _; rewrite lee_fin invr_ge0// betafun_ge0. - - by move=> x _; rewrite lee_fin beta_pdf_le_betafunV. -rewrite integral_cst//= lebesgue_measure_itv//=. -by rewrite lte01 oppr0 adde0 mule1 ltry. -Qed. - -Local Open Scope ring_scope. -Lemma bounded_beta_pdf_01 : - [bounded beta_pdf x : R^o | x in `[0%R, 1%R]%classic : set R]. -Proof. -exists (betafun a b)^-1; split; first by rewrite num_real. -move=> // y y1. -near=> M => /=. -rewrite (le_trans _ (ltW y1))//. -near: M. -move=> M /=. -rewrite in_itv/= => /andP[M0 M1]. -rewrite ler_norml; apply/andP; split. - rewrite lerNl (@le_trans _ _ 0%R)// ?invr_ge0 ?betafun_ge0//. - by rewrite lerNl oppr0 beta_pdf_ge0. -rewrite /beta_pdf ler_pdivrMr ?betafun_gt0//. -by rewrite mulVf ?XMonemX01_le1// gt_eqF// betafun_gt0. -Unshelve. all: by end_near. Qed. -Local Close Scope ring_scope. - -End beta_pdf. - -Section beta. -Local Open Scope ring_scope. -Context {R : realType}. -Variables a b : nat. - -Local Notation mu := (@lebesgue_measure R). - -Let beta_num (U : set (measurableTypeR R)) : \bar R := - \int[mu]_(x in U) (XMonemX01 a b x)%:E. - -Let beta_numT : beta_num setT = betafunEFin a b. -Proof. by rewrite /beta_num/= -/(betafunEFin a b) betafunEFinE. Qed. - -Let beta_num_lty U : measurable U -> (beta_num U < +oo)%E. -Proof. -move=> mU. -apply: (@le_lt_trans _ _ (\int[mu]_(x in U `&` `[0%R, 1%R]%classic) 1)%E); last first. - rewrite integral_cst//= ?mul1e//. - rewrite (le_lt_trans (measureIr _ _ _))//= lebesgue_measure_itv//= lte01//. - by rewrite EFinN sube0 ltry. - exact: measurableI. -rewrite /beta_num integral_XMonemX01 ge0_le_integral//=. -- exact: measurableI. -- by move=> x _; rewrite lee_fin XMonemX01_ge0. -- by apply/measurable_funTS/measurableT_comp => //; exact: measurable_XMonemX01. -- by move=> x _; rewrite lee_fin XMonemX01_le1. -Qed. - -Let beta_num0 : beta_num set0 = 0%:E. -Proof. by rewrite /beta_num integral_set0. Qed. - -Let beta_num_ge0 U : (0 <= beta_num U)%E. -Proof. -by rewrite /beta_num integral_ge0//= => x Ux; rewrite lee_fin XMonemX01_ge0. -Qed. - -Let beta_num_sigma_additive : semi_sigma_additive beta_num. -Proof. -move=> /= F mF tF mUF; rewrite /beta_num; apply: cvg_toP. - apply: ereal_nondecreasing_is_cvgn => m n mn. - apply: lee_sum_nneg_natr => // k _ _; apply: integral_ge0 => /= x Fkx. - by rewrite lee_fin; exact: XMonemX01_ge0. -rewrite ge0_integral_bigcup//=. -- by apply/measurable_funTS/measurableT_comp => //; exact: measurable_XMonemX01. -- by move=> x _; rewrite lee_fin XMonemX01_ge0. -Qed. - -HB.instance Definition _ := isMeasure.Build _ _ _ beta_num - beta_num0 beta_num_ge0 beta_num_sigma_additive. - -Definition beta_prob := - @mscale _ _ _ (invr_nonneg (NngNum (betafun_ge0 a b))) beta_num. - -HB.instance Definition _ := Measure.on beta_prob. - -Let beta_prob_setT : beta_prob setT = 1%:E. -Proof. -rewrite /beta_prob /= /mscale /= beta_numT betafunEFinE//. -by rewrite -EFinM mulVf// gt_eqF// betafun_gt0. -Qed. - -HB.instance Definition _ := - @Measure_isProbability.Build _ _ _ beta_prob beta_prob_setT. - -Lemma beta_prob01 : beta_prob `[0, 1] = 1%:E. -Proof. -rewrite /beta_prob /= /mscale/= /beta_num -betafunEFinT betafunEFinE//. -by rewrite -EFinM mulVf// gt_eqF// betafun_gt0. -Qed. - -Lemma beta_prob_fin_num U : measurable U -> beta_prob U \is a fin_num. -Proof. -move=> mU; rewrite ge0_fin_numE//. -rewrite /beta_prob/= /mscale/= /beta_num lte_mul_pinfty//. - by rewrite lee_fin// invr_ge0 betafun_ge0. -apply: (@le_lt_trans _ _ (betafunEFin a b)). - apply: ge0_subset_integral => //=. - by apply/measurable_EFinP; exact: measurable_XMonemX01. - by move=> x _; rewrite lee_fin XMonemX01_ge0. -by rewrite betafunEFin_lty. -Qed. - -Lemma integral_beta_pdf U : measurable U -> - (\int[mu]_(x in U) (beta_pdf a b x)%:E = beta_prob U :> \bar R)%E. -Proof. -move=> mU. -rewrite /beta_pdf. -under eq_integral do rewrite EFinM/=. -rewrite ge0_integralZr//=. -- by rewrite /beta_prob/= /mscale/= muleC. -- by apply/measurable_funTS/measurableT_comp => //; exact: measurable_XMonemX01. -- by move=> x _; rewrite lee_fin XMonemX01_ge0. -- by rewrite lee_fin invr_ge0// betafun_ge0. -Qed. - -End beta. -Arguments beta_prob {R}. - -Lemma integral_beta_prob_bernoulli_lty {R : realType} a b (f : R -> R) U : - measurable_fun setT f -> - (forall x, x \in `[0%R, 1%R] -> 0 <= f x <= 1)%R -> - (\int[beta_prob a b]_x `|bernoulli (f x) U| < +oo :> \bar R)%E. -Proof. -move=> mf /= f01. -apply: (@le_lt_trans _ _ (\int[beta_prob a b]_x cst 1 x))%E. - apply: ge0_le_integral => //=. - apply: measurableT_comp => //=. - by apply: (measurableT_comp (measurable_bernoulli2 _)). - by move=> x _; rewrite gee0_abs// probability_le1. -by rewrite integral_cst//= mul1e -ge0_fin_numE// beta_prob_fin_num. -Qed. - -Lemma integral_beta_prob_bernoulli_onemX_lty {R : realType} n a b U : - (\int[beta_prob a b]_x `|bernoulli (`1-x ^+ n) U| < +oo :> \bar R)%E. -Proof. -apply: integral_beta_prob_bernoulli_lty => //=. - by apply: measurable_funX => //; exact: measurable_funB. -move=> x; rewrite in_itv/= => /andP[x0 x1]. -rewrite exprn_ge0 ?subr_ge0//= exprn_ile1// ?subr_ge0//. -by rewrite lerBlDl -lerBlDr subrr. -Qed. - -Lemma integral_beta_prob_bernoulli_onem_lty {R : realType} n a b U : - (\int[beta_prob a b]_x `|bernoulli (1 - `1-x ^+ n) U| < +oo :> \bar R)%E. -Proof. -apply: integral_beta_prob_bernoulli_lty => //=. - apply: measurable_funB => //. - by apply: measurable_funX => //; exact: measurable_funB. -move=> x; rewrite in_itv/= => /andP[x0 x1]. -rewrite -lerBlDr opprK add0r. -rewrite andbC lerBlDl -lerBlDr subrr. -rewrite exprn_ge0 ?subr_ge0//= exprn_ile1// ?subr_ge0//. -by rewrite lerBlDl -lerBlDr subrr. -Qed. - -Local Open Scope ring_scope. - -Lemma beta_prob_uniform {R : realType} : beta_prob 1 1 = uniform_prob (@ltr01 R). -Proof. -apply/funext => U. -rewrite /beta_prob /uniform_prob. -rewrite /mscale/= betafun11 invr1 !mul1e. -rewrite integral_XMonemX01 integral_uniform_pdf. -under eq_integral. - move=> /= x. - rewrite inE => -[Ux/=]; rewrite in_itv/= => x10. - rewrite XMonemX01_11//=. - over. -rewrite /=. -under [RHS]eq_integral. - move=> /= x. - rewrite inE => -[Ux/=]; rewrite in_itv/= => x10. - rewrite /uniform_pdf x10 subr0 invr1. - over. -by []. -Qed. - -Lemma beta_prob_integrable {R :realType} a b a' b' : - (beta_prob a b).-integrable `[0, 1] - (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (XMonemX a' b' x)%:E). -Proof. -apply/integrableP; split. - by apply/measurableT_comp => //; exact: measurable_fun_XMonemX. -apply: (@le_lt_trans _ _ (\int[beta_prob a b]_(x in `[0%R, 1%R]) 1)%E). - apply: ge0_le_integral => //=. - do 2 apply/measurableT_comp => //. - exact: measurable_fun_XMonemX. - move=> x; rewrite in_itv/= => /andP[x0 x1]. - rewrite lee_fin ger0_norm; last first. - by rewrite !mulr_ge0// exprn_ge0// onem_ge0. - by rewrite mulr_ile1// ?exprn_ge0 ?onem_ge0// exprn_ile1// ?onem_ge0// onem_le1. -rewrite integral_cst//= mul1e. -by rewrite -ge0_fin_numE// beta_prob_fin_num. -Qed. - -Lemma beta_prob_integrable_onem {R : realType} a b a' b' : - (beta_prob a b).-integrable `[0, 1] - (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (`1-(XMonemX a' b' x))%:E). -Proof. -apply: (eq_integrable _ (cst 1 \- (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => - (XMonemX a' b' x)%:E))%E) => //. -apply: (@integrableB _ (g_sigma_algebraType R.-ocitv.-measurable)) => //=. - (* TODO: lemma? *) - apply/integrableP; split => //. - rewrite (eq_integral (fun x => (\1_setT x)%:E))/=; last first. - by move=> x _; rewrite /= indicT normr1. - rewrite integral_indic//= setTI /beta_prob /mscale/= lte_mul_pinfty//. - by rewrite lee_fin invr_ge0 betafun_ge0. - have /integrableP[_] := @integrable_XMonemX01 R a b. - under eq_integral. - move=> x _. - rewrite gee0_abs//; last first. - by rewrite lee_fin XMonemX01_ge0. - over. - by rewrite /= -/(betafunEFin a b) /= betafunEFinT. -exact: beta_prob_integrable. -Qed. - -Lemma beta_prob_integrable_dirac {R : realType} a b a' b' (c : bool) U : - (beta_prob a b).-integrable `[0, 1] - (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (XMonemX a' b' x)%:E * \d_c U)%E. -Proof. -apply: integrableMl => //=; last first. - exists 1; split => // x x1/= _ _; rewrite (le_trans _ (ltW x1))//. - by rewrite ger0_norm// indicE; case: (_ \in _). -exact: beta_prob_integrable. -Qed. - -Lemma beta_prob_integrable_onem_dirac {R : realType} a b a' b' (c : bool) U : - (beta_prob a b).-integrable `[0, 1] - (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => (`1-(XMonemX a' b' x))%:E * \d_c U)%E. -Proof. -apply: integrableMl => //=; last first. - exists 1; split => // x x1/= _ _; rewrite (le_trans _ (ltW x1))//. - by rewrite ger0_norm// indicE; case: (_ \in _). -exact: beta_prob_integrable_onem. -Qed. - -Local Close Scope ring_scope. - -Section integral_beta. -Context {R : realType}. -Local Notation mu := lebesgue_measure. - -Lemma beta_prob_dom a b : @beta_prob R a b `<< mu. -Proof. -move=> A mA muA0; rewrite /beta_prob /mscale/=. -apply/eqP; rewrite mule_eq0 eqe invr_eq0 gt_eqF/= ?betafun_gt0//; apply/eqP. -rewrite integral_XMonemX01; apply/eqP; rewrite eq_le; apply/andP; split; last first. - by apply: integral_ge0 => x _; rewrite lee_fin XMonemX01_ge0. -apply: (@le_trans _ _ (\int[mu]_(x in A `&` `[0%R, 1%R]%classic) 1)); last first. - rewrite integral_cst ?mul1e//=; last exact: measurableI. - by rewrite -[leRHS]muA0 measureIl. -apply: ge0_le_integral => //=; first exact: measurableI. -- by move=> x _; rewrite lee_fin XMonemX01_ge0. -- apply/measurable_funTS/measurableT_comp => //. - exact: measurable_XMonemX01. -- by move=> x _; rewrite lee_fin XMonemX01_le1. -Qed. - -Section beta_pdf_Beta. - -Local Open Scope charge_scope. - -(* beta_pdf is almost density function of Beta *) -Lemma beta_pdf_uniq_ae (a b : nat) : - ae_eq mu `[0%R, 1%R]%classic - ('d ((charge_of_finite_measure (@beta_prob R a b))) '/d mu) - (EFin \o (beta_pdf a b)). -Proof. -apply: integral_ae_eq => //. -- apply: (@integrableS _ _ _ _ setT) => //=. - apply: Radon_Nikodym_integrable => //=. - exact: beta_prob_dom. -- apply/measurable_funTS/measurableT_comp => //. - exact: measurable_beta_pdf. -- move=> E E01 mE. - rewrite integral_beta_pdf//. - apply/esym. - rewrite -Radon_Nikodym_integral//=. - exact: beta_prob_dom. -Qed. - -(* need to add lemma about radon-nikodym derivative of - lebesgue_stieltjes measure w.r.t. continuous density function *) -(*Lemma beta_pdf_uniq (a b : nat) : - {in `[0%R, 1%R]%classic, - ('d ((charge_of_finite_measure (@Beta R a b))) '/d mu) =1 - (EFin \o (beta_pdf a b))}. -Proof. Abort.*) - -End beta_pdf_Beta. - -Lemma integral_Beta a b f U : measurable U -> measurable_fun U f -> - \int[beta_prob a b]_(x in U) `|f x| < +oo -> - \int[beta_prob a b]_(x in U) f x = \int[mu]_(x in U) (f x * (beta_pdf a b x)%:E) :> \bar R. -Proof. -move=> mU mf finf. -rewrite -(Radon_Nikodym_change_of_variables (beta_prob_dom a b)) //=; last first. - by apply/integrableP; split. -apply: ae_eq_integral => //. -- apply: emeasurable_funM => //; apply: measurable_int. - apply: (integrableS _ _ (@subsetT _ _)) => //=. - by apply: (Radon_Nikodym_integrable _) => //=; exact: beta_prob_dom. -- apply: emeasurable_funM => //=; apply/measurableT_comp => //=. - by apply/measurable_funTS; exact: measurable_beta_pdf. -- apply: ae_eqe_mul2l => /=. - rewrite Radon_NikodymE//=; first exact: beta_prob_dom. - move=> ?. - case: cid => /= h [h1 h2 h3]. -(* uniqueness of Radon-Nikodym derivertive up to equal on non null sets of mu *) - apply: integral_ae_eq => //. - + apply: integrableS h2 => //. (* integrableST? *) - apply/measurable_funTS/measurableT_comp => //. - exact: measurable_beta_pdf. - + by move=> E E01 mE; rewrite -h3//= integral_beta_pdf. -Qed. - -End integral_beta. - -Section beta_prob_bernoulliE. -Context {R : realType}. -Local Notation mu := lebesgue_measure. Local Open Scope ring_scope. -Definition div_betafun a b c d : R := betafun (a + c) (b + d) / betafun a b. - -Lemma div_betafun_ge0 a b c d : 0 <= div_betafun a b c d. -Proof. by rewrite /div_betafun divr_ge0// betafun_ge0. Qed. - -Lemma div_betafun_le1 a b c d : (0 < a)%N -> (0 < b)%N -> - div_betafun a b c d <= 1. -Proof. -move=> a0 b0. -rewrite /div_betafun ler_pdivrMr// ?mul1r ?betafun_gt0//. -rewrite !betafunE. -rewrite addn_eq0 (gtn_eqF a0)/=. -rewrite addn_eq0 (gtn_eqF b0)/=. -rewrite ler_pdivrMr ?ltr0n ?fact_gt0//. -rewrite mulrAC. -rewrite ler_pdivlMr ?ltr0n ?fact_gt0//. -rewrite -!natrM ler_nat. -case: a a0 => //. -move=> n. -rewrite addSn. -case: b b0 => //. -move=> m. -rewrite [(n + c).+1.-1]/=. -rewrite [n.+1.-1]/=. -rewrite [m.+1.-1]/=. -rewrite addnS. -rewrite [(_ + m).+1.-1]/=. -rewrite (addSn m d). -rewrite [(m + _).+1.-1]/=. -rewrite (addSn (n + c)). -rewrite [_.+1.-1]/=. -rewrite addSn addnS. -by rewrite leq_fact2// leq_addr. -Qed. - -Definition beta_prob_bernoulli a b c d U : \bar R := - \int[beta_prob a b]_(y in `[0, 1]) bernoulli (XMonemX01 c.+1 d.+1 y) U. - -Lemma beta_prob_bernoulliE a b c d U : (a > 0)%N -> (b > 0)%N -> - beta_prob_bernoulli a b c d U = bernoulli (div_betafun a b c d) U. -Proof. -move=> a0 b0. -rewrite /beta_prob_bernoulli. -under eq_integral => x. - rewrite inE/= in_itv/= => x01. - rewrite bernoulliE/= ?XMonemX01_ge0 ?XMonemX01_le1//. - over. -rewrite /=. -rewrite [in RHS]bernoulliE/= ?div_betafun_ge0 ?div_betafun_le1//=. -under eq_integral => x x01. - rewrite /XMonemX01 patchE x01/=. - over. -rewrite /=. -rewrite integralD//=; last 2 first. - exact: beta_prob_integrable_dirac. - exact: beta_prob_integrable_onem_dirac. -congr (_ + _). - rewrite integralZr//=; last exact: beta_prob_integrable. - congr (_ * _)%E. - rewrite integral_Beta//; last 2 first. - by apply/measurableT_comp => //; exact: measurable_fun_XMonemX. - by have /integrableP[_] := @beta_prob_integrable R a b c d. - rewrite /beta_pdf. - under eq_integral do rewrite EFinM -muleA muleC -muleA. - rewrite /=. - transitivity ((betafun a b)^-1%:E * - \int[mu]_(x in `[0%R, 1%R]) (XMonemX01 (a + c) (b + d) x)%:E : \bar R)%E. - rewrite -integralZl//=; last first. - apply/integrableP; split. - apply/measurable_EFinP/measurable_funTS. - exact: measurable_XMonemX01. - under eq_integral. - move=> x x01. - rewrite gee0_abs; last by rewrite lee_fin XMonemX01_ge0. - over. - by rewrite /= -betafunEFinT betafunEFin_lty. - apply: eq_integral => x x01. - (* TODO: lemma? property of XMonemX? *) - rewrite muleA muleC muleA -(EFinM (x ^+ c)). - rewrite -/(XMonemX c d x) -EFinM mulrA XMonemX_XMonemX01//. - by rewrite -EFinM mulrC. - by rewrite -betafunEFinT betafunEFinE -EFinM mulrC. -under eq_integral do rewrite muleC. -rewrite /=. -rewrite integralZl//=; last exact: beta_prob_integrable_onem. -rewrite muleC; congr (_ * _)%E. -rewrite integral_Beta//=; last 2 first. - apply/measurableT_comp => //=. - by apply/measurable_funB => //; exact: measurable_fun_XMonemX. - by have /integrableP[] := @beta_prob_integrable_onem R a b c d. -rewrite /beta_pdf. -under eq_integral do rewrite EFinM muleA. -rewrite integralZr//=; last first. - apply: integrableMr => //=. - - by apply/measurable_funB => //=; exact: measurable_fun_XMonemX. - - apply/ex_bound => //. - + apply: (@globally_properfilter _ _ 0%R) => //=. - by apply: inferP; rewrite in_itv/= lexx ler01. - + exists 1 => t. - rewrite /= in_itv/= => t01. - apply: normr_onem; apply/andP; split. - by rewrite mulr_ge0// exprn_ge0// ?onem_ge0//; case/andP: t01. - by rewrite mulr_ile1// ?exprn_ge0 ?exprn_ile1// ?onem_ge0 ?onem_le1//; case/andP: t01. - - exact: integrableS (integrable_XMonemX01 _ _). -transitivity ( - (\int[mu]_(x in `[0%R, 1%R]) - ((XMonemX01 a b x)%:E - (XMonemX01 (a + c) (b + d) x)%:E) : \bar R) - * (betafun a b)^-1%:E)%E. - congr (_ * _)%E. - apply: eq_integral => x x01. - rewrite /onem -EFinM mulrBl mul1r EFinB; congr (_ - _)%E. - by rewrite XMonemX_XMonemX01. -rewrite integralB_EFin//=; last 2 first. - exact: integrableS (integrable_XMonemX01 _ _). - exact: integrableS (integrable_XMonemX01 _ _). -rewrite -!betafunEFinT !betafunEFinE. -rewrite -EFinM. -rewrite mulrBl /onem mulfV; last by rewrite gt_eqF// betafun_gt0. -by []. -Qed. - -End beta_prob_bernoulliE. - Declare Scope lang_scope. Delimit Scope lang_scope with P. @@ -1975,8 +821,8 @@ Inductive evalD : forall g t, exp D g t -> exp_var x H -D> acc_typ (map snd g) i ; measurable_acc_typ (map snd g) i | eval_bernoulli g e r mr : - e -D> r ; mr -> (exp_bernoulli e : exp D g _) -D> bernoulli \o r ; - measurableT_comp measurable_bernoulli mr + e -D> r ; mr -> (exp_bernoulli e : exp D g _) -D> bernoulli_prob \o r ; + measurableT_comp measurable_bernoulli_prob mr | eval_binomial g n e r mr : e -D> r ; mr -> (exp_binomial n e : exp D g _) -D> binomial_prob n \o r ; @@ -1991,8 +837,8 @@ Inductive evalD : forall g t, exp D g t -> | eval_poisson g n (e : exp D g _) f mf : e -D> f ; mf -> - exp_poisson n e -D> poisson_pdf n \o f ; - measurableT_comp (measurable_poisson_pdf n) mf + exp_poisson n e -D> poisson_pmf ^~ n \o f ; + measurableT_comp (measurable_poisson_pmf n measurableT) mf | eval_normalize g t (e : exp P g t) k : e -P> k -> @@ -2373,7 +1219,7 @@ all: rewrite {z g t}. by exists (snd \o f); eexists; exact: eval_proj2. - by move=> g x t tE; subst t; eexists; eexists; exact: eval_var. - move=> g e [p [mp H]]. - exists ((bernoulli : R -> pprobability bool R) \o p). + exists ((bernoulli_prob : R -> pprobability bool R) \o p). by eexists; exact: eval_bernoulli. - move=> g n e [p [mp H]]. exists ((binomial_prob n : R -> pprobability nat R) \o p). @@ -2381,7 +1227,7 @@ all: rewrite {z g t}. - by eexists; eexists; exact: eval_uniform. - by eexists; eexists; exact: eval_beta. - move=> g h e [f [mf H]]. - by exists (poisson_pdf h \o f); eexists; exact: eval_poisson. + by exists (poisson_pmf ^~ h \o f); eexists; exact: eval_poisson. - move=> g t e [k ek]. by exists (normalize_pt k); eexists; exact: eval_normalize. - move=> g t1 t2 x e1 [k1 ev1] e2 [k2 ev2]. @@ -2551,8 +1397,8 @@ Proof. by move=> i; apply/execD_evalD; exact: eval_var. Qed. Lemma execD_bernoulli g e : @execD g _ (exp_bernoulli e) = - existT _ ((bernoulli : R -> pprobability bool R) \o projT1 (execD e)) - (measurableT_comp measurable_bernoulli (projT2 (execD e))). + existT _ ((bernoulli_prob : R -> pprobability bool R) \o projT1 (execD e)) + (measurableT_comp measurable_bernoulli_prob (projT2 (execD e))). Proof. exact/execD_evalD/eval_bernoulli/evalD_execD. Qed. Lemma execD_binomial g n e : @@ -2579,8 +1425,9 @@ Proof. exact/execD_evalD/eval_normalize/evalP_execP. Qed. Lemma execD_poisson g n (e : exp D g Real) : execD (exp_poisson n e) = - existT _ (poisson_pdf n \o projT1 (execD e)) - (measurableT_comp (measurable_poisson_pdf n) (projT2 (execD e))). + existT _ (poisson_pmf ^~ n \o projT1 (execD e)) + (measurableT_comp (measurable_poisson_pmf n measurableT) + (projT2 (execD e))). Proof. exact/execD_evalD/eval_poisson/evalD_execD. Qed. Lemma execP_if g st e1 e2 e3 : @@ -2644,4 +1491,5 @@ f_equal. apply: eq_kernel => y V. exact: He. Qed. + Local Close Scope lang_scope. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index c72101c828..232f32c172 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -149,8 +149,8 @@ Proof. rewrite !execP_letin !execP_sample !execD_bernoulli !execP_return /=. rewrite execD_pair !exp_var'E (execD_var_erefl "x") (execD_var_erefl "y") /=. rewrite !execD_real//=. -do 2 (rewrite letin'E/= integral_bernoulli//=; last lra). -by rewrite letin'E/= integral_bernoulli//=; lra. +do 2 (rewrite letin'E/= integral_bernoulli_prob//=; last lra). +by rewrite letin'E/= integral_bernoulli_prob//=; lra. Qed. Lemma exec_sample_pair1213'_TandT : @@ -204,9 +204,9 @@ Proof. rewrite !execP_letin !execP_sample/= !execD_bernoulli execP_return /=. rewrite !(@execD_bin _ _ binop_and) !exp_var'E (execD_var_erefl "x"). rewrite (execD_var_erefl "y") /= !letin'E/= !execD_real/=. -rewrite integral_bernoulli//=; last lra. -rewrite !letin'E/= integral_bernoulli//=; last lra. -rewrite integral_bernoulli//=; last lra. +rewrite integral_bernoulli_prob//=; last lra. +rewrite !letin'E/= integral_bernoulli_prob//=; last lra. +rewrite integral_bernoulli_prob//=; last lra. rewrite /onem. rewrite muleDr// -addeA; congr (_ + _)%E. by rewrite !muleA; congr (_%:E); congr (_ * _); field. @@ -227,9 +227,9 @@ Proof. rewrite !execP_letin !execP_sample !execD_bernoulli !execP_return /=. rewrite !(@execD_bin _ _ binop_and) !exp_var'E (execD_var_erefl "x"). rewrite (execD_var_erefl "y") (execD_var_erefl "z") /= !execD_real/=. -do 3 (rewrite !letin'E/= integral_bernoulli//=; last lra). -do 2 (rewrite integral_bernoulli//=; last lra). -rewrite !letin'E/= integral_bernoulli//=; last lra. +do 3 (rewrite !letin'E/= integral_bernoulli_prob//=; last lra). +do 2 (rewrite integral_bernoulli_prob//=; last lra). +rewrite !letin'E/= integral_bernoulli_prob//=; last lra. rewrite !muleDr// -!addeA; congr (_ + _)%E. by rewrite !muleA; congr *%E; congr EFin; field. rewrite !muleA -!muleDl//; congr *%E; congr EFin. @@ -264,7 +264,7 @@ under eq_integral. under eq_integral do rewrite retE /=. over. rewrite /=. -rewrite integral_bernoulli//=; [|lra|by move=> b; rewrite integral_ge0]. +rewrite integral_bernoulli_prob//=; [|lra|by move=> b; rewrite integral_ge0]. rewrite iteE/= !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_indic//= !iteE/= /mscale/=. @@ -272,14 +272,14 @@ rewrite setTI !diracT !mule1. rewrite ger0_norm//. rewrite -EFinD/= eqe ifF; last first. by apply/negbTE/negP => /orP[/eqP|//]; rewrite /onem; lra. -rewrite integral_bernoulli//=; last lra. +rewrite integral_bernoulli_prob//=; last lra. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_dirac//= !diracT !mul1e ger0_norm//. rewrite exp_var'E (execD_var_erefl "x")/=. rewrite !indicT/= !mulr1. -rewrite bernoulliE//=; last lra. +rewrite bernoulli_probE//=; last lra. by rewrite muleDl//; congr (_ + _)%E; rewrite -!EFinM; congr (_%:E); rewrite !indicE /onem /=; case: (_ \in _); field. @@ -304,7 +304,7 @@ under eq_integral. rewrite !letin'E. under eq_integral do rewrite retE /=. over. -rewrite /= integral_bernoulli//=; [|lra|by move=> b; rewrite integral_ge0]. +rewrite /= integral_bernoulli_prob//=; [|lra|by move=> b; rewrite integral_ge0]. rewrite iteE/= !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_indic//= !iteE/= /mscale/=. @@ -313,13 +313,13 @@ rewrite ger0_norm//. rewrite -EFinD/= eqe ifF; last first. apply/negbTE/negP => /orP[/eqP|//]. by rewrite /onem; lra. -rewrite integral_bernoulli//=; last lra. +rewrite integral_bernoulli_prob//=; last lra. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_dirac//= !diracT !mul1e ger0_norm//. rewrite exp_var'E (execD_var_erefl "x")/=. -rewrite bernoulliE//=; last lra. +rewrite bernoulli_probE//=; last lra. rewrite !mul1r. rewrite muleDl//; congr (_ + _)%E; rewrite -!EFinM; @@ -348,7 +348,7 @@ under eq_integral. rewrite !letin'E. under eq_integral do rewrite retE /=. over. -rewrite /= integral_bernoulli//=; [|lra|by move=> b; rewrite integral_ge0]. +rewrite /= integral_bernoulli_prob//=; [|lra|by move=> b; rewrite integral_ge0]. rewrite iteE/= !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_cst//= !diracT !(mule1,mul1e). @@ -359,12 +359,12 @@ rewrite !diracT/= !mul1r. rewrite -EFinD/= eqe ifF; last first. apply/negbTE/negP => /orP[/eqP|//]. by rewrite /onem; lra. -rewrite integral_bernoulli//=; last lra. +rewrite integral_bernoulli_prob//=; last lra. rewrite !letin'E/= !iteE/=. rewrite !ge0_integral_mscale//=. rewrite ger0_norm//. rewrite !integral_dirac//= !diracT !mul1e ger0_norm//. -rewrite bernoulliE//=; last lra. +rewrite bernoulli_probE//=; last lra. rewrite !indicT. rewrite muleDl//; congr (_ + _)%E; rewrite -!EFinM; @@ -502,7 +502,7 @@ Proof. rewrite /guard 2!execP_letin execP_sample execD_bernoulli execD_real. rewrite execP_if/= !execP_return !exp_var'E !(execD_var_erefl "p") execD_unit. rewrite execP_score execD_real/=. -rewrite letin'E/= integral_bernoulli//=; last lra. +rewrite letin'E/= integral_bernoulli_prob//=; last lra. rewrite !letin'E !iteE/= integral_dirac// ge0_integral_mscale//=. by rewrite normr0 mul0e !mule0 !adde0 !diracT !mul1e. Qed. @@ -577,26 +577,29 @@ Local Notation mu := (@lebesgue_measure R). Lemma integrable_bernoulli_XMonemX01 a b U (mu : {measure set (g_sigma_algebraType R.-ocitv.-measurable) -> \bar R}) : measurable U -> (mu `[0%R, 1%R]%classic < +oo)%E -> - mu.-integrable `[0, 1] (fun x => bernoulli (XMonemX01 a b x) U). + mu.-integrable `[0, 1] (fun x => bernoulli_prob ((@XMonemX R a b) \_`[0,1] x) U). Proof. move=> mU mu01oo. apply/integrableP; split. - apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. - by apply: measurable_funTS; exact: measurable_XMonemX01. + apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. + apply/measurable_restrict => //=. + by rewrite setIidr//; exact: measurable_XMonemX. apply: (@le_lt_trans _ _ (\int[mu]_(x in `[0%R, 1%R]) cst 1 x)%E). apply: ge0_le_integral => //=. apply/measurable_funTS/measurableT_comp => //=. - apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. - exact: measurable_XMonemX01. + apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. + apply/measurable_restrict => //=. + by rewrite setTI//; exact: measurable_XMonemX. by move=> x _; rewrite gee0_abs// probability_le1. by rewrite integral_cst//= mul1e. Qed. Let measurable_bernoulli_XMonemX01 U : - measurable_fun setT (fun x : R => bernoulli (XMonemX01 2 1 x) U). + measurable_fun setT (fun x : R => bernoulli_prob ((@XMonemX R 1 0) \_ `[0,1] x) U). Proof. -apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. -exact: measurable_XMonemX01. +apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. +apply/measurable_restrict => //=. +by rewrite setTI//; exact: measurable_XMonemX. Qed. Lemma beta_bernoulli_bernoulli U : measurable U -> @@ -607,16 +610,16 @@ Proof. move=> mU. rewrite execP_letin !execP_sample execD_beta !execD_bernoulli/=. rewrite !execD_real/= exp_var'E (execD_var_erefl "p")/=. -transitivity (beta_prob_bernoulli 6 4 1 0 U : \bar R). - rewrite /beta_prob_bernoulli !letin'E/=. - rewrite integral_Beta//=; last 2 first. - exact: measurable_bernoulli2. - exact: integral_beta_prob_bernoulli_lty. - rewrite integral_Beta//=; last 2 first. +transitivity (beta_prob_bernoulli_prob 6 4 1 0 U : \bar R). + rewrite /beta_prob_bernoulli_prob !letin'E/=. + rewrite integral_beta_prob//=; last 2 first. + exact: measurable_bernoulli_prob2. + exact: integral_beta_prob_bernoulli_prob_lty. + rewrite integral_beta_prob//=; last 2 first. by apply: measurable_funTS => /=; exact: measurable_bernoulli_XMonemX01. - rewrite integral_Beta//=. + rewrite integral_beta_prob//=. + suff: mu.-integrable `[0%R, 1%R] - (fun x => bernoulli (XMonemX01 2 1 x) U * (beta_pdf 6 4 x)%:E)%E. + (fun x => bernoulli_prob ((@XMonemX R 1 0) \_`[0,1] x)%R U * (beta_pdf 6 4 x)%:E)%E. move=> /integrableP[_]. under eq_integral. move=> x _. @@ -635,18 +638,18 @@ transitivity (beta_prob_bernoulli 6 4 1 0 U : \bar R). + under eq_integral do rewrite gee0_abs//=. have : (beta_prob 6 4 `[0%R, 1%R] < +oo :> \bar R)%E. by rewrite -ge0_fin_numE// beta_prob_fin_num. - by move=> /(@integrable_bernoulli_XMonemX01 2 1 _ (beta_prob 6 4) mU) /integrableP[]. + by move=> /(@integrable_bernoulli_XMonemX01 1 0 _ (beta_prob 6 4) mU) /integrableP[]. rewrite [RHS]integral_mkcond. apply: eq_integral => x _ /=. rewrite patchE. case: ifPn => x01. - by rewrite /XMonemX01 patchE x01 XMonemX0' expr1. - by rewrite /beta_pdf /XMonemX01 patchE (negbTE x01) mul0r mule0. -rewrite beta_prob_bernoulliE// !bernoulliE//=; last 2 first. + by rewrite patchE x01 XMonemXn0 expr1. + by rewrite /beta_pdf patchE (negbTE x01) mul0r mule0. +rewrite beta_prob_bernoulli_probE// !bernoulli_probE//=; last 2 first. lra. - by rewrite div_betafun_ge0 div_betafun_le1. + by rewrite div_beta_fun_ge0 div_beta_fun_le1. by congr (_ * _ + _ * _)%:E; - rewrite /div_betafun/= /onem !betafunE/=; repeat rewrite !factE/=; field. + rewrite /div_beta_fun/= /onem !beta_funE/=; repeat rewrite !factE/=; field. Qed. End beta_bernoulli_bernoulli. @@ -707,13 +710,14 @@ Definition staton_bus_syntax0 : @exp R _ [::] _ := Definition staton_bus_syntax := [Normalize {staton_bus_syntax0}]. Let sample_bern : R.-sfker munit ~> mbool := - sample _ (measurableT_comp measurable_bernoulli (measurable_cst (2 / 7 : R)%R)). + sample _ (measurableT_comp measurable_bernoulli_prob (measurable_cst (2 / 7 : R)%R)). Let ite_3_10 : R.-sfker mbool * munit ~> measurableTypeR R := ite macc0of2 (@ret _ _ _ (measurableTypeR R) R _ (kr 3)) (@ret _ _ _ (measurableTypeR R) R _ (kr 10)). Let score_poisson4 : R.-sfker measurableTypeR R * (mbool * munit) ~> munit := - score (measurableT_comp (measurable_poisson_pdf 4) (@macc0of2 _ _ (measurableTypeR R) _)). + score (measurableT_comp (measurable_poisson_pmf 4 measurableT) + (@macc0of2 _ _ (measurableTypeR R) _)). Let kstaton_bus' := letin' sample_bern @@ -755,7 +759,7 @@ Lemma exec_staton_bus : execD staton_bus_syntax = existT _ (normalize_pt kstaton_bus') (measurable_normalize_pt _). Proof. by rewrite execD_normalize_pt exec_staton_bus0'. Qed. -Let poisson4 := @poisson_pdf R 4%N. +Let poisson4 := @poisson_pmf R ^~ 4%N. Let staton_bus_probability U := ((2 / 7)%:E * (poisson4 3)%:E * \d_true U + @@ -767,19 +771,19 @@ Proof. rewrite exec_staton_bus0' /staton_bus_probability /kstaton_bus'. rewrite /sample_bern. rewrite letin'E/=. -rewrite integral_bernoulli//=; last lra. +rewrite integral_bernoulli_prob//=; last lra. rewrite -!muleA; congr (_ * _ + _ * _)%E. - rewrite letin'_iteT//. rewrite letin'_retk//. rewrite letin'_kret//. rewrite /score_poisson4. - by rewrite /score/= /mscale/= ger0_norm//= poisson_pdf_ge0. + by rewrite /score/= /mscale/= ger0_norm//= poisson_pmf_ge0. - by rewrite onem27. - rewrite letin'_iteF//. rewrite letin'_retk//. rewrite letin'_kret//. rewrite /score_poisson4. - by rewrite /score/= /mscale/= ger0_norm//= poisson_pdf_ge0. + by rewrite /score/= /mscale/= ger0_norm//= poisson_pmf_ge0. Qed. End staton_bus. @@ -802,13 +806,14 @@ Definition staton_busA_syntax : exp _ [::] _ := [Normalize {staton_busA_syntax0}]. Let sample_bern : R.-sfker munit ~> mbool := - sample _ (measurableT_comp measurable_bernoulli (measurable_cst (2 / 7 : R)%R)). + sample _ (measurableT_comp measurable_bernoulli_prob (measurable_cst (2 / 7 : R)%R)). Let ite_3_10 : R.-sfker mbool * munit ~> measurableTypeR R := ite macc0of2 (@ret _ _ _ (measurableTypeR R) R _ (kr 3)) (@ret _ _ _ (measurableTypeR R) R _ (kr 10)). Let score_poisson4 : R.-sfker measurableTypeR R * (mbool * munit) ~> munit := - score (measurableT_comp (measurable_poisson_pdf 4) (@macc0of3' _ _ _ (measurableTypeR R) _ _)). + score (measurableT_comp (measurable_poisson_pmf 4 measurableT) + (@macc0of3' _ _ _ (measurableTypeR R) _ _)). (* same as kstaton_bus _ (measurable_poisson 4) but expressed with letin' instead of letin *) @@ -884,7 +889,7 @@ rewrite execP_return exp_var'E/= (execD_var_erefl "x") //=. by apply/eq_sfkernel => /= -[[] [a [b []]]] U0. Qed. -Let poisson4 := @poisson_pdf R 4%N. +Let poisson4 := @poisson_pmf R ^~ 4%N. Lemma exec_staton_busA0 U : execP staton_busA_syntax0 tt U = ((2 / 7%:R)%:E * (poisson4 3%:R)%:E * \d_true U + diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v index 26d61be498..ce241d96f0 100644 --- a/theories/lang_syntax_table_game.v +++ b/theories/lang_syntax_table_game.v @@ -150,20 +150,20 @@ End bounded. Lemma measurable_bernoulli_expn {R : realType} U n : measurable_fun [set: g_sigma_algebraType R.-ocitv.-measurable] - (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => bernoulli (`1-x ^+ n) U). + (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => bernoulli_prob (`1-x ^+ n) U). Proof. -apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. +apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. by apply: measurable_funX => //=; exact: measurable_funB. Qed. Lemma integrable_bernoulli_beta_pdf {R : realType} U : measurable U -> (@lebesgue_measure R).-integrable [set: g_sigma_algebraType R.-ocitv.-measurable] - (fun x => bernoulli (1 - `1-x ^+ 3) U * (beta_pdf 6 4 x)%:E)%E. + (fun x => bernoulli_prob (1 - `1-x ^+ 3) U * (beta_pdf 6 4 x)%:E)%E. Proof. move=> mU. have ? : measurable_fun [set: g_sigma_algebraType R.-ocitv.-measurable] - (fun x => bernoulli (1 - `1-x ^+ 3) U). - apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. + (fun x => bernoulli_prob (1 - `1-x ^+ 3) U). + apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. apply: measurable_funB => //; apply: measurable_funX => //. exact: measurable_funB. apply/integrableP; split => /=. @@ -186,15 +186,15 @@ apply: (@le_lt_trans _ _ (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (beta_pdf 6 by rewrite ge0_fin_numE// (le_lt_trans (probability_le1 _ _))// ltry. by rewrite lee_fin beta_pdf_ge0. by rewrite mule_ge0// lee_fin beta_pdf_ge0. - by rewrite /beta_pdf /XMonemX01 patchE (negbTE x01) mul0r mule0 abse0. + by rewrite /beta_pdf patchE (negbTE x01) mul0r mule0 abse0. apply: (@le_lt_trans _ _ - (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (betafun 6 4)^-1%:E)%E); last first. + (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (beta_fun 6 4)^-1%:E)%E); last first. by rewrite integral_cst//= lebesgue_measure_itv/= lte01 EFinN sube0 mule1 ltry. apply: ge0_le_integral => //=. - by move=> ? _; rewrite lee_fin beta_pdf_ge0. - by apply/measurable_funTS/measurableT_comp => //; exact: measurable_beta_pdf. -- by move=> ? _; rewrite lee_fin invr_ge0// betafun_ge0. -- by move=> x _; rewrite lee_fin beta_pdf_le_betafunV. +- by move=> ? _; rewrite lee_fin invr_ge0// beta_fun_ge0. +- by move=> x _; rewrite lee_fin beta_pdf_le_beta_funV. Qed. Section game_programs. @@ -389,9 +389,9 @@ Local Open Scope ereal_scope. Lemma measurable_bernoulli_onemXn {R : realType} U : measurable_fun [set: g_sigma_algebraType R.-ocitv.-measurable] - (fun x => bernoulli (1 - `1-x ^+ 3) U). + (fun x => bernoulli_prob (1 - `1-x ^+ 3) U). Proof. -apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. +apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. apply: measurable_funB => //. by apply: measurable_funX; exact: measurable_funB. Qed. @@ -402,31 +402,31 @@ Proof. exact/(bounded_norm _).1/boundedMl/bounded_XMonemX. Qed. Lemma integrable_bernoulli_XMonemX {R : realType} U : (beta_prob 1 1).-integrable [set: R] - (fun x => bernoulli (1 - `1-x ^+ 3) U * (normr (56 * XMonemX 5 3 x))%:E). + (fun x => bernoulli_prob (1 - `1-x ^+ 3) U * (normr (56 * XMonemX 5 3 x))%:E). Proof. apply/integrableP; split. apply: emeasurable_funM; first exact: measurable_bernoulli_onemXn. apply/measurable_EFinP => //; apply: measurableT_comp => //. - by apply: measurable_funM => //; exact: measurable_fun_XMonemX. + by apply: measurable_funM => //; exact: measurable_XMonemX. rewrite beta_prob_uniform integral_uniform//=. rewrite subr0 invr1 mul1e. suff : lebesgue_measure.-integrable `[0%R, 1%R] - (fun y : R => bernoulli (1 - `1-y ^+ 3) U * (normr (56 * XMonemX 5 3 y))%:E). + (fun y : R => bernoulli_prob (1 - `1-y ^+ 3) U * (normr (56 * XMonemX 5 3 y))%:E). by move=> /integrableP[]. apply: integrableMl => //=. - apply/integrableP; split. by apply: measurable_funTS; exact: measurable_bernoulli_onemXn. - have := @integral_beta_prob_bernoulli_onem_lty R 3 1%N 1%N U. + have := @integral_beta_prob_bernoulli_prob_onem_lty R 3 1%N 1%N U. rewrite beta_prob_uniform integral_uniform//=; last first. by apply: measurableT_comp => //=; exact: measurable_bernoulli_onemXn. by rewrite subr0 invr1 mul1e. - apply: @measurableT_comp => //=; apply: measurable_funM => //. - exact: measurable_fun_XMonemX. + exact: measurable_XMonemX. exact: bounded_norm_XnonemXn. apply: @measurableT_comp => //; apply: emeasurable_funM => //=. exact: measurable_bernoulli_onemXn. do 2 apply: @measurableT_comp => //=. -by apply: measurable_funM => //; exact: measurable_fun_XMonemX. +by apply: measurable_funM => //; exact: measurable_XMonemX. Qed. Section from_prog2_to_prog3. @@ -489,25 +489,25 @@ under [in RHS]eq_integral => y _. rewrite /=. rewrite [RHS]ge0_integral_mscale//=; last first. by move=> _ _; rewrite integral_ge0. -rewrite integral_Beta//=; last 2 first. +rewrite integral_beta_prob//=; last 2 first. - apply: emeasurable_funM => //=. exact: measurable_bernoulli_onemXn. apply/measurable_EFinP; apply: measurableT_comp => //. - by apply: measurable_funM => //; exact: measurable_fun_XMonemX. + by apply: measurable_funM => //; exact: measurable_XMonemX. - by have /integrableP[] := @integrable_bernoulli_XMonemX R U. rewrite ger0_norm// integral_dirac// diracT mul1e. -rewrite integral_Beta/=; [|by []|exact: measurable_bernoulli_onemXn - |exact: integral_beta_prob_bernoulli_onem_lty]. +rewrite integral_beta_prob/=; [|by []|exact: measurable_bernoulli_onemXn + |exact: integral_beta_prob_bernoulli_prob_onem_lty]. rewrite -integralZl//=; last exact: integrable_bernoulli_beta_pdf. apply: eq_integral => y _. rewrite [in RHS]muleCA -[in LHS]muleA; congr *%E. -rewrite /beta_pdf /XMonemX01 2!patchE; case: ifPn => [y01|_]; last first. +rewrite /beta_pdf 2!patchE; case: ifPn => [y01|_]; last first. by rewrite !mul0r 2!mule0. rewrite ger0_norm; last first. by rewrite mulr_ge0// XMonemX_ge0//; rewrite inE in y01. rewrite [X in _ = _ * X]EFinM [in RHS]muleCA. rewrite /= XMonemX00 mul1r [in LHS](mulrC 56%R) [in LHS]EFinM -[in LHS]muleA; congr *%E. -by rewrite !betafunE/=; repeat rewrite !factE/=; rewrite -EFinM; congr EFin; lra. +by rewrite !beta_funE/=; repeat rewrite !factE/=; rewrite -EFinM; congr EFin; lra. Qed. End from_prog2_to_prog3. @@ -517,23 +517,23 @@ Local Open Scope ereal_scope. Lemma int_beta_prob01 {R : realType} (f : R -> R) a b U : measurable_fun [set: R] f -> (forall x, x \in `[0%R, 1%R] -> 0 <= f x <= 1)%R -> - \int[beta_prob a b]_y bernoulli (f y) U = - \int[beta_prob a b]_(y in `[0%R, 1%R] : set R) bernoulli (f y) U. + \int[beta_prob a b]_y bernoulli_prob (f y) U = + \int[beta_prob a b]_(y in `[0%R, 1%R] : set R) bernoulli_prob (f y) U. Proof. move=> mf f01. -rewrite [LHS]integral_Beta//=; last 2 first. +rewrite [LHS]integral_beta_prob//=; last 2 first. apply: measurable_funTS. - by apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. - exact: integral_beta_prob_bernoulli_lty. -rewrite [RHS]integral_Beta//; last 2 first. + by apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. + exact: integral_beta_prob_bernoulli_prob_lty. +rewrite [RHS]integral_beta_prob//; last 2 first. apply/measurable_funTS => //=. - by apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. - apply: (le_lt_trans _ (lang_syntax.integral_beta_prob_bernoulli_lty a b U mf f01)). + by apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. + apply: (le_lt_trans _ (integral_beta_prob_bernoulli_prob_lty a b U mf f01)). apply: ge0_subset_integral => //=; apply: measurableT_comp => //=. - by apply: (measurableT_comp (measurable_bernoulli2 _)) => //=. + by apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. rewrite [RHS]integral_mkcond/=; apply: eq_integral => x _ /=. rewrite !patchE; case: ifPn => // x01. -by rewrite /beta_pdf /XMonemX01 patchE (negbTE x01) mul0r mule0. +by rewrite /beta_pdf patchE (negbTE x01) mul0r mule0. Qed. Lemma expr_onem_01 {R : realType} (x : R) : x \in `[0%R, 1%R] -> @@ -545,19 +545,19 @@ by rewrite lerBlDl -lerBlDr subrr. Qed. Lemma int_beta_prob_bernoulli {R : realType} (U : set (@mtyp R Bool)) : - \int[beta_prob 6 4]_y bernoulli (`1-y ^+ 3) U = bernoulli (1 / 11) U :> \bar R. + \int[beta_prob 6 4]_y bernoulli_prob (`1-y ^+ 3) U = bernoulli_prob (1 / 11) U :> \bar R. Proof. rewrite int_beta_prob01//; last 2 first. by apply: measurable_funX => //; exact: measurable_funB. exact: expr_onem_01. -have := @beta_prob_bernoulliE R 6 4 0 3 U isT isT. -rewrite /beta_prob_bernoulli. +have := @beta_prob_bernoulli_probE R 6 4 0 3 U isT isT. +rewrite /beta_prob_bernoulli_prob. under eq_integral. move=> x x0. - rewrite /XMonemX01 patchE x0 XMonemX0. + rewrite patchE x0 XMonemX0n. over. -rewrite /= => ->; congr bernoulli. -by rewrite /div_betafun addn0 !betafunE/=; repeat rewrite !factE/=; field. +rewrite /= => ->; congr bernoulli_prob. +by rewrite /div_beta_fun addn0 !beta_funE/=; repeat rewrite !factE/=; field. Qed. Lemma dirac_bool {R : realType} (U : set bool) : @@ -572,10 +572,10 @@ have [| | |] := set_bool U => /eqP ->; rewrite !diracE. Qed. Lemma int_beta_prob_bernoulli_onem {R : realType} (U : set (@mtyp R Bool)) : - \int[beta_prob 6 4]_y bernoulli (`1-(`1-y ^+ 3)) U = bernoulli (10 / 11) U :> \bar R. + \int[beta_prob 6 4]_y bernoulli_prob (`1-(`1-y ^+ 3)) U = bernoulli_prob (10 / 11) U :> \bar R. Proof. -transitivity (\d_false U + \d_true U - bernoulli (1 / 11) U : \bar R)%E; last first. - rewrite /bernoulli ifT; last lra. +transitivity (\d_false U + \d_true U - bernoulli_prob (1 / 11) U : \bar R)%E; last first. + rewrite /bernoulli_prob ifT; last lra. rewrite ifT; last lra. apply/eqP; rewrite sube_eq//; last first. rewrite ge0_adde_def// inE. @@ -600,18 +600,18 @@ rewrite [X in _ == _ + X]int_beta_prob01; last 2 first. by apply: measurable_funX => //; exact: measurable_funB. exact: expr_onem_01. rewrite -ge0_integralD//=; last 2 first. - apply: (@measurableT_comp _ _ _ _ _ _ (bernoulli ^~ U)) => /=. - exact: measurable_bernoulli2. + apply: (@measurableT_comp _ _ _ _ _ _ (bernoulli_prob ^~ U)) => /=. + exact: measurable_bernoulli_prob2. apply: measurable_funB => //=; apply: measurable_funX => //=. exact: measurable_funB. - apply: (@measurableT_comp _ _ _ _ _ _ (bernoulli ^~ U)) => /=. - exact: measurable_bernoulli2. + apply: (@measurableT_comp _ _ _ _ _ _ (bernoulli_prob ^~ U)) => /=. + exact: measurable_bernoulli_prob2. by apply: measurable_funX => //=; exact: measurable_funB. apply/eqP; transitivity (\int[beta_prob 6 4]_(x in `[0%R, 1%R]) (\d_false U + \d_true U) : \bar R). by rewrite integral_cst//= beta_prob01 mule1 EFinD. apply: eq_integral => /= x x01. -rewrite /bernoulli subr_ge0 lerBlDr -lerBlDl subrr andbC. +rewrite /bernoulli_prob subr_ge0 lerBlDr -lerBlDl subrr andbC. rewrite (_ : (_ <= _ <= _)%R); last first. by apply: expr_onem_01; rewrite inE in x01. rewrite -fsbig_split//=. @@ -689,12 +689,12 @@ rewrite !execP_sample !execD_bernoulli !execD_real/=. apply: funext=> x. apply: eq_probability=> /= U. rewrite !normalizeE/=. -rewrite !bernoulliE//=; [|lra..]. +rewrite !bernoulli_probE//=; [|lra..]. rewrite !diracT !mule1 -EFinD add_onemK onee_eq0/=. rewrite !letin'E. under eq_integral. move=> A _ /=. - rewrite !bernoulliE//=; [|lra..]. + rewrite !bernoulli_probE//=; [|lra..]. rewrite !diracT !mule1 -EFinD add_onemK. over. rewrite !ge0_integral_mscale//= (ger0_norm (ltW p0))//. @@ -702,7 +702,7 @@ rewrite integral_dirac// !diracT !indicT /= !mule1 !mulr1. rewrite add_onemK invr1 mule1. rewrite gt_eqF ?lte_fin//=. rewrite integral_dirac//= diracT mul1e. -by rewrite muleAC -EFinM divff ?gt_eqF// mul1e bernoulliE. +by rewrite muleAC -EFinM divff ?gt_eqF// mul1e bernoulli_probE. Qed. Lemma prog45 : execD (@prog4 R) = execD (@prog5 R). diff --git a/theories/prob_lang.v b/theories/prob_lang.v index 0647e16482..f2639adff6 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -18,8 +18,6 @@ From mathcomp Require Import lebesgue_integral trigo probability kernel charge. (* ESOP 2017 *) (* *) (* ``` *) -(* poisson_pdf == Poisson pdf *) -(* exponential_pdf == exponential distribution pdf *) (* measurable_sum X Y == the type X + Y, as a measurable type *) (* ``` *) (* *) @@ -113,61 +111,26 @@ subst p2. by f_equal. Qed. -(* NB: to be PRed to probability.v *) -Section poisson_pdf. -Variable R : realType. -Local Open Scope ring_scope. - -(* density function for Poisson *) -Definition poisson_pdf k r : R := - if r > 0 then r ^+ k / k`!%:R^-1 * expR (- r) else 1%:R. - -Lemma poisson_pdf_ge0 k r : 0 <= poisson_pdf k r. -Proof. -rewrite /poisson_pdf; case: ifPn => r0//. -by rewrite mulr_ge0 ?expR_ge0// mulr_ge0// exprn_ge0 ?ltW. -Qed. +Definition poisson3 {R : realType} := @poisson_pmf R 3%:R 4. (* 0.168 *) +Definition poisson10 {R : realType} := @poisson_pmf R 10%:R 4. (* 0.019 *) -Lemma poisson_pdf_gt0 k r : 0 < r -> 0 < poisson_pdf k.+1 r. +(* TODO: move *) +Lemma poisson_pmf_gt0 {R : realType} k (r : R) : + (0 < r -> 0 < poisson_pmf r k.+1)%R. Proof. -move=> r0; rewrite /poisson_pdf r0 mulr_gt0 ?expR_gt0//. +move=> r0; rewrite /poisson_pmf r0 mulr_gt0 ?expR_gt0//. by rewrite divr_gt0// ?exprn_gt0// invr_gt0 ltr0n fact_gt0. Qed. -Lemma measurable_poisson_pdf k : measurable_fun setT (poisson_pdf k). -Proof. -rewrite /poisson_pdf; apply: measurable_fun_if => //. - exact: measurable_fun_ltr. -by apply: measurable_funM => /=; - [exact: measurable_funM|exact: measurableT_comp]. -Qed. - -Definition poisson3 := poisson_pdf 4 3%:R. (* 0.168 *) -Definition poisson10 := poisson_pdf 4 10%:R. (* 0.019 *) - -End poisson_pdf. - -Section exponential_pdf. -Variable R : realType. -Local Open Scope ring_scope. - -(* density function for exponential *) -Definition exponential_pdf x r : R := r * expR (- r * x). - -Lemma exponential_pdf_gt0 x r : 0 < r -> 0 < exponential_pdf x r. -Proof. by move=> r0; rewrite /exponential_pdf mulr_gt0// expR_gt0. Qed. - -Lemma exponential_pdf_ge0 x r : 0 <= r -> 0 <= exponential_pdf x r. -Proof. by move=> r0; rewrite /exponential_pdf mulr_ge0// expR_ge0. Qed. - -Lemma measurable_exponential_pdf x : measurable_fun setT (exponential_pdf x). +Lemma exponential_pdf_gt0 {R : realType} (r : R) x : + (0 < r -> 0 < x -> 0 < exponential_pdf r x)%R. Proof. -apply: measurable_funM => //=; apply: measurableT_comp => //. -exact: measurable_funM. +move=> r0 x0; rewrite /exponential_pdf/=. +rewrite patchE/= ifT; last first. + by rewrite inE/= in_itv/= (ltW x0). +by rewrite mulr_gt0// expR_gt0. Qed. -End exponential_pdf. - (* X + Y is a measurableType if X and Y are *) HB.instance Definition _ (X Y : pointedType) := isPointed.Build (X + Y)%type (@inl X Y point). @@ -412,7 +375,7 @@ rewrite big1 ?adde0//= => j jk. rewrite ifF// lte_fin lee_fin. move: jk; rewrite neq_ltn/= => /orP[|] jr. - suff : (j.+1%:R <= `|f t|)%R by rewrite leNgt => /negbTE ->; rewrite andbF. - rewrite (_ : j.+1%:R = j.+1%:~R)// floor_ge_int//. + rewrite (_ : j.+1%:R = j.+1%:~R)// -floor_ge_int_tmp//. move: jr; rewrite -lez_nat => /le_trans; apply. by rewrite -[leRHS](@ger0_norm _ (floor `|f t|)) ?floor_ge0. - suff : (`|f t| < j%:R)%R by rewrite ltNge => /negbTE ->. @@ -1481,12 +1444,13 @@ Qed. (* hard constraints to express score below 1 *) Lemma score_fail (r : R) : (0 <= r <= 1)%R -> score (kr r) = - letin (sample_cst (bernoulli r) : R.-pker T ~> _) + letin (sample_cst (bernoulli_prob r) : R.-pker T ~> _) (ite (@macc1of2 _ _ _ _) (ret ktt) fail). Proof. move=> /andP[r0 r1]; apply/eq_sfkernel => x U. rewrite letinE/= /sample; unlock. -by rewrite /mscale/= ger0_norm// integral_bernoulli ?r0//= 2!iteE//= failE mule0 adde0. +rewrite /mscale/= ger0_norm// integral_bernoulli_prob ?r0//=. +by rewrite 2!iteE//= failE mule0 adde0. Qed. End insn1_lemmas. @@ -1612,9 +1576,9 @@ Lemma letin_sample_bernoulli d d' (T : measurableType d) (T' : measurableType d') (R : realType) (r : R) (u : R.-sfker [the measurableType _ of (T * bool)%type] ~> T') x y : (0 <= r <= 1)%R -> - letin (sample_cst (bernoulli r)) u x y = + letin (sample_cst (bernoulli_prob r)) u x y = r%:E * u (x, true) y + (`1- r)%:E * u (x, false) y. -Proof. by move=> r01; rewrite letinE/= integral_bernoulli. Qed. +Proof. by move=> r01; rewrite letinE/= integral_bernoulli_prob. Qed. Section sample_and_return. Import Notations. @@ -1622,7 +1586,7 @@ Context d (T : measurableType d) (R : realType). Definition sample_and_return : R.-sfker T ~> _ := letin - (sample_cst (bernoulli (2 / 7))) (* T -> B *) + (sample_cst (bernoulli_prob (2 / 7))) (* T -> B *) (ret macc1of2) (* T * B -> B *). Lemma sample_and_returnE t U : sample_and_return t U = @@ -1643,7 +1607,7 @@ Context d (T : measurableType d) (R : realType). return r *) Definition sample_and_branch : R.-sfker T ~> _ := letin - (sample_cst (bernoulli (2 / 7))) (* T -> B *) + (sample_cst (bernoulli_prob (2 / 7))) (* T -> B *) (ite macc1of2 (ret (@k3 _ _ R)) (ret k10)). Lemma sample_and_branchE t U : sample_and_branch t U = @@ -1660,12 +1624,12 @@ Context d (T : measurableType d) (R : realType). Import Notations. Definition bernoulli_and : R.-sfker T ~> mbool := - (letin (sample_cst (bernoulli (1 / 2))) - (letin (sample_cst (bernoulli (1 / 2))) + (letin (sample_cst (bernoulli_prob (1 / 2))) + (letin (sample_cst (bernoulli_prob (1 / 2))) (ret (measurable_and macc1of3 macc2of3)))). Lemma bernoulli_andE t U : - bernoulli_and t U = sample_cst (bernoulli (1 / 4)) t U. + bernoulli_and t U = sample_cst (bernoulli_prob (1 / 4)) t U. Proof. rewrite /bernoulli_and. rewrite letin_sample_bernoulli; last lra. @@ -1675,7 +1639,7 @@ rewrite muleDr//= -muleDl//. rewrite !muleA -addeA -muleDl// -!EFinM !onem1S/= -splitr mulr1. have -> : (1 / 2 * (1 / 2) = 1 / 4%:R :> R)%R by rewrite mulf_div mulr1// -natrM. rewrite [in RHS](_ : 1 / 4 = (1 / 4)%:nng%:num)%R//. -rewrite bernoulliE/=; last lra. +rewrite bernoulli_probE/=; last lra. rewrite -!EFinM; congr( _ + (_ * _)%:E). by rewrite /onem; lra. Qed. @@ -1687,7 +1651,7 @@ Import Notations. Context d (T : measurableType d) (R : realType) (h : R -> R). Hypothesis mh : measurable_fun setT h. Definition kstaton_bus : R.-sfker T ~> mbool := - letin (sample_cst (bernoulli (2 / 7))) + letin (sample_cst (bernoulli_prob (2 / 7))) (letin (letin (ite macc1of2 (ret k3) (ret k10)) (score (measurableT_comp mh macc2of3))) @@ -1704,8 +1668,8 @@ End staton_bus. Section staton_bus_poisson. Import Notations. Context d (T : measurableType d) (R : realType). -Let poisson4 := @poisson_pdf R 4%N. -Let mpoisson4 := @measurable_poisson_pdf R 4%N. +Let poisson4 r := @poisson_pmf R r 4%N. +Let mpoisson4 := @measurable_poisson_pmf R setT 4%N measurableT. Definition kstaton_bus_poisson : R.-sfker R ~> mbool := kstaton_bus _ mpoisson4. @@ -1720,12 +1684,12 @@ rewrite -!muleA; congr (_ * _ + _ * _). - rewrite letin_kret//. rewrite letin_iteT//. rewrite letin_retk//. - by rewrite scoreE//= => r r0; exact: poisson_pdf_ge0. + by rewrite scoreE//= => r r0; exact: poisson_pmf_ge0. - by rewrite onem27. rewrite letin_kret//. rewrite letin_iteF//. rewrite letin_retk//. - by rewrite scoreE//= => r r0; exact: poisson_pdf_ge0. + by rewrite scoreE//= => r r0; exact: poisson_pmf_ge0. Qed. (* true -> 2/7 * 0.168 = 2/7 * 3^4 e^-3 / 4! *) @@ -1740,7 +1704,7 @@ Lemma staton_busE P (t : R) U : Proof. rewrite /staton_bus normalizeE !kstaton_bus_poissonE !diracT !mule1 ifF //. apply/negbTE; rewrite gt_eqF// lte_fin. -by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// poisson_pdf_gt0// ltr0n. +by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// poisson_pmf_gt0// ltr0n. Qed. End staton_bus_poisson. @@ -1770,12 +1734,14 @@ rewrite -!muleA; congr (_ * _ + _ * _). - rewrite letin_kret//. rewrite letin_iteT//. rewrite letin_retk//. - rewrite scoreE//= => r r0; exact: exponential_pdf_ge0. + rewrite scoreE//= => r r0; apply: exponential_pdf_ge0 => //. + by rewrite ler0q; lra. - by rewrite onem27. rewrite letin_kret//. rewrite letin_iteF//. rewrite letin_retk//. - by rewrite scoreE//= => r r0; exact: exponential_pdf_ge0. + rewrite scoreE//= => r r0; apply: exponential_pdf_ge0. + by rewrite ler0q; lra. Qed. (* true -> 5/7 * 0.019 = 5/7 * 10^4 e^-10 / 4! *) @@ -1791,12 +1757,12 @@ Proof. rewrite /staton_bus. rewrite normalizeE /= !kstaton_bus_exponentialE !diracT !mule1 ifF //. apply/negbTE; rewrite gt_eqF// lte_fin. -by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n// exponential_pdf_gt0 ?ltr0n. +by rewrite addr_gt0// mulr_gt0//= ?divr_gt0// ?ltr0n//; + rewrite exponential_pdf_gt0 ?ltr0n// ltr0q; lra. Qed. End staton_bus_exponential. - Section von_neumann_trick. Context d {T : measurableType d} {R : realType}. @@ -1845,12 +1811,12 @@ Lemma trickE gamma X : trick gamma X = (r *+ (inl tt \in X) + q *+ ((inr true \in X) + (inr false \in X)))%:E. Proof. -have Dbernoulli : D =1 bernoulli p by exact/eq_bernoulli/Dtrue. +have Dbernoulli : D =1 bernoulli_prob p by exact/eq_bernoulli/Dtrue. have p_itv01 : (0 <= p <= 1)%R. by rewrite -2!lee_fin -Dtrue?measure_ge0 ?probability_le1. pose eqbern := eq_measure_integral _ (fun x _ _ => Dbernoulli x). rewrite /trick/= /kcomp. -do 2?rewrite ?eqbern ?integral_bernoulli//= /kcomp/=. +do 2?rewrite ?eqbern ?integral_bernoulli_prob//= /kcomp/=. rewrite !integral_dirac ?diracT//= ?mul1e. rewrite !iteE//= ?diracE/= /kcomp/=. rewrite !integral_dirac /acc1of4/= ?diracT ?diracE ?mul1e//. @@ -1911,7 +1877,8 @@ HB.instance Definition _ gamma := Measure_isProbability.Build _ _ _ HB.instance Definition _ := Kernel_isProbability.Build _ _ _ _ _ kvon_neumann_trick von_neumann_trick_prob_kernelT. -Theorem von_neumann_trickP gamma : von_neumann_trick gamma =1 bernoulli 2^-1. +Theorem von_neumann_trickP gamma : + von_neumann_trick gamma =1 bernoulli_prob 2^-1. Proof. by apply: eq_bernoulli; rewrite von_neumann_trick_prob_kernel. Qed. End von_neumann_trick_proof. @@ -2113,7 +2080,7 @@ End letin'A. Lemma letin'_sample_bernoulli d d' (T : measurableType d) (T' : measurableType d') (R : realType) (r : R) (r01 : (0 <= r <= 1)%R) (u : R.-sfker bool * T ~> T') x y : - letin' (sample_cst (bernoulli r)) u x y = + letin' (sample_cst (bernoulli_prob r)) u x y = r%:E * u (true, x) y + (`1- r)%:E * u (false, x) y. Proof. by rewrite letin'_letin letin_sample_bernoulli. Qed. @@ -2175,12 +2142,12 @@ Arguments fail' {d d' X Y R}. Lemma score_fail' d (X : measurableType d) {R : realType} (r : R) (r01 : (0 <= r <= 1)%R) : score (kr r) = - letin' (sample_cst (bernoulli r) : R.-pker X ~> _) + letin' (sample_cst (bernoulli_prob r) : R.-pker X ~> _) (ite macc0of2 (ret ktt) fail'). Proof. move: r01 => /andP[r0 r1]; apply/eq_sfkernel => x U. rewrite letin'E/= /sample; unlock. -rewrite integral_bernoulli ?r0//=. +rewrite integral_bernoulli_prob ?r0//=. by rewrite /mscale/= iteE//= iteE//= fail'E mule0 adde0 ger0_norm. Qed. @@ -2194,7 +2161,7 @@ Definition gauss_pdf := @normal_pdf R 0 1. Lemma normal_pdf_gt0 m s x : 0 < s -> 0 < normal_pdf m s x :> R. Proof. move=> s0; rewrite /normal_pdf gt_eqF// mulr_gt0 ?expR_gt0// invr_gt0. -by rewrite sqrtr_gt0 pmulrn_rgt0// mulr_gt0 ?pi_gt0 ?exprn_gt0. +by rewrite sqrtr_gt0 mulrn_wgt0// mulr_gt0 ?pi_gt0 ?exprn_gt0. Qed. Lemma gauss_pdf_gt0 x : 0 < gauss_pdf x. diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 537236ed40..2b422b33cb 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -28,19 +28,19 @@ Variable R : realType. Local Open Scope ring_scope. Notation mu := (@lebesgue_measure R). Hypothesis integral_poisson_density : forall k, - (\int[mu]_x (@poisson_pdf R k x)%:E = 1%E)%E. + (\int[mu]_x (@poisson_pmf R x k)%:E = 1%E)%E. (* density function for poisson *) -Definition poisson1 := @poisson_pdf R 1%N. +Definition poisson1 r := @poisson_pmf R r 1%N. Lemma poisson1_ge0 (x : R) : 0 <= poisson1 x. -Proof. exact: poisson_pdf_ge0. Qed. +Proof. exact: poisson_pmf_ge0. Qed. Definition mpoisson1 (V : set R) : \bar R := (\int[lebesgue_measure]_(x in V) (poisson1 x)%:E)%E. Lemma measurable_fun_poisson1 : measurable_fun setT poisson1. -Proof. exact: measurable_poisson_pdf. Qed. +Proof. exact: measurable_poisson_pmf. Qed. Let mpoisson10 : mpoisson1 set0 = 0%E. Proof. by rewrite /mpoisson1 integral_set0. Qed. @@ -48,7 +48,7 @@ Proof. by rewrite /mpoisson1 integral_set0. Qed. Lemma mpoisson1_ge0 A : (0 <= mpoisson1 A)%E. Proof. apply: integral_ge0 => x Ax. -by rewrite lee_fin poisson1_ge0. +rewrite lee_fin poisson1_ge0//. Qed. Let mpoisson1_sigma_additive : semi_sigma_additive mpoisson1. @@ -57,12 +57,12 @@ move=> /= F mF tF mUF. rewrite /mpoisson1/= integral_bigcup//=; last first. apply/integrableP; split. apply/measurable_EFinP. - exact: measurable_funS (measurable_poisson_pdf _). + exact: measurable_funS (measurable_poisson_pmf _ measurableT). rewrite (_ : (fun x => _) = (EFin \o poisson1)); last first. by apply/funext => x; rewrite gee0_abs// lee_fin poisson1_ge0//. apply: le_lt_trans. apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. - by apply/measurable_EFinP; exact: measurable_poisson_pdf. + by apply/measurable_EFinP; exact: measurable_poisson_pmf. by move=> ? _; rewrite lee_fin poisson1_ge0//. by rewrite /= integral_poisson_density// ltry. apply: is_cvg_ereal_nneg_natsum_cond => n _ _. @@ -90,19 +90,19 @@ Variable R : realType. Notation mu := (@lebesgue_measure R). Import Notations. Hypothesis integral_poisson_density : forall k, - (\int[mu]_x (@poisson_pdf R k x)%:E = 1%E)%E. + (\int[mu]_x (@poisson_pmf R x k)%:E = 1%E)%E. -Let f1 x := ((poisson1 (x : R)) ^-1)%R. +Let f1 x := ((poisson1 (x : R))^-1)%R. Let mf1 : measurable_fun setT f1. -rewrite /f1 /poisson1 /poisson_pdf. +rewrite /f1 /poisson1 /poisson_pmf. apply: (measurable_comp (F := [set r : R | r != 0%R])) => //. - exact: open_measurable. - move=> /= r [t ? <-]. by case: ifPn => // t0; rewrite gt_eqF ?mulr_gt0 ?expR_gt0//= invrK ltr0n. - apply: open_continuous_measurable_fun => //. by apply/in_setP => x /= x0; exact: inv_continuous. -- exact: measurable_poisson_pdf. +- exact: measurable_poisson_pmf. Qed. Definition staton_counting : R.-sfker X ~> _ := From 4707c82030ca4e901d0be8f60e44c1a2ce7cb0ca Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 10 Oct 2024 22:35:14 +0900 Subject: [PATCH 32/48] adjustment --- coq-mathcomp-analysis.opam | 1 - theories/lang_syntax.v | 2 +- theories/prob_lang.v | 7 +++++-- 3 files changed, 6 insertions(+), 4 deletions(-) diff --git a/coq-mathcomp-analysis.opam b/coq-mathcomp-analysis.opam index ed4e7d5181..f074956f7a 100644 --- a/coq-mathcomp-analysis.opam +++ b/coq-mathcomp-analysis.opam @@ -20,7 +20,6 @@ depends: [ "coq-mathcomp-field" "coq-mathcomp-bigenough" { (>= "1.0.0") } "coq-mathcomp-algebra-tactics" { (>= "1.2.3") } - "coq-mathcomp-zify" { (>= "1.5.0") } ] tags: [ diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index ae334c2698..2a2b658d8d 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -8,7 +8,6 @@ From mathcomp Require Import reals ereal topology normedtype sequences exp. From mathcomp Require Import esum measure lebesgue_measure numfun derive realfun. From mathcomp Require Import lebesgue_integral probability ftc kernel charge. From mathcomp Require Import prob_lang lang_syntax_util. -From mathcomp Require Import ring lra. (**md**************************************************************************) (* # Syntax and Evaluation for a Probabilistic Programming Language *) @@ -285,6 +284,7 @@ Qed. End lt0. End increasing_change_of_variables_from_decreasing. + Lemma decreasing_nonincreasing {R : realType} (F : R -> R) (J : interval R) : {in J &, {homo F : x y /~ (x < y)%R}} -> {in J &, {homo F : x y /~ (x <= y)%R}}. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index f2639adff6..d0a9370569 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -114,6 +114,10 @@ Qed. Definition poisson3 {R : realType} := @poisson_pmf R 3%:R 4. (* 0.168 *) Definition poisson10 {R : realType} := @poisson_pmf R 10%:R 4. (* 0.019 *) +Definition dep_uncurry (A : Type) (B : A -> Type) (C : Type) : + (forall a : A, B a -> C) -> {a : A & B a} -> C := + fun f p => let (a, Ba) := p in f a Ba. + (* TODO: move *) Lemma poisson_pmf_gt0 {R : realType} k (r : R) : (0 < r -> 0 < poisson_pmf r k.+1)%R. @@ -190,7 +194,6 @@ Lemma invr_nonneg_proof (R : numDomainType) (p : {nonneg R}) : (0 <= (p%:num)^-1)%R. Proof. by rewrite invr_ge0. Qed. -(* TODO: move *) Definition invr_nonneg (R : numDomainType) (p : {nonneg R}) := NngNum (invr_nonneg_proof p). @@ -2185,7 +2188,7 @@ Section gauss_lebesgue. Context d (T : measurableType d) (R : realType). Notation mu := (@lebesgue_measure R). -Let f1 (x : g_sigma_algebraType (R.-ocitv.-measurable)) := (gauss_pdf x)^-1%R. +Let f1 (x : measurableTypeR R) := (gauss_pdf x)^-1%R. Let f1E (x : R) : f1 x = (Num.sqrt (pi *+ 2) * expR (- (- x ^+ 2 / 2)))%R. Proof. From 7b8d3bd960374815d3528b82f5800e683b941005 Mon Sep 17 00:00:00 2001 From: IshiguroYoshihiro Date: Tue, 14 Jan 2025 09:46:53 +0900 Subject: [PATCH 33/48] progress - exp_normal - clean measurable_normal_prob2 - add lang_syntax_hello_wip.v - integral_normal_prob_dirac - execP_helloRight0_to_1 - wip helloRight12 - continuousT_integral - normr_exp_coeff_near_nonincreasing - near_ereal_nondecreasing_is_cvgn - near_monotone_convergence - wip measurable_normal_probi{,2} - exp_coeff2_near_increasing - radon_nikodym_crestr_fin - integration_by_substitution_shift --- theories/lang_syntax.v | 233 ++++- theories/lang_syntax_hello_wip.v | 1135 ++++++++++++++++++++++++ theories/probability.v | 1388 ++++++++++++++++++++++++++++++ 3 files changed, 2746 insertions(+), 10 deletions(-) create mode 100644 theories/lang_syntax_hello_wip.v diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 2a2b658d8d..c8c9aad7d8 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -203,6 +203,8 @@ End integral_indicator_function. End integral_indicator_function. (* TODO: naming *) +(* duplicated, cvg_compNP *) +(* Lemma cvg_atNP {T : topologicalType} {R : numFieldType} (f : R -> T) (a : R) (l : T) : f x @[x --> a] --> l <-> (f \o -%R) x @[x --> (- a)%R] --> l. Proof. @@ -210,6 +212,7 @@ rewrite nbhsN. have <-// : f x @[x --> a] = fmap [eta f \o -%R] ((- x)%R @[x --> a]). by apply/seteqP; split=> A; move=> [/= e e0 H]; exists e => //= B /H/=; rewrite opprK. Qed. +*) Lemma derivable_oo_bnd_id {R : numFieldType} (a b : R) : derivable_oo_LRcontinuous (@id R^o) a b. @@ -259,6 +262,8 @@ move=> ab cf. by rewrite /= opprK in cf. Qed. +(* duplicated, integration_by_substitution_oppr *) +(* Lemma oppr_change (f : R -> R) a b : (a < b)%R -> {within `[a, b], continuous f} -> \int[mu]_(x in `[a, b]) (f x)%:E = @@ -280,11 +285,12 @@ rewrite integration_by_substitution_decreasing//. by apply/funext => y; rewrite /= opprK. by rewrite !opprK. Qed. +*) End lt0. End increasing_change_of_variables_from_decreasing. - +(* duplicated? *) Lemma decreasing_nonincreasing {R : realType} (F : R -> R) (J : interval R) : {in J &, {homo F : x y /~ (x < y)%R}} -> {in J &, {homo F : x y /~ (x <= y)%R}}. @@ -469,6 +475,42 @@ Context {R : realType}. Inductive flag := D | P. +(* +Section uniop. + +Inductive uniop := +| uniop_not +| uniop_neg | uniop_inv. + +Definition type_of_uniop (u : uniop) : typ := +match u with +| uniop_not => Bool +| uniop_neg => Real +| uniop_inv => Real +end. + +Definition fun_of_uniop g (u : uniop) : (mctx g -> mtyp (type_of_uniop u)) -> + @mctx R g -> @mtyp R (type_of_uniop u) := +match u with +| uniop_not => (fun f x => f x && f x : mtyp Bool) +| uniop_neg => (fun f => (\- f)%R) +| uniop_inv => (fun f => (f ^-1)%R) +end. + +Definition mfun_of_uniop g b + (f : @mctx R g -> @mtyp R (type_of_uniop b)) (mf : measurable_fun setT f) + measurable_fun [set: @mctx R g] (fun_of_uniop f). +destruct b. +exact: measurable_and mf1 mf2. +exact: measurable_or mf1 mf2. +exact: measurable_funD. +exact: measurable_funB. +exact: measurable_funM. +Defined. + +End uniop. +*) + Section binop. Inductive binop := @@ -508,6 +550,8 @@ Defined. End binop. +(* TODO: rename *) +(* TODO: generalize? *) Section relop. Inductive relop := | relop_le | relop_lt | relop_eq . @@ -532,16 +576,43 @@ Defined. End relop. +Section relop_Real. +Inductive relop_real := +| relop_real_le | relop_real_lt | relop_real_eq . + +Definition fun_of_relop_real g (r : relop_real) : (@mctx R g -> @mtyp R Real) -> + (mctx g -> mtyp Real) -> @mctx R g -> @mtyp R Bool := +match r with +| relop_real_le => (fun f1 f2 x => (f1 x <= f2 x)%R) +| relop_real_lt => (fun f1 f2 x => (f1 x < f2 x)%R) +| relop_real_eq => (fun f1 f2 x => (f1 x == f2 x)%R) +end. + +Definition mfun_of_relop_real g r + (f1 : @mctx R g -> @mtyp R Real) (mf1 : measurable_fun setT f1) + (f2 : @mctx R g -> @mtyp R Real) (mf2 : measurable_fun setT f2) : + measurable_fun [set: @mctx R g] (fun_of_relop_real r f1 f2). +destruct r. +exact: measurable_fun_ler. +exact: measurable_fun_ltr. +exact: measurable_fun_eqr. +Defined. + +End relop_Real. + Inductive exp : flag -> ctx -> typ -> Type := | exp_unit g : exp D g Unit | exp_bool g : bool -> exp D g Bool | exp_nat g : nat -> exp D g Nat | exp_real g : R -> exp D g Real -| exp_pow g : nat -> exp D g Real -> exp D g Real +| exp_pow g : nat -> exp D g Real -> exp D g Real (* TODO: nat should be at second *) +| exp_pow_real g : R (* base *) -> exp D g Real -> exp D g Real | exp_bin (b : binop) g : exp D g (type_of_binop b) -> exp D g (type_of_binop b) -> exp D g (type_of_binop b) | exp_rel (r : relop) g : exp D g Nat -> exp D g Nat -> exp D g Bool +| exp_rel_real (r : relop_real) g : exp D g Real -> + exp D g Real -> exp D g Bool | exp_pair g t1 t2 : exp D g t1 -> exp D g t2 -> exp D g (Pair t1 t2) | exp_proj1 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t1 | exp_proj2 g t1 t2 : exp D g (Pair t1 t2) -> exp D g t2 @@ -551,6 +622,8 @@ Inductive exp : flag -> ctx -> typ -> Type := | exp_uniform g (a b : R) (ab : (a < b)%R) : exp D g (Prob Real) | exp_beta g (a b : nat) (* NB: should be R *) : exp D g (Prob Real) | exp_poisson g : nat -> exp D g Real -> exp D g Real +| exp_normal g (s : R) (s0 : (s != 0)%R) + : exp D g Real -> exp D g (Prob Real) (* NB: 0 < s *) | exp_normalize g t : exp P g t -> exp D g (Prob t) | exp_letin g t1 t2 str : exp P g t1 -> exp P ((str, t1) :: g) t2 -> exp P g t2 @@ -576,9 +649,11 @@ Arguments exp_unit {R g}. Arguments exp_bool {R g}. Arguments exp_nat {R g}. Arguments exp_real {R g}. -Arguments exp_pow {R g}. +Arguments exp_pow {R g} &. +Arguments exp_pow_real {R g} &. Arguments exp_bin {R} b {g} &. Arguments exp_rel {R} r {g} &. +Arguments exp_rel_real {R} r {g} &. Arguments exp_pair {R g} & {t1 t2}. Arguments exp_var {R g} _ {t} & H. Arguments exp_bernoulli {R g} &. @@ -586,10 +661,11 @@ Arguments exp_binomial {R g} &. Arguments exp_uniform {R g} &. Arguments exp_beta {R g} &. Arguments exp_poisson {R g}. +Arguments exp_normal {R g _} &. Arguments exp_normalize {R g _}. Arguments exp_letin {R g} & {_ _}. Arguments exp_sample {R g} & {t}. -Arguments exp_score {R g}. +Arguments exp_score {R g} &. Arguments exp_return {R g} & {_}. Arguments exp_if {R z g t} &. Arguments exp_weak {R} z g h {t} x. @@ -606,6 +682,8 @@ Notation "r ':R'" := (@exp_real _ _ r%R) (in custom expr at level 1, format "r :R") : lang_scope. Notation "e ^+ n" := (exp_pow n e) (in custom expr at level 1) : lang_scope. +Notation "e `^ r" := (exp_pow_real e r) + (in custom expr at level 1) : lang_scope. Notation "e1 && e2" := (exp_bin binop_and e1 e2) (in custom expr at level 2) : lang_scope. Notation "e1 || e2" := (exp_bin binop_or e1 e2) @@ -620,8 +698,12 @@ Notation "e1 <= e2" := (exp_rel relop_le e1 e2) (in custom expr at level 2) : lang_scope. Notation "e1 == e2" := (exp_rel relop_eq e1 e2) (in custom expr at level 4) : lang_scope. +Notation "e1 <=R e2" := (exp_rel_real relop_real_le e1 e2) + (in custom expr at level 2) : lang_scope. +Notation "e1 ==R e2" := (exp_rel_real relop_real_eq e1 e2) + (in custom expr at level 4) : lang_scope. Notation "'return' e" := (@exp_return _ _ _ e) - (in custom expr at level 6) : lang_scope. + (in custom expr at level 7) : lang_scope. (*Notation "% str" := (@exp_var _ _ str%string _ erefl) (in custom expr at level 1, format "% str") : lang_scope.*) (* Notation "% str H" := (@exp_var _ _ str%string _ H) @@ -637,21 +719,23 @@ Notation "\pi_1 e" := (exp_proj1 e) Notation "\pi_2 e" := (exp_proj2 e) (in custom expr at level 1) : lang_scope. Notation "'let' x ':=' e 'in' f" := (exp_letin x e f) - (in custom expr at level 5, + (in custom expr at level 6, x constr, - f custom expr at level 5, + f custom expr at level 6, left associativity) : lang_scope. Notation "{ c }" := c (in custom expr, c constr) : lang_scope. Notation "x" := x (in custom expr at level 0, x ident) : lang_scope. Notation "'Sample' e" := (exp_sample e) - (in custom expr at level 5) : lang_scope. + (in custom expr at level 6) : lang_scope. Notation "'Score' e" := (exp_score e) - (in custom expr at level 5) : lang_scope. + (in custom expr at level 6) : lang_scope. Notation "'Normalize' e" := (exp_normalize e) (in custom expr at level 0) : lang_scope. Notation "'if' e1 'then' e2 'else' e3" := (exp_if e1 e2 e3) - (in custom expr at level 6) : lang_scope. + (in custom expr at level 7) : lang_scope. +Notation "( e )" := e + (in custom expr at level 1) : lang_scope. Section free_vars. Context {R : realType}. @@ -663,8 +747,10 @@ Fixpoint free_vars k g t (e : @exp R k g t) : seq string := | exp_nat _ _ => [::] | exp_real _ _ => [::] | exp_pow _ _ e => free_vars e + | exp_pow_real _ _ e => free_vars e | exp_bin _ _ e1 e2 => free_vars e1 ++ free_vars e2 | exp_rel _ _ e1 e2 => free_vars e1 ++ free_vars e2 + | exp_rel_real _ _ e1 e2 => free_vars e1 ++ free_vars e2 | exp_pair _ _ _ e1 e2 => free_vars e1 ++ free_vars e2 | exp_proj1 _ _ _ e => free_vars e | exp_proj2 _ _ _ e => free_vars e @@ -674,6 +760,7 @@ Fixpoint free_vars k g t (e : @exp R k g t) : seq string := | exp_uniform _ _ _ _ => [::] | exp_beta _ _ _ => [::] | exp_poisson _ _ e => free_vars e + | exp_normal _ _ _ e => free_vars e | exp_normalize _ _ e => free_vars e | exp_letin _ _ _ x e1 e2 => free_vars e1 ++ rem x (free_vars e2) | exp_sample _ _ _ => [::] @@ -781,6 +868,21 @@ Local Open Scope lang_scope. Local Open Scope ring_scope. +(* TODO: PR *) +Lemma measurable_powRr b : measurable_fun setT (@powR R b). +Proof. +rewrite /powR. +apply: measurable_fun_if => //. + rewrite preimage_true setTI/=. + case: (b == 0); rewrite ?set_true ?set_false. + apply: measurableT_comp => //. + exact: measurable_fun_eqr. + exact: measurable_fun_set0. +rewrite preimage_false setTI. +apply: measurableT_comp => //. +exact: mulrr_measurable. +Qed. + Inductive evalD : forall g t, exp D g t -> forall f : dval R g t, measurable_fun setT f -> Prop := | eval_unit g : ([TT] : exp D g _) -D> cst tt ; ktt @@ -794,6 +896,9 @@ Inductive evalD : forall g t, exp D g t -> | eval_pow g n (e : exp D g _) f mf : e -D> f ; mf -> [e ^+ {n}] -D> (fun x => f x ^+ n) ; (measurable_funX n mf) +| eval_pow_real g (e : exp D g _) r f mf : e -D> f ; mf -> + [{r} `^ e] -D> (fun x => r `^ (f x)) ; measurableT_comp (measurable_powRr r) mf + | eval_bin g bop (e1 : exp D g _) f1 mf1 e2 f2 mf2 : e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> exp_bin bop e1 e2 -D> fun_of_binop f1 f2 ; mfun_of_binop mf1 mf2 @@ -802,6 +907,10 @@ Inductive evalD : forall g t, exp D g t -> e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> exp_rel rop e1 e2 -D> fun_of_relop rop f1 f2 ; mfun_of_relop rop mf1 mf2 +| eval_rel_real g rop (e1 : exp D g _) f1 mf1 e2 f2 mf2 : + e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> + exp_rel_real rop e1 e2 -D> fun_of_relop_real rop f1 f2 ; mfun_of_relop_real rop mf1 mf2 + | eval_pair g t1 (e1 : exp D g t1) f1 mf1 t2 (e2 : exp D g t2) f2 mf2 : e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> [(e1, e2)] -D> fun x => (f1 x, f2 x) ; measurable_fun_pair mf1 mf2 @@ -840,6 +949,13 @@ Inductive evalD : forall g t, exp D g t -> exp_poisson n e -D> poisson_pmf ^~ n \o f ; measurableT_comp (measurable_poisson_pmf n measurableT) mf +(* TODO *) +| eval_normal g s (s0 : (s != 0)%R) (e : exp D g _) r mr : + e -D> r ; mr -> + (exp_normal s0 e : exp D g _) -D> + (fun x => @normal_prob _ (r x) s) ; + measurableT_comp (measurable_normal_prob2 s0 ) mr + | eval_normalize g t (e : exp P g t) k : e -P> k -> [Normalize e] -D> normalize_pt k ; measurable_normalize_pt k @@ -923,6 +1039,11 @@ all: (rewrite {g t e u v mu mv hu}). inj_ex H4; subst v. inj_ex H2; subst e0. by move: H3 => /IH <-. +- move=> g e r f mf ev IH {}v {}mv. + inversion 1; subst g0 r0. + inj_ex H4; subst v. + inj_ex H2; subst e0. + by move: H3 => /IH <-. - move=> g bop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. inversion 1; subst g0 bop0. inj_ex H10; subst v. @@ -935,6 +1056,12 @@ all: (rewrite {g t e u v mu mv hu}). inj_ex H1; subst e1. inj_ex H3; subst e3. by move: H6 H7 => /IH1 <- /IH2 <-. +- move=> g rop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. + inversion 1; subst g0 rop0. + inj_ex H5; subst v. + inj_ex H1; subst e1. + inj_ex H3; subst e3. + by move: H6 H7 => /IH1 <- /IH2 <-. - move=> g t1 e1 f1 mf1 t2 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. simple inversion 1 => //; subst g0. case: H3 => ? ?; subst t0 t3. @@ -981,6 +1108,11 @@ all: (rewrite {g t e u v mu mv hu}). inj_ex H2; subst e0. inj_ex H4; subst v. by rewrite (IH _ _ H3). +- move=> g s s0 e r mr ev IH {}v {}mv. + inversion 1; subst g0 s1. + inj_ex H2; subst e0. + inj_ex H3; subst v. + by rewrite (IH _ _ H5). - move=> g t e k ev IH f mf. inversion 1; subst g0 t0. inj_ex H2; subst e0. @@ -1071,6 +1203,11 @@ all: rewrite {g t e u v eu}. inj_ex H4; subst v. inj_ex H2; subst e0. by move: H3 => /IH <-. +- move=> g e b f mf ev IH {}v {}mv. + inversion 1; subst g0 r. + inj_ex H4; subst v. + inj_ex H2; subst e0. + by move: H3 => /IH <-. - move=> g bop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. inversion 1; subst g0 bop0. inj_ex H10; subst v. @@ -1083,6 +1220,12 @@ all: rewrite {g t e u v eu}. inj_ex H1; subst e1. inj_ex H3; subst e3. by move: H6 H7 => /IH1 <- /IH2 <-. +- move=> g rop e1 f1 mf1 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. + inversion 1; subst g0 rop0. + inj_ex H5; subst v. + inj_ex H1; subst e1. + inj_ex H3; subst e3. + by move: H6 H7 => /IH1 <- /IH2 <-. - move=> g t1 e1 f1 mf1 t2 e2 f2 mf2 ev1 IH1 ev2 IH2 {}v {}mv. simple inversion 1 => //; subst g0. case: H3 => ? ?; subst t0 t3. @@ -1131,6 +1274,11 @@ all: rewrite {g t e u v eu}. inj_ex H4; subst v. inj_ex H5; subst mv. by rewrite (IH _ _ H3). +- move=> g s s0 e r mr ev IH {}v {}mv. + inversion 1; subst g0 s1. + inj_ex H2; subst e0. + inj_ex H3; subst v. + by rewrite (IH _ _ H5). - move=> g t e k ev IH {}v {}mv. inversion 1; subst g0 t0. inj_ex H2; subst e0. @@ -1207,10 +1355,14 @@ all: rewrite {z g t}. - by do 2 eexists; exact: eval_real. - move=> g n e [f [mf H]]. by exists (fun x => (f x ^+ n)%R); eexists; exact: eval_pow. +- move=> g r e [f [mf H]]. + by exists (fun x => (r `^ (f x))%R); eexists; exact: eval_pow_real. - move=> b g e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. by exists (fun_of_binop f1 f2); eexists; exact: eval_bin. - move=> r g e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. by exists (fun_of_relop r f1 f2); eexists; exact: eval_rel. +- move=> r g e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. + by exists (fun_of_relop_real r f1 f2); eexists; exact: eval_rel_real. - move=> g t1 t2 e1 [f1 [mf1 H1]] e2 [f2 [mf2 H2]]. by exists (fun x => (f1 x, f2 x)); eexists; exact: eval_pair. - move=> g t1 t2 e [f [mf H]]. @@ -1228,6 +1380,9 @@ all: rewrite {z g t}. - by eexists; eexists; exact: eval_beta. - move=> g h e [f [mf H]]. by exists (poisson_pmf ^~ h \o f); eexists; exact: eval_poisson. +- move=> g s s0 e [r [mr H]]. + exists (fun x => @normal_prob _ (r x) s : pprobability _ _). + by eexists; exact: eval_normal. - move=> g t e [k ek]. by exists (normalize_pt k); eexists; exact: eval_normalize. - move=> g t1 t2 x e1 [k1 ev1] e2 [k2 ev2]. @@ -1338,6 +1493,14 @@ Proof. by move=> f mf; apply/execD_evalD/eval_pow/evalD_execD. Qed. +Lemma execD_pow_real g r (e : exp D g _) : + let f := projT1 (execD e) in let mf := projT2 (execD e) in + execD (exp_pow_real r e) = + @existT _ _ (fun x => r `^ f x) (measurableT_comp (measurable_powRr r) mf). +Proof. +by move=> f mf; apply/execD_evalD/eval_pow_real/evalD_execD. +Qed. + Lemma execD_bin g bop (e1 : exp D g _) (e2 : exp D g _) : let f1 := projT1 (execD e1) in let f2 := projT1 (execD e2) in let mf1 := projT2 (execD e1) in let mf2 := projT2 (execD e2) in @@ -1356,6 +1519,15 @@ Proof. by move=> f1 f2 mf1 mf2; apply/execD_evalD/eval_rel; exact: evalD_execD. Qed. +Lemma execD_rel_real g rop (e1 : exp D g _) (e2 : exp D g _) : + let f1 := projT1 (execD e1) in let f2 := projT1 (execD e2) in + let mf1 := projT2 (execD e1) in let mf2 := projT2 (execD e2) in + execD (exp_rel_real rop e1 e2) = + @existT _ _ (fun_of_relop_real rop f1 f2) (mfun_of_relop_real rop mf1 mf2). +Proof. +by move=> f1 f2 mf1 mf2; apply/execD_evalD/eval_rel_real; exact: evalD_execD. +Qed. + Lemma execD_pair g t1 t2 (e1 : exp D g t1) (e2 : exp D g t2) : let f1 := projT1 (execD e1) in let f2 := projT1 (execD e2) in let mf1 := projT2 (execD e1) in let mf2 := projT2 (execD e2) in @@ -1417,6 +1589,13 @@ Lemma execD_beta g a b : existT _ (cst [the probability _ _ of beta_prob a b]) (measurable_cst _). Proof. exact/execD_evalD/eval_beta. Qed. +Lemma execD_normal g s s0 e : + let f := projT1 (execD e) in let mf := projT2 (execD e) in + @execD g _ (@exp_normal _ _ s s0 e) = + existT _ (fun x => [the probability _ _ of @normal_prob _ (f x) s]) + (measurableT_comp (measurable_normal_prob2 s0) mf). +Proof. exact/execD_evalD/eval_normal/evalD_execD. Qed. + Lemma execD_normalize_pt g t (e : exp P g t) : @execD g _ [Normalize e] = existT _ (normalize_pt (execP e) : _ -> pprobability _ _) @@ -1492,4 +1671,38 @@ apply: eq_kernel => y V. exact: He. Qed. +Lemma congr_letinl_new {R : realType} g t1 t2 str (e1 e2 : @exp _ _ g t1) + (e : @exp _ _ (_ :: g) t2) x U : + (forall y V, measurable V -> execP e1 y V = execP e2 y V) -> + measurable U -> + @execP R g t2 [let str := e1 in e] x U = + @execP R g t2 [let str := e2 in e] x U. +Proof. +Abort. +(* by move=> + mU; move/eq_sfkernel => He; rewrite !execP_letin He. Qed. *) + +Lemma congr_letinr_new {R : realType} g t1 t2 str (e : @exp _ _ _ t1) + (e1 e2 : @exp _ _ (_ :: g) t2) x U : + (forall y V, measurable V -> execP e1 (y, x) V = execP e2 (y, x) V) -> + @execP R g t2 [let str := e in e1] x U = @execP R g t2 [let str := e in e2] x U. +Proof. +Abort. +(* + by move=> He; rewrite !execP_letin !letin'E; apply: eq_integral => ? _; exact: He. +Qed. +*) +Lemma congr_normalize_new {R : realType} g t (e1 e2 : @exp R _ g t) : + (forall x U, measurable U -> execP e1 x U = execP e2 x U) -> + execD [Normalize e1] = execD [Normalize e2]. +Proof. +Abort. +(* +move=> He; apply: eq_execD. +rewrite !execD_normalize_pt /=. +f_equal. +apply: eq_kernel => y V. +exact: He. +Qed. +*) + Local Close Scope lang_scope. diff --git a/theories/lang_syntax_hello_wip.v b/theories/lang_syntax_hello_wip.v new file mode 100644 index 0000000000..3c79472ce7 --- /dev/null +++ b/theories/lang_syntax_hello_wip.v @@ -0,0 +1,1135 @@ +From Coq Require Import String. +From HB Require Import structures. +From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. +From mathcomp.classical Require Import mathcomp_extra boolp. +From mathcomp Require Import ring lra. +From mathcomp Require Import classical_sets functions cardinality fsbigop. +From mathcomp Require Import interval_inference reals ereal topology normedtype. +From mathcomp Require Import sequences esum measure lebesgue_measure numfun exp. +From mathcomp Require Import trigo realfun charge lebesgue_integral kernel. +From mathcomp Require Import probability prob_lang. +From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. + +(**md**************************************************************************) +(* # Inferring behavior from text-massage data example *) +(* *) +(* ref: *) +(* - Chung-chieh Shan, Equational reasoning for probabilistic programming, *) +(* POPL TutorialFest 2018 *) +(* https://homes.luddy.indiana.edu/ccshan/rational/equational-handout.pdf *) +(* - *) +(* *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import Order.TTheory GRing.Theory Num.Def Num.ExtraDef Num.Theory. +Import numFieldTopology.Exports. + +Local Open Scope classical_set_scope. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. + + +Definition neq01 {R : realType} := (@GRing.oner_neq0 R). +Definition exp_normal1 {R g} := (@exp_normal R g _ neq01). + +Definition neq0Vsqrt2 {R : realType} : ((@Num.sqrt R 2)^-1 != 0)%R. +Proof. exact: lt0r_neq0. Qed. +Definition exp_normal_Vsqrt2 {R g} := (@exp_normal R g _ neq0Vsqrt2). + +Definition neq0sqrt32 {R : realType} : ((@Num.sqrt R (3 / 2)) != 0)%R. +Proof. exact: lt0r_neq0. Qed. +Definition exp_normal_sqrt32 {R g} := (@exp_normal R g _ neq0sqrt32). + +(* +Section eq_lebesgue_measure_measurable. +Context {R : realType}. + +Local Notation mu := lebesgue_measure. + +Variable f : cumulative R. +Hypothesis intf : integrable mu setT (EFin \o f). + +Local Notation lsmf := (lebesgue_stieltjes_measure f). + +Let fin_lsmf : fin_num_fun lsmf. +Proof. +apply: lty_fin_num_fun. +rewrite /=/lsmf/measure_extension/mu_ext. +Admitted. + +HB.instance Definition _ := @Measure_isFinite.Build _ _ _ lsmf fin_lsmf. + +Lemma Radon_Nikodym_ae_unique : + ae_eq mu setT (EFin \o f) ('d lsmf '/d mu). +Proof. +apply: integral_ae_eq => //. + admit. +move=> E _ mE. +rewrite -Radon_Nikodym_integral//; last first. + admit. + + +Lemma eq_lebesgue_measure_measurable (m' : measure _ R) : +sigma_additive m' -> +(forall U : set _, measurable U -> m' U = mu U) -> +m' = mu. +Proof. +move=> sm' H. +apply: eq_measure. +apply/funext => U. +have : setT U by []. +rewrite -(setUv measurable). +move=> [mU|nmU]; first exact: H. +pose e := m' U. +have : m' U = e by []. +Abort. + +End eq_lebesgue_measure_measurable. +*) + +(* TODO: PR *) +Section ge0_fin_measurable_probability_integrable. +Context d {T : measurableType d} {R : realType} {p : probability T R} + {f : T -> \bar R}. + +Lemma ge0_fin_measurable_probability_integrable : + (forall x, 0 <= f x) -> + (exists M : R, forall x, f x <= M%:E) -> + measurable_fun setT f -> + p.-integrable setT f. +Proof. +move=> f0 [M fleM] mf. +apply/integrableP; split => //. +apply (@le_lt_trans _ _ (\int[p]_x M%:E)). + apply: ge0_le_integral => //=. + exact: measurableT_comp. + move=> t _. + exact: (@le_trans _ _ (f t)). + move=> t _. + by rewrite gee0_abs. +by rewrite integral_cst// probability_setT mule1 ltry. +Qed. + +End ge0_fin_measurable_probability_integrable. + +(* TODO: PR *) +Section pkernel_probability_integrable. +Context d d' {T : measurableType d} {T' : measurableType d'} {R : realType} + {p : probability T R} {f : R.-pker T ~> T'}. + +Lemma pkernel_probability_integrable V : + measurable V -> + p.-integrable setT (fun x => f x V). +Proof. +move=> mV. +apply: ge0_fin_measurable_probability_integrable => //. + exists 1%R. + move=> x. + apply: (@le_trans _ _ (f x setT)). + by apply: le_measure; rewrite ?inE. + by rewrite prob_kernel. +exact: measurable_kernel. +Qed. + +End pkernel_probability_integrable. + +(* TODO: move to probability? *) +Section normal_prob_properties. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +Local Open Scope charge_scope. + +Lemma normal_pdf_uniq_ae (m s : R) (s0 : (0 < s)%R) : + ae_eq mu setT + ('d ((charge_of_finite_measure (@normal_prob R m s))) '/d mu) + (EFin \o (@normal_pdf R m s)). +Proof. +apply: integral_ae_eq => //. + apply: Radon_Nikodym_integrable => /=. + exact: normal_prob_dominates. + apply/measurable_EFinP. + exact: measurable_normal_pdf. +move=> /= E _ mE. +by rewrite -Radon_Nikodym_integral//=; apply: normal_prob_dominates. +Qed. + +End normal_prob_properties. + +(* +Section disintegration_measure. +Context {R : realType}. +(* define lebesgue syntax *) + +End disintegration_measure. + +Section disintegration_program. + +Variable l : ctx. + +Definition lhs : exp R _ l _ := + [let "t" := Sample {exp_normal ltr01 (exp_real 0)} in + let "p" := (* ? TODO *) in + return (#{"t"}, #{"p"})]. + +Definition rhs : exp R _ l _ := + [let "t" := Sample exp_lebesgue(* TODO *) in + let "p" := Score (exp_normal_pdf [#{"t"}])) (* TODO *) in + return (#{"t"}, #{"p"})]. + +Lemma disintegration_normal : +execD lhs = exec rhs. + +End disintegration_program. +*) + +Section normal_prob_lemmas. +Context {R: realType}. +Local Notation mu := lebesgue_measure. + +(* TODO: move *) +Lemma emeasurable_normal_prob (s : R) U : s != 0%R -> measurable U -> + measurable_fun setT (fun x => normal_prob x s U). +Proof. +move=> s0 mU. +under [X in _ _ X]eq_fun. + move=> x. + rewrite -(@fineK _ (_ x _ _)); last first. + rewrite ge0_fin_numE//. + apply: (@le_lt_trans _ _ (normal_prob x s setT)). + by apply: le_measure; rewrite ?inE. + by rewrite probability_setT ltry. + over. +apply/measurable_EFinP. +apply: (continuous_measurable_fun). +exact: normal_prob_continuous. +Qed. + +(* TODO: s != 0 *) +Lemma integral_normal_prob (m s : R) (s0 : (0 < s)%R) f U : + measurable U -> + (normal_prob m s).-integrable U f -> + \int[@normal_prob _ m s]_(x in U) f x = + \int[mu]_(x in U) (f x * (normal_pdf m s x)%:E). +Proof. +move=> mU intf. +move/integrableP : (intf) => [mf intf_lty]. +rewrite -(Radon_Nikodym_change_of_variables (normal_prob_dominates m s))//=. +apply: ae_eq_integral => //=. + apply: emeasurable_funM => //. + apply: measurable_funTS. + have : charge_of_finite_measure (normal_prob m s) `<< mu. + exact: normal_prob_dominates m s. + move/Radon_Nikodym_integrable. + by move/integrableP => []. + apply: emeasurable_funM => //. + apply/measurable_EFinP. + apply: measurable_funTS. + exact: measurable_normal_pdf. +apply: ae_eq_mul2l. +apply: (ae_eq_subset (@subsetT _ U)). +exact: (normal_pdf_uniq_ae m s0). +Qed. + +(* +Definition tail1 : @exp R _ [:: ("x", Real)] _ := + [let "z" := Sample {exp_normal ltr01 [#{"x"}]} in + return #{"z"}]. + +Lemma helloRight01 : +execP [let "x" := Sample {exp_normal ltr01 (exp_real 0)} in + let "_" := Score {expR 1} `^ + ({0}:R - (#{"y0"} - #{"x"}) ^+ {2} * {2^-1}:R) + * {(Num.sqrt 2 * pi)^-1}:R in + {exp_weak _ [::] [:: ("x", Real); ("y0", Real)] ("_", Unit) + (exp_weak _ [:: ("x", Real)] [::] ("y0", Real) tail1 _)}] = + [let "_" := Score {expR 1} `^ ({0}:R - #{"y0"} ^+ {2} * {4^-1}:R) * + {(Num.sqrt (4 * pi))^-1}:R in + let "x" := Sample {exp_normal ltr0Vsqrt2 [#{"y0"} * {2^-1}:R ]} in + {tail1}]. +*) + +Lemma normal_prob_integrable_dirac (m s : R) (V : set R): measurable V -> + (normal_prob m s).-integrable setT (fun x => \d_x V). +Proof. +move=> mV. +apply/integrableP; split; first exact: measurable_fun_dirac. +rewrite -(setUv V). +rewrite ge0_integral_setU => //; first last. + exact/disj_setPCl. + rewrite setUv. + apply: measurableT_comp => //. + exact: measurable_fun_dirac. + exact: measurableC. +under eq_integral. + move=> x Vx. + rewrite diracE Vx/= normr1. + over. +under [X in _ + X < _]eq_integral. + move=> x. + rewrite inE/= => nVx. + have {}nVx := (memNset nVx). + rewrite indicE nVx/= normr0. + over. +rewrite !integral_cst//=; last exact: measurableC. +rewrite mul1e mul0e adde0. +apply: (le_lt_trans (probability_le1 (normal_prob m s) mV)). +exact: ltey. +Qed. + +Lemma integral_normal_prob_dirac (s : R) (m : R) V : + (0 < s)%R -> + measurable V -> + \int[normal_prob m s]_x0 (\d_x0 V) = normal_prob m s V. +Proof. +move=> s0 mV. +rewrite integral_normal_prob => //; last exact: (normal_prob_integrable_dirac m s). +under eq_integral do rewrite diracE. +rewrite /normal_prob. +rewrite [in RHS]integral_mkcond. +under [in RHS]eq_integral do rewrite patchE. +rewrite /=. +apply: eq_integral => x _. +by case: ifP => xV/=; rewrite ?mul1e ?mul0e. +Qed. + +End normal_prob_lemmas. + +Section hello_programs. +Local Open Scope lang_scope. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +Definition guard_real {g} str (r : R) : + @exp R P [:: (str, _) ; g] _ := + [if #{str} ==R {r}:R then return TT else Score {0}:R]. + +Definition helloWrong (y0 : R) : @exp R _ [::] _ := + [Normalize + let "x" := Sample {exp_normal1 (exp_real 0)} in + let "y" := Sample {exp_normal1 [#{"x"}]} in + let "_" := {guard_real "y" y0} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. + +Definition helloRight : @exp R _ [:: ("y0", Real)] _ := + [Normalize + let "x" := Sample {exp_normal1 (exp_real 0)} in + let "_" := Score ({expR 1} `^ + ({0}:R - (#{"y0"} - #{"x"}) ^+ {2} * {2^-1}:R)) + * {(Num.sqrt 2 * pi)^-1}:R in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. + +Definition helloJoint : @exp R _ [::] _ := + [Normalize + let "x" := Sample {exp_normal1 (exp_real 0)} in + let "y" := Sample {exp_normal1 [#{"x"}]} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return (#{"y"}, #{"z"})]. + +End hello_programs. + +Section helloRight_subproofs. +Local Open Scope lang_scope. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +Local Definition int_normal_helloRightP := + (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + \int[normal_prob 0 1]_x (fun x0 => + `|expR (- ((y.1 - x0) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * normal_prob x0 1 V) x). + +Local Definition int_mu_helloRightP := + (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + \int[mu]_x + (`|expR (- ((y.1 - x) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * normal_prob x 1 V * + (normal_pdf 0 1 x)%:E)). + +Local Definition int_normal_helloRight1P := +(fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => +(\int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x +(((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))))%:num)%:E * + normal_prob x 1 V)%E)). + +Local Definition int_mu_helloRight1P := +(fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + \int[mu]_x + (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E * + (normal_prob x 1 V * (normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 x)%:E))). + +Local Definition int_normal_helloRight2P := +(fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + \int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x normal_prob x 1 V). + +Local Definition int_mu_helloRight2P := +(fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + normal_prob (y.1 / 2) (Num.sqrt (3 / 2)) V). + +Lemma int_change_helloRightP y V : measurable V -> +int_normal_helloRightP y V = int_mu_helloRightP y V. +Proof. +move=> mV; rewrite /int_normal_helloRightP. +rewrite integral_normal_prob//; first last. + apply: integrableMr => //. + apply: (@measurableT_comp _ _ _ _ _ _ (@normr _ R^o)) => //. + apply: measurable_funM => //. + apply: measurableT_comp => //. + apply: measurableT_comp => //. + apply: measurable_funM => //. + apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => (t ^+ 2%R)%R)) => //. + exact: measurable_funB. + exists (Num.sqrt (2 * pi))^-1%R; split; first exact: num_real. + move=> x x_gt. + move=> p _. + rewrite /= normr_id. + have /normr_idP -> : (0 <= expR (- ((y.1 - p) ^+ 2%R / 2)) / (Num.sqrt (2 * pi)))%R. + apply: mulr_ge0; first exact: expR_ge0. + by rewrite invr_ge0 sqrtr_ge0// mulr_ge0// pi_ge0. + apply/ltW. + apply: le_lt_trans x_gt. + rewrite -[leRHS]mul1r ler_pM//. + exact: expR_ge0. + by rewrite -expR0 ler_expR oppr_le0 mulr_ge0// ?sqr_ge0// expR0 invr_ge0. + apply/integrableP; split. + exact: emeasurable_normal_prob. + apply/abse_integralP => //. + exact: emeasurable_normal_prob. + have/gee0_abs -> : 0 <= \int[normal_prob 0 1]_x normal_prob x 1 V. + exact: integral_ge0. + apply: (@le_lt_trans _ _ (\int[normal_prob 0 1]_x (cst 1%R x)%:E)). + apply: ge0_le_integral => //. + exact: emeasurable_normal_prob. + apply/measurable_EFinP. + exact: measurable_cst. + move=> x _. + exact: probability_le1. + by rewrite /= integral_cst// mul1e probability_setT ltry. +Qed. + +Lemma int_change_helloRight1P y V : measurable V -> +int_normal_helloRight1P y V = int_mu_helloRight1P y V. +Proof. +move=> mV; rewrite /int_normal_helloRight1P. +rewrite integral_normal_prob//; first last. + apply/integrableP; split. + apply: measurable_funeM. + exact: emeasurable_normal_prob. + apply: (@le_lt_trans _ _ + (\int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x + (((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / + Num.sqrt (4 * pi))))%:num)%:E * (cst 1%R x)%:E))). + apply: ge0_le_integral => //. + apply: measurableT_comp => //. + apply: measurable_funeM. + exact: emeasurable_normal_prob. + by move=> x _; rewrite /= mule1. + rewrite /= mule1. + exact/measurable_EFinP. + move=> x _/=. + rewrite gee0_abs; last exact: mule_ge0. + rewrite lee_pmul//. + exact: probability_le1. + rewrite integralZr//; last exact: finite_measure_integrable_cst. + by rewrite integral_cst// mule1 probability_setT ltry. +by under eq_integral do rewrite -muleA. +Qed. + +Lemma int_change_helloRight2P y V : + measurable V -> + int_normal_helloRight2P y V = int_mu_helloRight2P y V. +Proof. +move=> mV. +rewrite /int_normal_helloRight2P. +rewrite /int_mu_helloRight2P. +(* rewrite by fubini *) +transitivity (\int[mu]_(z in V) + \int[mu]_x ( + (normal_pdf x 1 z * normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 x)%:E)). + rewrite integral_normal_prob//; last first. + apply: ge0_fin_measurable_probability_integrable => //. + exists 1%R => ?; exact: probability_le1. + exact: emeasurable_normal_prob. + under eq_integral. + move=> x _. + rewrite /normal_prob. + rewrite -integralZr//; last first. + apply: (integrableS measurableT) => //. + exact: integrable_normal_pdf. + under eq_integral. + move=> z _. + rewrite -EFinM. + rewrite [X in X%:E](_:_= + (fun xz : R * R => (normal_pdf xz.1 1 xz.2 * normal_pdf (y.1 / 2) + (Num.sqrt 2)^-1 xz.1)%R) (x, z)); last by []. + over. + rewrite integral_mkcond. + rewrite epatch_indic. + over. + rewrite /=. + rewrite (@fubini_tonelli _ _ _ _ _ mu mu +(EFin \o ((fun xz : R * R => (normal_pdf xz.1 1 xz.2 * + normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 xz.1)%R * \1_V xz.2)%R))); + first last. + move=> x. + rewrite lee_fin/=. + apply: mulr_ge0 => //. + apply: mulr_ge0; exact: normal_pdf_ge0. + apply/measurable_EFinP. + apply: measurable_funM; last first. + apply: measurable_indic. + rewrite -[X in measurable X]setTX. + exact: measurableX. + apply: measurable_funM. + under [X in measurable_fun _ X]eq_fun. + move=> x. + rewrite normal_pdfE//. + rewrite normal_pdf0_center/=. + over. + rewrite /=. + apply: measurableT_comp. + exact: measurable_normal_pdf0. + exact: measurable_funB. + apply: measurableT_comp => //. + exact: measurable_normal_pdf. + under eq_integral. + move=> x _. + rewrite /=. + under eq_integral do rewrite EFinM. + rewrite -epatch_indic. + over. + rewrite /=. + rewrite [RHS]integral_mkcond/=. + apply: eq_integral. + move=> /= x _. + rewrite patchE. + case: ifPn => xV. + apply: eq_integral => z _/=. + by rewrite ifT. + apply: integral0_eq => /= z _. + rewrite ifF//. + apply/negbTE. + rewrite notin_setE/=. + move/negP in xV. + by move/mem_set. +apply: eq_integral. +move=> x _. +transitivity ((2 * pi * (Num.sqrt 2))^-1%:E * + \int[mu]_x0 ((expR (- (x - x0) ^+ 2 / 2))%R * expR (- (x0 - (y.1 / 2)) ^+ 2 ))%:E). + under eq_integral. + move=> z _. + rewrite !normal_pdfE//. + rewrite /normal_pdf0. + + over. + admit. +rewrite normal_pdfE// /normal_pdf0. +evar (C : Real.sort R). +transitivity (((2 * pi * Num.sqrt 2)^-1)%:E * + \int[mu]_x0 (expR (- (3 / 2)%R * (x0 - C) ^+ 2 - (x - y.1 / 2) ^+ 2 / (3 / 2 *+ 2)))%:E). + admit. +(* gauss integral *) +admit. +Admitted. + + +End helloRight_subproofs. + +Section helloRight_proof. +Local Open Scope lang_scope. +Context {R : realType}. + +Local Notation mu := lebesgue_measure. + +(* TODO: remove P, helloRight -> helloRight0 *) +(* NB: exp_powR level setting is mistaken? *) +(* ((_ `^ _) * _) cannot write as (_ `^ _ * _) *) +Definition helloRightP : @exp R P [:: ("y0", Real)] Real := +[let "x" := Sample {exp_normal1 (exp_real 0)} in + let "_" := Score ({expR 1} `^ + ({0}:R - (#{"y0"} - #{"x"}) ^+ {2} * {2^-1}:R)) + * {(Num.sqrt (2 * pi))^-1}:R in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. + +(* rename *) +Definition helloRight1P : @exp R _ [:: ("y0", Real)] Real := + [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2} * {4^-1}:R)) * + {(Num.sqrt (4 * pi))^-1}:R in + let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R ]} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. + +Definition helloRight1 : @exp R _ [:: ("y0", Real)] _ := + [Normalize {helloRight1P}]. + +Definition helloRight2P : @exp R _ [:: ("y0", Real)] _ := + [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2} * {4^-1}:R)) * + {(Num.sqrt (4 * pi))^-1}:R in + Sample {exp_normal_sqrt32 [#{"y0"} * {2^-1}:R]}]. (* ? *) + +Definition helloRight2 : @exp R _ [:: ("y0", Real)] _ := + [Normalize {helloRight2P}]. + +(* TODO: separate program and proof each line +Local Definition helloRightP_line1 : + forall _, @exp R P [:: ("y0", Real)] Real := +(fun (t : @exp R P [:: ("x", Real); ("y0", Real)] Real) => +[let "x" := Sample {exp_normal1 (exp_real 0)} in + {t}]). +*) + +Lemma execP_helloRight12 y V : measurable V -> + @execP R [:: ("_", Unit); ("y0", Real)] _ + [let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R ]} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}] y V = + @execP R [:: ("_", Unit); ("y0", Real)] _ + [Sample {exp_normal_sqrt32 [#{"y0"} * {2^-1}:R]}] y V. +Proof. +move=> mV. +rewrite !execP_letin. +rewrite !execP_sample !execD_normal/=. +rewrite (@execD_bin _ _ binop_mult) execD_real/=. +rewrite execP_return. +rewrite !exp_var'E (execD_var_erefl "y0") (execD_var_erefl "x") (execD_var_erefl "z")/=. +rewrite !letin'E/=. +under eq_integral do rewrite letin'E/=. +rewrite /=. +(* *) +under eq_integral do rewrite integral_normal_prob_dirac//. +exact: (int_change_helloRight2P _ mV). +Qed. + +(* ref: https://homes.luddy.indiana.edu/ccshan/rational/equational-handout.pdf + * p.2, equation (7) + * change sampling order by bayes' theorem. + *) + +Local Definition executed_helloRightP := + (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + letin' + (sample (fun=> normal_prob 0 1) + (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) (kr 0))) + (letin' + (score + (measurable_funM + (measurableT_comp (measurable_powRr (expR 1)) + (measurable_funB (kr 0) + (measurable_funM + (measurable_funX 2%R + (measurable_funB (measurable_acc_typ [:: Real; Real] 1) + (measurable_acc_typ [:: Real; Real] 0))) + (kr 2^-1)))) (kr (Num.sqrt (2 * pi))^-1))) + (letin' + (sample + (fun + x : unit * + (g_sigma_algebraType R.-ocitv.-measurable * + (g_sigma_algebraType R.-ocitv.-measurable * unit)) => + normal_prob x.2.1 1) + (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) + (measurable_acc_typ [:: Unit; Real; Real] 1))) + (ret (measurable_acc_typ [:: Real; Unit; Real; Real] 0)))) y V : \bar R). + +Local Definition executed_helloRight1P := + (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + letin' + (score + (measurable_funM + (measurableT_comp (measurable_powRr (expR 1)) + (measurable_funB (kr 0) + (measurable_funM (measurable_funX 2%R (measurable_acc_typ [:: Real] 0)) (kr 4^-1)))) + (kr (Num.sqrt (4 * pi))^-1))) + (letin' + (sample + (fun x : unit * (g_sigma_algebraType R.-ocitv.-measurable * unit) => + probability_normal_prob__canonical__measure_Probability (x.2.1 / 2) (Num.sqrt 2)^-1) + (measurableT_comp (measurable_normal_prob2 neq0Vsqrt2) + (measurable_funM (measurable_acc_typ [:: Unit; Real] 1) (kr 2^-1)))) + (letin' + (sample + (fun + x : g_sigma_algebraType R.-ocitv.-measurable * + (unit * (g_sigma_algebraType R.-ocitv.-measurable * unit)) => + probability_normal_prob__canonical__measure_Probability x.1 1) + (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) + (measurable_acc_typ [:: Real; Unit; Real] 0))) + (ret (measurable_acc_typ [:: Real; Real; Unit; Real] 0)))) y V : \bar R). + +(* TODO: rename, this is step0 of lhs *) +Lemma lhs_execPE y V : + execP helloRightP y V = executed_helloRightP y V. +Proof. +rewrite !execP_letin. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite execD_real/=. +rewrite execP_score. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow_real/=. +rewrite (@execD_bin _ _ binop_minus)/=. +rewrite execD_real/=. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow/=. +rewrite (@execD_bin _ _ binop_minus)/=. +rewrite exp_var'E/= (execD_var_erefl "y0")/=. +rewrite exp_var'E/= (execD_var_erefl "x")/=. +rewrite execD_real/=. +rewrite execD_real/=. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite exp_var'E/= (execD_var_erefl "x")/=. +rewrite execP_return/=. +rewrite exp_var'E/= (execD_var_erefl "z")/=. +done. +Qed. + +Lemma rhs_execPE y V : + execP helloRight1P y V = executed_helloRight1P y V. +Proof. +rewrite execP_letin. +rewrite execP_score. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow_real/=. +rewrite (@execD_bin _ _ binop_minus)/=. +rewrite execD_real/=. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow/=. +rewrite exp_var'E/= (execD_var_erefl "y0")/=. +rewrite execD_real/=. +rewrite execD_real/=. +rewrite execP_letin. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite exp_var'E/= (execD_var_erefl "y0")/=. +rewrite execD_real/=. +rewrite execP_letin. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite exp_var'E/= (execD_var_erefl "x")/=. +rewrite execP_return. +rewrite exp_var'E/= (execD_var_erefl "z")/=. +done. +Qed. + +Lemma helloRightP_intE y V : measurable V -> + executed_helloRightP y V = int_normal_helloRightP y V. +Proof. +move=> mV. +rewrite /executed_helloRightP. +rewrite [in LHS]letin'E/=. +under [in LHS]eq_integral. + move=> x _. + rewrite letin'E/=. + under eq_integral. + move=> u _. + rewrite letin'E/=. + rewrite integral_normal_prob_dirac//. + over. + rewrite /=. + rewrite ge0_integral_mscale/=; first last. + - by move=> ? _. + - by []. + - exact: measurableT. + rewrite integral_dirac; first last. + - by []. + - exact: measurableT. + rewrite diracT mul1e. + rewrite sub0r. + rewrite -expRM mul1r. + over. +rewrite /=. +done. +Qed. + +Lemma helloRight1P_intE y V : measurable V -> + executed_helloRight1P y V = int_normal_helloRight1P y V. +Proof. +move=> mV; rewrite /executed_helloRight1P. +rewrite letin'E/=. +under eq_integral. + move=> u _. + rewrite letin'E/=. + under eq_integral. + move=> x _. + rewrite letin'E/=. + rewrite integral_normal_prob_dirac//. + over. + over. +rewrite ge0_integral_mscale//; first last. + move=> ? _. + exact: integral_ge0. +rewrite integral_dirac//. +rewrite diracT mul1e. +rewrite sub0r -expRM mul1r. +rewrite /=. +rewrite -(@ge0_integralZl _ _ R (normal_prob _ _) _ measurableT _ _ (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E) )//; first last. +exact: emeasurable_normal_prob. +Qed. + +(* TODO: rename *) +Lemma execP_helloRight0_to_1 y V: +measurable V -> +execP helloRightP y V = execP helloRight1P y V. +Proof. +move=> mV. +(* lhs *) +rewrite lhs_execPE. +rewrite (helloRightP_intE _ mV). +rewrite (int_change_helloRightP _ mV). +(* rhs *) +rewrite rhs_execPE. +rewrite (helloRight1P_intE _ mV). +rewrite (int_change_helloRight1P _ mV). +(* eq_integral *) +apply: eq_integral. +move=> x _. +rewrite [in LHS]muleAC. +rewrite [in RHS](muleC (normal_prob x 1 V)) muleA. +congr (fun t => (t * normal_prob x 1 V)%E). +rewrite [in LHS]ger0_norm; last first. +- by rewrite mulr_ge0// expR_ge0. +rewrite [in RHS]ger0_norm; last first. +- by rewrite mulr_ge0// expR_ge0. +rewrite -!EFinM; congr EFin. +rewrite !normal_pdfE// /normal_pdf0. +rewrite mulrA. +rewrite mulrAC. +rewrite -(@sqrtrV _ 2)//. +rewrite sqr_sqrtr; last by rewrite invr_ge0. +rewrite subr0. +rewrite [X in X / _ = _]mulrC mulrA -expRD -mulrA -invfM. +rewrite [RHS]mulrAC [X in _ = _ * X / _]mulrC mulrA -mulrA -expRD -invfM. +congr *%R. + congr expR; lra. +congr GRing.inv. +rewrite expr1n mul1r -mulrnAr -(mulr_natl pi) mulrA mulVf// mul1r -expr2. +rewrite sqr_sqrtr; last by rewrite mulr_ge0// pi_ge0. +rewrite !sqrtrM// mulrCA -expr2 sqr_sqrtr; last exact: pi_ge0. +congr *%R. +have /normr_idP <- : (@GRing.zero R <= 2)%R by []. +rewrite -(sqrtr_sqr 2%R). +congr Num.sqrt. +lra. +Qed. + +Lemma helloRight0_to_2 y V : measurable V -> +execP helloRightP y V = execP helloRight2P y V. +Proof. +move=> mV. +rewrite execP_helloRight0_to_1//. +rewrite /helloRight1P/helloRight2P. +rewrite execP_letin/= [in RHS]execP_letin/=. +rewrite letin'E [RHS]letin'E. +by under eq_integral do rewrite execP_helloRight12//. +Qed. + +(* +Lemma helloRight0_to_1 : +execD helloRight = execD helloRight1. +Proof. +apply: congr_normalize => y V. +(* lhs *) +rewrite ![in LHS]execP_letin. +rewrite [in LHS]execP_sample. +rewrite [in LHS]execD_normal/=. +rewrite [in LHS]execD_real/=. +rewrite [in LHS]execP_score. +rewrite [in LHS]execD_pow_real/=. +rewrite [in LHS](@execD_bin _ _ binop_mult)/=. +rewrite [in LHS](@execD_bin _ _ binop_minus)/=. +rewrite [in LHS]execD_real/=. +rewrite [in LHS](@execD_bin _ _ binop_mult)/=. +rewrite [in LHS]execD_pow/=. +rewrite [in LHS](@execD_bin _ _ binop_minus)/=. +rewrite [in LHS]exp_var'E/= (execD_var_erefl "y0")/=. +rewrite [in LHS]exp_var'E/= (execD_var_erefl "x")/=. +rewrite [in LHS]execD_real/=. +rewrite [in LHS]execD_real/=. +rewrite [in LHS]execP_sample. +rewrite [in LHS]execD_normal/=. +rewrite [in LHS]exp_var'E/= (execD_var_erefl "x")/=. +rewrite [in LHS]execP_return/=. +rewrite [in LHS]exp_var'E/= (execD_var_erefl "z")/=. +(* rhs *) +rewrite [in RHS]execP_letin. +rewrite [in RHS]execP_score. +rewrite [in RHS]execD_pow_real/=. +rewrite [in RHS](@execD_bin _ _ binop_mult)/=. +rewrite [in RHS](@execD_bin _ _ binop_minus)/=. +rewrite [in RHS]execD_real/=. +rewrite [in RHS](@execD_bin _ _ binop_mult)/=. +rewrite [in RHS]execD_pow/=. +rewrite [in RHS]exp_var'E/= (execD_var_erefl "y0")/=. +rewrite [in RHS]execD_real/=. +rewrite [in RHS]execD_real/=. +rewrite [in RHS]execP_letin. +rewrite [in RHS]execP_sample. +rewrite [in RHS]execD_normal/=. +rewrite [in RHS](@execD_bin _ _ binop_mult)/=. +rewrite [in RHS]exp_var'E/= (execD_var_erefl "y0")/=. +rewrite [in RHS]execD_real/=. +rewrite [in RHS]execP_letin. +rewrite [in RHS]execP_sample. +rewrite [in RHS]execD_normal/=. +rewrite [in RHS]exp_var'E/= (execD_var_erefl "x")/=. +rewrite [in RHS]execP_return. +rewrite [in RHS]exp_var'E/= (execD_var_erefl "z")/=. +(* lhs *) +rewrite [in LHS]letin'E/=. +under [in LHS]eq_integral. + move=> x _. + rewrite letin'E/=. + under eq_integral. + move=> u _. + rewrite letin'E/=. + rewrite integral_normal_prob_dirac; last first. + admit. (* ? *) + over. + rewrite /=. + rewrite ge0_integral_mscale/=; first last. + - by move=> ? _. + - by []. + - exact: measurableT. + rewrite integral_dirac; first last. + - by []. + - exact: measurableT. + rewrite diracT mul1e. + rewrite sub0r. + rewrite -expRM mul1r. + over. +rewrite /=. +rewrite [in LHS]integral_normal_prob; first last. +- apply: integrableMr => //. + admit. + exists (Num.sqrt 2 * pi)^-1; split; first exact: num_real. + move=> x x_ge_Vsq2pi. + move=> p _. + rewrite /= normr_id. + have /normr_idP -> : 0 <= expR (- ((y.1 - p) ^+ 2%R / 2) / (Num.sqrt 2 * pi)). +red. + admit. +have := (@integrableZl _ _ _ (normal_prob 0 1) _ (@measurableT _ R) _ (fun x => normal_prob x 1 V)). +(* ? *) admit. +- apply: emeasurable_funM; last first. +(* TODO: Define instance of normal_prob is kernel *) + (* apply: measurable_kernel. *) +rewrite /=. + under [X in _ _ X]eq_fun. + move=> x. + rewrite -(@fineK _ (normal_prob _ _ _)); last first. + admit. + over. + apply/measurable_EFinP. + apply: continuous_measurable_fun. + apply: continuous_normal_prob. +have := (@continuous_measurable_fun _ (fun x => normal_prob x 1 V)). + have := measurable_normal_prob2. + apply/measurability + admit. + apply/measurable_EFinP. + admit. +- exact: measurableT. +- exact: ltr01. +(* rhs *) +rewrite [in RHS]letin'E/=. +under [in RHS]eq_integral. + move=> u _. + rewrite letin'E/=. + under eq_integral. + move=> x _. + rewrite letin'E/=. + rewrite integral_normal_prob_dirac. + over. + over. +rewrite ge0_integral_mscale; first last. +- move=> ? _. + exact: integral_ge0. +- exact: measurable_cst. +- exact: measurableT. +rewrite integral_dirac/=; first last. +- exact: measurable_cst. +- exact: measurableT. +rewrite diracT mul1e. +rewrite sub0r -expRM mul1r. +rewrite [in RHS]integral_normal_prob; first last. +- (* ok *) admit. +- (* ? *)admit. (* TODO1: ish *) +- exact: measurableT. +rewrite -ge0_integralZl; first last. +- by []. +- move=> ? _. + apply: mule_ge0 => //. + rewrite lee_fin. + exact: normal_pdf_ge0. +- apply: emeasurable_funM. + + (* TODO1: ish *) admit. + + (* ? *)admit. +- exact: measurableT. +(* eq_integral *) +apply: eq_integral. +move=> x _. +rewrite [in LHS]muleAC. +rewrite [in RHS](muleC (normal_prob1 x V)) muleA. +congr *%E. +rewrite [in LHS]ger0_norm; last first. +- (* ok *)admit. +rewrite [in RHS]ger0_norm; last first. +- (* ok *)admit. +rewrite /normal_pdf mul1r subr0 divr1. +(* ? *) (* TODO2: ish *) +rewrite -!EFinM. +congr EFin. +rewrite mulrA. +rewrite mulrAC. +rewrite -expRD. +rewrite [RHS]mulrA. +rewrite [RHS]mulrAC. +rewrite -expRD. +congr *%R. +- congr expR. + rewrite [in RHS]expr_div_n. + rewrite -[in RHS](@sqrtrV _ 2)//. + rewrite sqr_sqrtr//. + (* ? *) + admit. +admit. +Admitted. +*) + +(* ref : https://homes.luddy.indiana.edu/ccshan/rational/equational-handout.pdf + * p.2, equation (9) + * sum independent random variables that are normally distributed + *) +(* +Lemma helloRight1_to_2 : execD helloRight1 = execD helloRight2. +Proof. +apply: congr_normalize => y V. +(* TODO: split rewriting on LHS and RHS *) +(* lhs *) +rewrite [in LHS]execP_letin. +rewrite !execP_score. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow_real/=. +rewrite (@execD_bin _ _ binop_minus)/=. +rewrite execD_real. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow/=. +rewrite exp_var'E/= (execD_var_erefl "y0")/=. +rewrite !execD_real/=. +rewrite execP_letin. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite !(@execD_bin _ _ binop_mult)/=. +rewrite exp_var'E/= (execD_var_erefl "y0")/= !execD_real/=. +rewrite execP_letin. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite exp_var'E/= (execD_var_erefl "x")/=. +rewrite execP_return/=. +rewrite exp_var'E/= (execD_var_erefl "z")/=. +(* rhs *) +rewrite execP_letin. +rewrite execP_score/=. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_real/=. +rewrite execD_pow_real/=. +rewrite (@execD_bin _ _ binop_minus)/= execD_real/=. +rewrite (@execD_bin _ _ binop_mult)/= execD_real/=. +rewrite exp_var'E/=. +rewrite execD_pow/=. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite (@execD_bin _ _ binop_mult)/= execD_real/=. +rewrite exp_var'E/= 2!(execD_var_erefl "y0")/=. +(* lhs *) +rewrite letin'E/=. +under eq_integral. + move=> x _. + rewrite letin'E/=. + under eq_integral. + move=> z _. + rewrite letin'E/=. + rewrite integral_normal_prob_dirac//; . + admit. + over. + over. +rewrite ge0_integral_mscale//=; last first. + move=> ? _; exact: integral_ge0. +rewrite integral_dirac//. +rewrite diracT mul1e. +(* rhs *) +rewrite letin'E/=. +rewrite ge0_integral_mscale//=. +rewrite integral_dirac//= diracT mul1e. +Abort. +*) + +End helloRight_proof. + +(* TODO: move *) +Section exponential_pdf. +Context {R : realType}. +Notation mu := lebesgue_measure. +Variable (mean : R). +Hypothesis mean0 : (0 < mean)%R. + +Definition exponential_pdf' (x : R) := (mean^-1 * expR (- x / mean))%R. +Definition exponential_pdf := exponential_pdf' \_ `[0%R, +oo[. + +Lemma exponential_pdf_ge0 (x : R) : (0 <= exponential_pdf x)%R. +Proof. +apply: restrict_ge0 => {}x _. +apply: mulr_ge0; last exact: expR_ge0. +by rewrite invr_ge0 ltW. +Qed. + +End exponential_pdf. + +Definition exponential {R : realType} (m : R) + of \int[@lebesgue_measure R]_x (exponential_pdf m x)%:E = 1%E : set _ -> \bar R := + fun V => (\int[lebesgue_measure]_(x in V) (exponential_pdf m x)%:E)%E. + +Section exponential. +Context {R : realType}. +Local Open Scope ring_scope. + +Notation mu := lebesgue_measure. +Variable (mean : R). +Hypothesis mean0 : (0 < mean)%R. + +Hypothesis integrable_exponential : forall (m : R), (0 < m)%R -> + mu.-integrable [set: _] (EFin \o (exponential_pdf m)). + +Hypothesis integral_exponential_pdf : forall (m : R), (0 < m)%R -> + (\int[mu]_x (exponential_pdf m x)%:E = 1)%E. + +Local Notation exponential := (exponential (integral_exponential_pdf mean0)). + +Let exponential0 : exponential set0 = 0%E. +Proof. by rewrite /exponential integral_set0. Qed. + +Let exponential_ge0 A : (0 <= exponential A)%E. +Proof. +by rewrite /exponential integral_ge0//= => x _; rewrite lee_fin exponential_pdf_ge0. +Qed. + +Let exponential_sigma_additive : semi_sigma_additive exponential. +Proof. +Admitted. + +HB.instance Definition _ := isMeasure.Build _ _ _ + exponential exponential0 exponential_ge0 exponential_sigma_additive. + +Let exponential_setT : exponential [set: _] = 1%E. +Proof. by rewrite /exponential integral_exponential_pdf. Qed. + +HB.instance Definition _ := @Measure_isProbability.Build _ _ R exponential exponential_setT. + +End exponential. diff --git a/theories/probability.v b/theories/probability.v index 440a509b52..49805ed32f 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -2941,3 +2941,1391 @@ by rewrite -!EFin_beta_fun -EFinM divff// gt_eqF// beta_fun_gt0. Qed. End beta_prob_bernoulliE. + +Section near_lt_lim. +Variable R : realFieldType. +Implicit Types u : R ^nat. + +Lemma near_lt_lim u (M : R) : + (\forall N \near \oo, {in [set n | (N <= n)%N] &, nondecreasing_seq u}) -> + cvgn u -> M < limn u -> \forall n \near \oo, M <= u n. +Proof. +move=> [] N _ Hnear. +move=> cu Ml; have [[n Mun]|/forallNP Mu] := pselect (exists n, M <= u n). + exists (maxn N n) => //. + move=> k/=. + rewrite geq_max => /andP. +(* + near=> m; suff : u n <= u m by exact: le_trans. + apply/(Hnear m). + near: m; exists n.+1 => // p q; apply/(Hnear n)/ltnW => //. + + +have {}Mu : forall x, M > u x by move=> x; rewrite ltNge; apply/negP. +have : limn u <= M by apply: limr_le => //; near=> m; apply/ltW/Mu. +by move/(lt_le_trans Ml); rewrite ltxx. +Unshelve. all: by end_near. Qed. +*) +Abort. + +End near_lt_lim. + +Section near_ereal_nondecreasing_is_cvgn. + +Let G N := ([set n | (N <= n)%N]). + +Lemma near_ereal_nondecreasing_is_cvgn (R : realType) (u_ : (\bar R) ^nat) : + (\forall N \near \oo, {in G N &, nondecreasing_seq u_ }) -> cvgn u_. +Proof. +move=> [N _ H]. +apply/cvg_ex. +exists (ereal_sup (range (fun n => u_ (n + N)))). +rewrite -(cvg_shiftn N). +apply: ereal_nondecreasing_cvgn. +move=> n m nm. +apply: (H N); rewrite /G ?inE//=. +- exact: leq_addl. +- exact: leq_addl. +- exact: leq_add. +Qed. + +Lemma near_ereal_nondecreasing_cvgn (R : realType) (u_ : (\bar R)^nat) : +(* + (\forall N \near \oo, {in G N &, nondecreasing_seq u_ }) + -> u_ @ \oo --> limn u_. (* ereal_sup range ? *) +*) +\forall N \near \oo, {in G N &, nondecreasing_seq u_ } + -> u_ @ \oo --> ereal_sup (range (fun n => u_ (n + N))). +Proof. +near=> N. +(* +move=> [N _ H]. +apply/cvg_ex. +exists (limn (fun n => u_ (n + N))). +rewrite -(cvg_shiftn N). +apply: ereal_nondecreasing_cvgn. +*) +Abort. + + +End near_ereal_nondecreasing_is_cvgn. + +Section near_monotone_convergence. +Local Open Scope ereal_scope. + +Context d (T : measurableType d) (R : realType). +Variable mu : {measure set T -> \bar R}. +Variables (D : set T) (mD : measurable D) (g' : (T -> \bar R)^nat). +Hypothesis mg' : forall n, measurable_fun D (g' n). +Hypothesis near_g'0 : \forall n \near \oo, forall x, D x -> 0 <= g' n x. +Hypothesis near_nd_g' : \forall N \near \oo, (forall x : T, D x -> + {in [set k| (N <= k)%N]&, {homo g'^~ x : n m / (n <= m)%N >-> (n <= m)%E}}). +Let f' := fun x => limn (g'^~ x). + +Lemma near_monotone_convergence : +(\int[mu]_(x in D) (fun x0 : T => limn (g'^~ x0)) x)%E = +limn (fun n : nat => (\int[mu]_(x in D) g' n x)%E). +Proof. +have [N0 _ H0] := near_g'0. +have [N1 _ H1] := near_nd_g'. +pose N := maxn N0 N1. +under eq_integral. + move=> x; rewrite inE/= => Dx. + have <- : limn (fun n : nat => g' (n + N) x) = limn (g'^~ x). + apply/cvg_lim => //. + rewrite (cvg_shiftn _ (g'^~ x) _). + apply: (@near_ereal_nondecreasing_is_cvgn _ (g'^~ x)). + exists N1 => //. + move=> n /= N1n. + exact: H1. + over. +apply/esym/cvg_lim => //. +rewrite -(cvg_shiftn N). +apply: cvg_monotone_convergence => //. + move=> n x Dx. + apply: H0 => //=. + apply: (leq_trans (leq_maxl N0 N1)). + exact: leq_addl. +move=> x Dx n m nm. +apply: (H1 N) => //; rewrite ?inE/=. +- exact: leq_maxr. +- exact: leq_addl. +- exact: leq_addl. +- exact: leq_add. +Qed. + +Lemma cvg_near_monotone_convergence : + \int[mu]_(x in D) g' n x @[n \oo] --> \int[mu]_(x in D) f' x. +Proof. +have [N0 _ Hg'0] := near_g'0. +have [N1 _ Hndg'] := near_nd_g'. +pose N := maxn N0 N1. +have N0N : (N0 <= N)%N by apply: (leq_maxl N0 N1). +have N1N : (N1 <= N)%N by apply: (leq_maxr N0 N1). +have g'_ge0 n x : D x -> (N <= n)%N -> 0 <= g' n x. + move=> + Nn. + apply: Hg'0 => /=. + exact: (leq_trans N0N). +have ndg' n m x : D x -> (N <= n)%N -> (n <= m)%N -> g' n x <= g' m x. + move=> Dx Nn nm. + apply: (Hndg' N); rewrite ?inE//=. + exact: leq_trans nm. +rewrite near_monotone_convergence. +apply: near_ereal_nondecreasing_is_cvgn. +exists N => //. +move=> k/= Nk n m; rewrite !inE/= => kn km nm. +apply: ge0_le_integral => // t Dt; [| |]. +- apply: g'_ge0 => //. + exact: leq_trans kn. +- apply: g'_ge0 => //. + exact: leq_trans km. +- apply: ndg' => //. + exact: leq_trans kn. +Qed. + +End near_monotone_convergence. + +Section exp_coeff_properties. +Context {R : realType}. + +(* not used, TODO:PR *) +Lemma exp_coeff_gt0 (x : R) n : 0 < x -> 0 < exp_coeff x n. +Proof. +move=> x0. +rewrite /exp_coeff/=. +apply: divr_gt0. + exact: exprn_gt0. +rewrite (_:0%R = 0%:R)// ltr_nat. +exact: fact_gt0. +Abort. + +Lemma series_exp_coeff_near_ge0 (x : R) : + \forall n \near \oo, 0 <= (series (exp_coeff x)) n. +Proof. +apply: (cvgr_ge (expR x)); last exact: expR_gt0. +exact: is_cvg_series_exp_coeff. +Abort. + +Lemma normr_exp_coeff_near_nonincreasing (x : R) : + \forall n \near \oo, + `|exp_coeff x n.+1| <= `|exp_coeff x n|. +Proof. +exists `|archimedean.Num.Def.ceil x |%N => //. +move=> n/= H. +rewrite exp_coeffE. +rewrite exprS mulrA normrM [leRHS]normrM ler_pM//. +rewrite factS mulnC natrM invfM -mulrA normrM ger_pMr; last first. + rewrite normr_gt0. + by rewrite invr_neq0//. +rewrite normrM normfV. +rewrite ler_pdivrMl; last first. + rewrite normr_gt0. + by rewrite lt0r_neq0. +rewrite mulr1. +apply: (le_trans (abs_ceil_ge _)). +rewrite gtr0_norm//. +by rewrite ler_nat ltnS. +Qed. + +Lemma exp_coeff2_near_nondecreasing (x : R) : + \forall N \near \oo, nondecreasing_seq (fun n => (series (exp_coeff x) (2 * (n + N))%N)). +Proof. +have := normr_exp_coeff_near_nonincreasing x. +move=> [N _] Hnear. +exists N => //n/= Nn. +apply/nondecreasing_seqP => k. +rewrite /series/=. +have N0 : (0 <= N)%N by []. +rewrite addSn mulnS add2n. +rewrite !big_nat_recr//=. +rewrite -addrA lerDl. +rewrite -[X in _ <= _ + X]opprK subr_ge0. +rewrite (le_trans (ler_norm _))// normrN. +have : (N <= (2 * (k + n)))%N. + rewrite mulnDr -(add0n N) leq_add//. + by rewrite mulSn mul1n -(add0n N) leq_add. +move/Hnear => H. +apply: (le_trans H). +rewrite ler_norml lexx andbT. +suff Hsuff : 0 <= exp_coeff x (2 * (k + n))%N. + by apply: (le_trans _ Hsuff); rewrite lerNl oppr0. +rewrite /exp_coeff/=. +apply: mulr_ge0 => //. +apply: exprn_even_ge0. +by rewrite mul2n odd_double. +Qed. + +Lemma exp_coeff2_near_in_increasing (x : R) : + \forall N \near \oo, {in [set k | (N <= k)%N] &, +nondecreasing_seq (fun n => (series (exp_coeff x) (2 * n)%N))}. +Proof. +have := normr_exp_coeff_near_nonincreasing x. +move=> [N _] Hnear. +exists N => //k/= Nk. +move=> n m; rewrite !inE/= => kn km nm. +have kn2 : (2 * k <= 2 * n)%N by rewrite leq_pmul2l. +have km2 : (2 * k <= 2 * m)%N by rewrite leq_pmul2l. +rewrite /series/=. +rewrite (big_cat_nat _ kn2)//=. +rewrite (big_cat_nat _ km2)//=. +rewrite lerD2. +have nm2 : (2 * n <= 2 * m)%N by rewrite leq_pmul2l. +rewrite (big_cat_nat _ nm2)//=. +rewrite lerDl. +rewrite -(add0n (2 * n)%N). +rewrite big_addn. +rewrite -mulnBr. +elim: (m - n)%N. + rewrite muln0. + rewrite big_mkord. + by rewrite big_ord0. +move=> {km nm km2 nm2} {}m IH. +rewrite mul2n. +rewrite doubleS. +rewrite big_nat_recr//=. +rewrite big_nat_recr//=. +rewrite -addrA. +rewrite addr_ge0//. + by rewrite -mul2n. +rewrite -[X in _ <= _ + X]opprK subr_ge0. +rewrite (le_trans (ler_norm _))// normrN. +rewrite -mul2n addSn -mulnDr. +have : (N <= (2 * (m + n)))%N. + rewrite mulnDr -(add0n N) leq_add//. + by rewrite (leq_trans _ kn2)// (leq_trans Nk)// leq_pmull. +move/Hnear => H. +apply: (le_trans H). +rewrite ler_norml lexx andbT. +suff Hsuff : 0 <= exp_coeff x (2 * (m + n))%N. + by apply: (le_trans _ Hsuff); rewrite lerNl oppr0. +rewrite /exp_coeff/=. +apply: mulr_ge0 => //. +apply: exprn_even_ge0. +by rewrite mul2n odd_double. +Qed. + +End exp_coeff_properties. + +Section normal_density. +Context {R : realType}. +Local Open Scope ring_scope. +Local Import Num.ExtraDef. + +(* NB: If s = 0 then normal_pdf0 is cst 0 for all m and x + * because divr is defined by x / 0 = 0. + *) +Definition normal_pdf0 (m s x : R) : R := + (sqrtr (s ^+2 * pi *+ 2))^-1 * expR (- (x - m) ^+ 2 / (s ^+ 2*+ 2)). + +Lemma normal_pdf0_center m s : normal_pdf0 m s = normal_pdf0 0 s \o center m. +Proof. +apply/funext => x/=. +by rewrite /normal_pdf0 subr0. +Qed. + +Lemma normal_pdf0_ge0 m s x : 0 <= normal_pdf0 m s x. +Proof. by rewrite mulr_ge0 ?expR_ge0// invr_ge0 mulr_ge0. Qed. + +Lemma normal_pdf0_gt0 m s x : s != 0 -> 0 < normal_pdf0 m s x. +Proof. +move=> s0. +rewrite mulr_gt0 ?expR_gt0// invr_gt0//. +by rewrite sqrtr_gt0 pmulrn_rgt0// pmulr_rgt0// ?pi_gt0// exprn_even_gt0/=. +Qed. + +Lemma measurable_normal_pdf0 m s : measurable_fun setT (normal_pdf0 m s). +Proof. +apply: measurable_funM => //=; apply: measurableT_comp => //=. +apply: measurable_funM => //=; apply: measurableT_comp => //=. +apply: measurableT_comp (exprn_measurable _) _ => /=. +exact: measurable_funD. +Qed. + +Lemma continuous_normal_pdf0 m s : continuous (normal_pdf0 m s). +Proof. +move=> x. +apply: cvgM; first exact: cvg_cst. +apply: (@cvg_comp _ R^o _ _ _ _ + (nbhs (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)))); last exact: continuous_expR. +apply: cvgM; last exact: cvg_cst. +apply: (@cvgN _ _ _ _ _ _ ((x - m) ^+ 2 : R^o)). +apply: (@cvg_comp _ _ _ _ (fun x : R => x ^+ 2)%R _ (nbhs (x - m))). + apply: (@cvgB _ R^o) => //. + exact: cvg_cst. +exact: sqr_continuous. +Qed. + +Lemma normal_pdf0_ub m s x : normal_pdf0 m s x <= (sqrtr (s ^+ 2 * pi *+ 2))^-1. +Proof. +rewrite -[leRHS]mulr1. +rewrite ler_wpM2l ?invr_ge0// ?sqrtr_ge0. +rewrite -[leRHS]expR0 ler_expR mulNr oppr_le0 mulr_ge0// ?sqr_ge0//. +by rewrite invr_ge0 mulrn_wge0// sqr_ge0. +Qed. + +End normal_density. + +Section normal_probability. +Variables (R : realType) (m sigma : R). +Local Open Scope ring_scope. +Notation mu := lebesgue_measure. + +Definition normal_pdf_part x := (expR (- (x - m) ^+ 2 / (sigma ^+ 2 *+ 2))). + +Let measurable_normal_pdf_part : measurable_fun setT normal_pdf_part. +Proof. +apply: measurableT_comp => //=. +apply: measurable_funM => //=. +apply: measurableT_comp => //=. +apply: measurable_funX => //=. +exact: measurable_funD. +Qed. + +Let integral_normal_pdf_part : sigma != 0 -> + (\int[mu]_x (normal_pdf_part x)%:E = + (Num.sqrt (sigma ^+ 2 * pi *+ 2))%:E)%E. +Proof. +move=> sigma_gt0. +pose F (x : R^o) := (x - m) / (`|sigma| * Num.sqrt 2). +have F'E : F^`()%classic = cst (`|sigma| * Num.sqrt 2)^-1. + apply/funext => x; rewrite /F derive1E deriveM// deriveD// derive_cst scaler0. + by rewrite add0r derive_id derive_cst addr0 scaler1. +have := @integralT_gauss R. +rewrite (@increasing_ge0_integration_by_substitutionT _ F gauss_fun)//=; first last. +- by move=> x; rewrite gauss_fun_ge0. +- exact: continuous_gauss_fun. +- apply/gt0_cvgMly; last exact: cvg_addrr. + by rewrite invr_gt0// mulr_gt0// normr_gt0. +- apply/gt0_cvgMlNy; last exact: cvg_addrr_Ny. + by rewrite invr_gt0// mulr_gt0// normr_gt0. +- by rewrite F'E; exact: is_cvg_cst. +- by rewrite F'E; exact: is_cvg_cst. +- by rewrite F'E => ?; exact: cvg_cst. +- by move=> x y; rewrite /F ltr_pM2r ?invr_gt0 ?mulr_gt0 ?normr_gt0// ltrBlDr subrK. +move=> /(congr1 (fun x => x * (`|sigma| * Num.sqrt 2)%:E)%E). +rewrite -ge0_integralZr//=; last first. + by move=> x _; rewrite lee_fin mulr_ge0//= ?gauss_fun_ge0// F'E/= invr_ge0. + rewrite F'E; apply/measurable_EFinP/measurable_funM => //. + apply/measurableT_comp => //; first exact: measurable_gauss_fun. + by apply: measurable_funM => //; exact: measurable_funD. +move=> int_gauss. +apply: eq_trans. + apply: eq_trans; last exact: int_gauss. + apply: eq_integral => /= x _. + rewrite F'E !fctE/= EFinM -muleA -EFinM mulVf ?mulr1; last first. + by rewrite gt_eqF ?mulr_gt0 ?normr_gt0. + rewrite /gauss_fun /F exprMn exprVn -[in RHS]mulNr. + by rewrite exprMn real_normK ?num_real// sqr_sqrtr// mulr_natr. +rewrite -mulrnAl sqrtrM ?mulrn_wge0 ?sqr_ge0//. +by rewrite -[in RHS]mulr_natr sqrtrM ?sqr_ge0// sqrtr_sqr !EFinM muleC. +Qed. + +Local Notation normal_pdf := (normal_pdf m sigma). +Local Notation normal_prob := (normal_prob m sigma). + +Let normal0 : normal_prob set0 = 0%E. +Proof. by rewrite /normal_prob integral_set0. Qed. + +Let normal_ge0 A : (0 <= normal_prob A)%E. +Proof. +rewrite /normal_prob integral_ge0//= => x _. +by rewrite lee_fin normal_pdf_ge0 ?ltW. +Qed. + +Let normal_sigma_additive : semi_sigma_additive normal_prob. +Proof. +move=> /= F mF tF mUF. +rewrite /normal_prob/= integral_bigcup//=; last first. + apply/integrableP; split. + apply/measurable_funTS/measurableT_comp => //. + exact: measurable_normal_pdf. + rewrite (_ : (fun x => _) = EFin \o normal_pdf); last first. + by apply/funext => x; rewrite gee0_abs// lee_fin normal_pdf_ge0 ?ltW. + apply: le_lt_trans. + apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. + by apply/measurable_EFinP; exact: measurable_normal_pdf. + by move=> ? _; rewrite lee_fin normal_pdf_ge0 ?ltW. + by rewrite integral_normal_pdf // ltey. +apply: is_cvg_ereal_nneg_natsum_cond => n _ _. +by apply: integral_ge0 => /= x ?; rewrite lee_fin normal_pdf_ge0 ?ltW. +Qed. + +HB.instance Definition _ := isMeasure.Build _ _ _ + normal_prob normal0 normal_ge0 normal_sigma_additive. + +Let normal_setT : normal_prob [set: _] = 1%E. +Proof. by rewrite /normal_prob integral_normal_pdf. Qed. + +HB.instance Definition _ := + @Measure_isProbability.Build _ _ R normal_prob normal_setT. + +End normal_probability. + +(* TODO: PR *) +Section shift_properties. + +Variable R : realType. + +Local Open Scope ring_scope. + +Notation mu := lebesgue_measure. + +Lemma derive_shift (v k : R) : 'D_v (shift k : R^o -> R^o) = cst v. +Proof. +apply/funext => x/=. +by rewrite deriveD// derive_id derive_cst addr0. +Qed. + +Lemma is_derive_shift x v (k : R) : is_derive x v (shift k : R^o -> R^o) v. +Proof. +split => //. +by rewrite derive_val addr0. +Qed. + +(* TODO: In integration_by_substitution, (f : R -> R) => (f : R -> \bar R) *) +Lemma ge0_integration_by_substitution_shift_itvNy (f : R -> R) (r e : R) : +{within `]-oo, r + e], continuous f} -> +{in `]-oo, r + e[, forall x : R, 0 <= f x} -> +(\int[mu]_(x in `]-oo, (r + e)%R]) (f x)%:E = +\int[mu]_(x in `]-oo, r]) ((f \o (shift e)) x)%:E)%E. +Proof. +move=> cf f0. +have := (derive_shift 1 e). +have <- := (funext (@derive1E R _ (shift e : R^o -> R^o))). +move=> dshiftE. +rewrite (@increasing_ge0_integration_by_substitutionNy _ (shift e))//; first last. +- exact: cvg_addrr_Ny. +- split. + move=> x _. + exact/ex_derive. + apply: cvg_at_left_filter. + apply: cvgD => //. + exact: cvg_cst. +- rewrite dshiftE. + exact: cvg_cst. +- rewrite dshiftE. + exact: is_cvg_cst. +- rewrite dshiftE. + move=> ? _; apply: cst_continuous. +- by move=> x y _ _ xy; rewrite ltr_leD. +by rewrite dshiftE mulr1/=. +Qed. + +Lemma ge0_integration_by_substitution_shift_itvy (f : R -> R) (r e : R) : +{within `[r + e, +oo[, continuous f} -> +{in `]r + e, +oo[, forall x : R, 0 <= f x} -> +(\int[mu]_(x in `[r + e, +oo[) (f x)%:E = +\int[mu]_(x in `[r, +oo[) ((f \o (shift e)) x)%:E)%E. +Proof. +move=> cf f0. +have := (derive_shift 1 e). +have <- := (funext (@derive1E R _ (shift e : R^o -> R^o))). +move=> dshiftE. +rewrite (@increasing_ge0_integration_by_substitutiony _ (shift e))//=; first last. +- exact: cvg_addrr. +- split. + move=> x _. + exact/ex_derive. + apply: cvg_at_right_filter. + apply: cvgD => //. + exact: cvg_cst. +- rewrite dshiftE. + exact: is_cvg_cst. +- rewrite dshiftE. + exact: is_cvg_cst. +- rewrite dshiftE. + move=> ? _; apply: cst_continuous. +- by move=> x y _ _ xy; rewrite ltr_leD. +by rewrite dshiftE mulr1/=. +Qed. + +End shift_properties. + +Section normal_kernel. + +Variable R : realType. +Variables s : R. +Hypothesis s0 : s != 0. +Local Open Scope ring_scope. +Notation mu := lebesgue_measure. + +Let normal_pdfE m x : normal_pdf m s x = + (Num.sqrt (s^+2 * pi *+ 2))^-1 * expR (- (x - m) ^+ 2 / (s^+2 *+ 2)). +Proof. +rewrite /normal_pdf ifF//. +exact/negP/negP. +Qed. + +Local Definition normal_prob2 := (fun m => normal_prob m s) : _ -> pprobability _ _. + +Lemma bij_shift x : bijective (id \+ @cst R R x). +Proof. +apply: (@Bijective _ _ _ (id \- cst x)). +- by move=> z;rewrite /=addrK. +- by move=> z; rewrite /= subrK. +Qed. + +Lemma shift_ocitv (x a b : R) : + (shift x) @` `]a, b]%classic = `]a + x, b + x]%classic. +Proof. +rewrite eqEsubset; split => r/=. + move=> [r' + <-]. + rewrite in_itv/=; move/andP => [ar' r'b]. + by rewrite in_itv/=; apply/andP; split; rewrite ?lerD2 ?ltrD2. +rewrite in_itv/=; move/andP => [axr rbx]. +exists (r - x); last by rewrite subrK. +rewrite in_itv/=; apply/andP; split. +- by rewrite ltrBrDr. +- by rewrite lerBlDr. +Qed. + +Lemma shift_preimage (x : R) U : + (shift x) @^-1` U = (shift (- x)) @` U. +Proof. +rewrite eqEsubset; split => r. + rewrite /= => Urx. + by exists (r + x) => //; rewrite addrK. +by move=> [z Uz <-]/=; rewrite subrK. +Qed. + +Lemma pushforward_shift_itv (mu : measure (measurableTypeR R) R) (a b x : R) : + (pushforward mu (fun z => z + x) + `]a, b]) = + mu `]a - x, b - x]%classic. +Proof. +rewrite /pushforward. +rewrite shift_preimage. +by rewrite shift_ocitv. +Qed. + +Lemma pushforward_shift_measurable (mu : measure (measurableTypeR R) R) (x : R) (U : set R) : + (pushforward mu (fun z => z + x) + U) = + mu ((center x) @` U). +Proof. +by rewrite /pushforward shift_preimage. +Qed. + +From mathcomp Require Import charge. +Open Scope charge_scope. + +(* +Lemma radon_nikodym_crestr_fin U (mU : measurable U) +(Uoo : (@lebesgue_measure R U < +oo)%E) : + ae_eq lebesgue_measure setT ('d charge_of_finite_measure (mfrestr mU Uoo) '/d + [the sigma_finite_measure _ _ of @lebesgue_measure R]) + (EFin \o \1_U). +Proof. +apply: integral_ae_eq => //=. +- admit. +- admit. +move=> E _ mE. +rewrite -Radon_Nikodym_integral. +rewrite integral_indic/=. +by rewrite /mfrestr/mrestr setIC. +Admitted. +*) + +(* +Lemma radon_nikodym_crestr U (mU : measurable U) : + ae_eq lebesgue_measure setT ('d charge_of_finite_measure (mfrestr mU Uoo) '/d + [the sigma_finite_measure _ _ of @lebesgue_measure R]) + (EFin \o \1_U). +Proof. +*) + +(* +rewrite [RHS](_:_= ('d charge_of_finite_measure (mfrestr mU Uoo) '/d + [the sigma_finite_measure _ _ of @lebesgue_measure R]) + (EFin \o \1_U) + move=> x _. + rewrite epatch_indic. + rewrite -radon_nikodym_crestr. +rewrite [RHS]integral_mkcond. +under [RHS]eq_integral do rewrite epatch_indic. + +rewrite -integral_pushforward. +apply: eq_integral. +move=> x _. +Admitted. +*) + +(* +Lemma normal_shift0 x : +normal_prob2 x = + @pushforward _ _ _ + (measurableTypeR R) _ (normal_prob2 0%R) (fun z => z + x) + :> (set R -> \bar R). +Proof. +apply: funext. +move=> U. +rewrite /normal_prob2/=. +rewrite /pushforward/=. +rewrite /normal_prob. +rewrite shift_preimage. +rewrite integration_by_substitution_shift/=. +apply: eq_integral. +move=> z Uz. +congr EFin. +rewrite /normal_pdf/=. +rewrite ifF; last exact/negP/negP. +rewrite ifF; last exact/negP/negP. +rewrite {2}/normal_fun. +by rewrite subr0. +Qed. +*) + +(* +Lemma measurable_normal_prob2_ocitv a b: + measurable_fun [set: R] (normal_prob2 ^~ `]a, b]%classic). +Proof. +apply: (@measurability _ _ _ _ _ _ + (@pset _ _ _ : set (set (pprobability _ R)))) => //. +move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. + +rewrite /normal_prob2/=. +rewrite /normal_prob. + +under [X in measurable_fun _ X]eq_fun. + move=> x. + rewrite (_: normal_kernel _ _ = (fine (normal_kernel x `]a, b]%classic))%:E); last first. + rewrite fineK//. + rewrite ge0_fin_numE//. + apply: (@le_lt_trans _ _ 1%E); last exact: ltey. + exact: probability_le1. + rewrite normal_shift0/=. + over. +apply: measurableT_comp; last by []. +apply: measurableT_comp; first exact: EFin_measurable. +rewrite /=. +under [X in measurable_fun _ X]eq_fun. + move=> x. + rewrite /normal_prob. +(pushforward_shift_itv (normal_kernel 0) a b x). +apply: continuous_measurable_fun. +*) + +Local Open Scope ereal_scope. + +Lemma integral_normal_prob (f : R -> \bar R) (m : R) U : measurable U -> measurable_fun U f -> + \int[normal_prob2 m]_(x in U) `|f x| < +oo -> + \int[normal_prob2 m]_(x in U) f x = \int[mu]_(x in U) (f x * (normal_pdf m s x)%:E) :> \bar R. +Proof. +Abort. + +Local Close Scope ereal_scope. + +(* +Lemma continuousT_integralT (f : R -> R -> R) : +(forall l, mu.-integrable setT (fun x => (f l x)%:E)) -> +{ae mu, forall x, {for x, continuous (fun l => f l x)}} -> +(exists g : R -> R, forall l x, `|f l x| <= g x) -> +continuous (fun l => \int[mu]_x f l x). +Proof. +Abort. +*) + + +Lemma normal_prob_fin_num (m : R) U : normal_prob m s U \is a fin_num. +Proof. +Admitted. + +Lemma integrable_normal_pdf0 U z : mu.-integrable U + (fun x : R => (normal_pdf 0 s (x - z))%:E). +Proof. +Abort. + +(* +Lemma shift_continuous (x : R) : continuous (shift x). +Proof. +Admitted. + +Lemma add_continuous (x : R) : continuous (+%R x). +Proof. +Admitted. +*) + +Local Import Num.ExtraDef. + +Local Definition f (m x : R) := + (fun n => let sigma := s ^+ 2 in + ((sqrtr (sigma * pi *+ 2))^-1 * + series (exp_coeff (- ((x - m) ) ^+ 2 / (sigma *+ 2))) (2 * n)%N)%:E). + +Lemma near_f_ge0 : \forall n \near \oo, forall m x, (0 <= f m x n)%E. +Proof. +Admitted. +(* +apply: lt_lim. +- move => n0 n1 n01. + rewrite ler_pM2l; last first. + rewrite invr_gt0. + rewrite sqrtr_gt0. + rewrite pmulrn_rgt0//. + rewrite mulr_gt0//. + by rewrite exprn_even_gt0. + exact: pi_gt0. + set X := (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)). + have [N] := (exp_coeff2_near_in_increasing X). +(* +near=> n. +apply: mulr_ge0. + rewrite invr_ge0. + exact: sqrtr_ge0. +have : series (exp_coeff (- (x - m) ^+ 2 / (s ^+ 2 *+ 2))) (2 * n)%N @[n --> \oo] --> expR (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)). +*) + admit. + admit. +rewrite /expR. +Admitted. +*) + +Lemma near_nd_f : \forall N \near \oo, forall m x, + {in [set n | (N <= n)%N]&, nondecreasing (f m x)}. +Proof. +near=> N. +move=> m x n0 n1; rewrite !inE/= => Nn0 Nn1 n01. +rewrite /f lee_fin ler_pM2l; last first. + rewrite invr_gt0 sqrtr_gt0. + rewrite pmulrn_lgt0//. + rewrite mulr_gt0//. + by rewrite exprn_even_gt0. + exact: pi_gt0. +Admitted. + +Lemma f_first_continuous (n : nat) (x : R) : + continuous (fun m => fine (f m x n)). +Proof. +Admitted. + +Lemma f_second_continuous (n : nat) (x : R) : + continuous (fine \o (f x ^~ n)). +Proof. +Admitted. + +Lemma measurable_f_second (m : R) n : measurable_fun setT (f m ^~ n). +Proof. +Abort. + +Lemma cvg_f m x : (f m x n) @[n --> \oo] --> (normal_pdf m s x)%:E. +Proof. +rewrite /f/normal_pdf. +rewrite ifF; last exact/negP/negP. +apply/cvg_EFin. + exact/nearW. +apply: cvgMr. +have : (2 * n)%N @[n --> \oo] --> \oo. + under eq_cvg do rewrite mulnC. + exact: cvg_mulnr. +move/cvg_comp; apply. +exact: is_cvg_series_exp_coeff. +Qed. + + +Lemma integral_f_fin_num x Ys n : (\int[mu]_(x0 in Ys) f x x0 n)%E \is a fin_num. +Admitted. + +Lemma cvg_integral_f m Ys : measurable Ys -> + (\int[mu]_(x in Ys) f m x n)%E @[n --> \oo] --> normal_prob m s Ys. +Proof. +move=> mYs. +rewrite /normal_prob. +rewrite [X in _ --> X](_ : _ = (\int[mu]_(x in Ys) (limn (f m x)))%E); last first. + apply: eq_integral => y _. + apply/esym/cvg_lim => //. + exact: cvg_f. +apply: cvg_near_monotone_convergence => //. +move=> n. +rewrite /f. +(* ge0_emeasurable_sum *) +- admit. +- near=> n. + move=> x Ysx. + rewrite /f. +admit. +near=> n. +move=> x Ysx. +Admitted. + +Lemma integrable_f x Ys n : mu.-integrable Ys ((f x)^~ n). +Admitted. + +(* outline of proof: + 1. It is enough to prove that `(fun x => normal_prob x s Ys)` is continuous for + all measurable set `Ys`. + 2. Continuity is obtained by continuity under integral from continuity of + `normal_pdf`. + 3. Fix a point `a` in `R` and `e` with `0 < e`. Then take the function + `g : R -> R` as that `g x` is the maximum value of + `normal_pdf a s x` at a point within `e` of `x`. + Then `g x` is equal to `normal_pdf a s 0` if `x` in `ball a e`, + `normal_pdf a s (x - e)` for x > a + e, + and `normal_pdf a s (x + e)` for x < a - e. + 4. Integrability of `g` is checked by calculating integration. + By integration by substitution, the integral of `g` on ]-oo, a - e] + is equal to the integral of `normal_pdf a s` on `]-oo, a], + and it on `[a + e, +oo[ similarly. + So the integral of `g` on ]-oo, +oo[ is the integral of `f` on ]-oo, +oo[ + added by the integral of `normal_pdf a s x` on ]a - e, a + e[ + *) + +Let g' a e : R -> R := fun x => if (x \in (ball a e : set R^o)) then + (Num.sqrt (s ^+ 2 * pi *+ 2))^-1 * + expR (- 0%R ^+ 2 / (s ^+ 2 *+ 2)) +else + (Num.sqrt (s ^+ 2 * pi *+ 2))^-1 * + expR (- (`|x - a| - e) ^+ 2 / (s ^+ 2 *+ 2)). + +Let ballFE_le (a e x : R) : x <= (a - e)%R -> (x \in (ball a e : set R^o)) = false. +Proof. +move=> xae. +apply: memNset. +rewrite ball_itv/= in_itv/=; apply/negP/andP/not_andP; left. +by apply/negP; rewrite -leNgt. +Qed. + +Let ballFE_ge (a e x : R) : a + e <= x -> (x \in (ball a e : set R^o)) = false. +Proof. +move=> xae. +apply: memNset. +rewrite ball_itv/= in_itv/=; apply/negP/andP/not_andP; right. +by apply/negP; rewrite -leNgt. +Qed. + +Lemma g'a0 (a : R) : g' a 0 = normal_pdf0 a s. +Proof. +rewrite /g'. +apply/funext => x. +have /orP [x0|x0] := le_total x a. + rewrite ballFE_le; last by rewrite subr0. + by rewrite subr0 real_normK// num_real. +rewrite ballFE_ge; last by rewrite addr0. +by rewrite subr0 real_normK// num_real. +Qed. + +Lemma mg' a e : measurable_fun setT (g' a e). +Proof. +apply: measurable_fun_if => //. + apply: (measurable_fun_bool true) => /=. + rewrite setTI preimage_mem_true. + exact: measurable_ball. +apply: measurable_funTS => /=. +apply: measurableT_comp => //. +apply: measurableT_comp => //. +apply: measurable_funM => //. +apply: measurableT_comp => //. +apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)) => //. +apply: measurable_funB => //. +apply: measurableT_comp => //. +exact: measurable_funD. +Qed. + +Lemma g'_ge0 a e x : 0 <= g' a e x. +Proof. +rewrite /g'; case: ifP => _. + rewrite -[in leRHS](@subr0 _ 0) -normal_pdfE ?subr0; apply: normal_pdf_ge0. +rewrite -normal_pdfE; exact: normal_pdf_ge0. +Qed. + +Lemma continuous_g' (a e : R) : (0 <= e) -> continuous (g' a e). +Proof. +move=> e0. +have tmp k : k < a - e -> ((`|k - a| - e) ^+ 2 = (k - (a - e)) ^+ 2). + move=> kae. + rewrite -normrN opprB. + have /normr_idP -> : (0 <= a - k). + by rewrite subr_ge0 ltW// (lt_le_trans kae)// gerBl. + by rewrite -sqrrN !opprB addrCA. +have tmp2 k : a + e < k -> ((`|k - a| - e) ^+ 2 = (k - (a + e)) ^+ 2). + move=> kae; rewrite opprD addrA. + have /normr_idP -> // : (0 <= k - a). + by rewrite subr_ge0 ltW// (le_lt_trans _ kae)// lerDl. +apply: (@in1TT R). +rewrite -continuous_open_subspace; last exact: openT. +rewrite (_:[set: R] = + `]-oo, (a - e)%R] `|` `[(a - e)%R, a + e] `|` `[a + e, +oo[); last first. + rewrite -setUitv1// -setUA setUAC setUA -itv_bndbnd_setU//; last first. + by rewrite bnd_simp lerD// ge0_cp. + rewrite -setUitv1// -setUA setUAC setUA -itv_bndbnd_setU//. + by rewrite set_itvE !setTU. +apply: withinU_continuous. + rewrite -setUitv1//. + rewrite -setUA setUCA. + rewrite -itv_bndbnd_setU//; last first. + by rewrite bnd_simp lerD// ge0_cp. + rewrite setUidr//. + move=> _/= ->; rewrite in_itv/=. + by rewrite lerD// ge0_cp. + by apply: interval_closed. + apply: withinU_continuous; first exact: interval_closed. + exact: interval_closed. + apply/continuous_within_itvNycP; split. + move=> x. + rewrite in_itv/= => xae. + apply/(@cvgrPdist_le _ R^o R _ _ (g' a e) (g' a e x)). + rewrite /=. + move=> eps eps0. + near=> t. + have tae : t < a - e. + near: t. + exact: lt_nbhsl. + rewrite /g'. + rewrite !ballFE_le ?(@ltW _ _ _ (a - e))//. + rewrite !tmp//. + move=> {tae}; near: t. + move: eps eps0. + apply/(@cvgrPdist_le _ _ _ (nbhs x)). + exact: continuous_normal_pdf0. + apply/(@cvgrPdist_lt _ R^o). + move=> eps eps0. + near=> t. + rewrite /g' !ballFE_le//. + rewrite -addrAC subrr sub0r normrN; have /normr_idP -> := e0. + rewrite subrr -(subrr (a - e)) tmp//. + near: t; move: eps eps0. + apply/(@cvgrPdist_lt _ R^o). + apply: cvg_at_left_filter. + exact: continuous_normal_pdf0. + move: e0; rewrite le_eqVlt => /predU1P; case => [<- | e0]. + rewrite g'a0. + apply: continuous_subspaceT. + exact: continuous_normal_pdf0. + apply/continuous_within_itvP; first by rewrite -(opprK e) ler_ltB// opprK gtrN. + split. + move=> x xae. + rewrite /continuous_at. + rewrite /g' ifT; last by rewrite ball_itv inE/=. + apply/(@cvgrPdist_le _ R^o). + move=> eps eps0. + near=> t. + rewrite ifT; first by rewrite subrr normr0 ltW. + rewrite ball_itv inE/= in_itv/=; apply/andP; split. + near: t. + apply: lt_nbhsr. + by move: xae; rewrite in_itv/= => /andP[]. + near: t. + apply: lt_nbhsl. + by move: xae; rewrite in_itv/= => /andP[]. + apply/(@cvgrPdist_le _ R^o). + move=> eps eps0. + near=> t. + rewrite /g' ballFE_le// ifT; last first. + rewrite ball_itv inE/= in_itv/=; apply/andP => []; split => //. + near: t. + apply: nbhs_right_lt. + by rewrite -(opprK e) ler_ltB// opprK gtrN. + rewrite addrAC subrr sub0r normrN; have /ltW/normr_idP -> := e0. + by rewrite !subrr normr0 ltW. + apply/(@cvgrPdist_le _ R^o). + move=> eps eps0. + near=> t. + rewrite /g' ballFE_ge// ifT; last first. + rewrite ball_itv inE/= in_itv/=; apply/andP => []; split => //. + near: t. + apply: nbhs_left_gt. + by rewrite -(opprK e) ler_ltB// opprK gtrN. + rewrite addrAC subrr add0r; have /ltW/normr_idP -> := e0. + by rewrite !subrr normr0 ltW. +apply/continuous_within_itvcyP; split. + move=> x. + rewrite in_itv/= andbT => aex. + apply/(@cvgrPdist_le _ R^o). + rewrite /=. + move=> eps eps0. + near=> t. + have tae : a + e < t. + near: t. + exact: lt_nbhsr. + rewrite /g'. + rewrite !ballFE_ge ?(@ltW _ _ (a + e)%E)//. + rewrite !tmp2// ?(@ltW _ _ (a + e)). + move=> {tae}; near: t. + move: eps eps0. + apply/(@cvgrPdist_le _ _ _ (nbhs x)). + exact: continuous_normal_pdf0. +apply/(@cvgrPdist_le _ R^o). +move=> eps eps0. +near=> t. +rewrite /g' !ballFE_ge//. +rewrite addrAC subrr add0r; have /normr_idP -> := e0. +rewrite subrr -(subrr (a + e)). +rewrite tmp2//. +near: t. +move: eps eps0. +apply/cvgrPdist_le. +apply: cvg_at_right_filter. +apply: continuous_normal_pdf0. +Unshelve. all: end_near. Qed. + +Lemma gE_Ny a e : (0 <= e) -> + (\int[mu]_(x in `]-oo, (a - e)%R]) `|g' a e x|%:E = + \int[mu]_(x in `]-oo, a]) `|normal_pdf a s x|%:E)%E. +Proof. +move=> e0. +rewrite ge0_integration_by_substitution_shift_itvNy => /=; first last. +- move=> ? _; exact: normr_ge0. +- apply/continuous_subspaceT. + move=> x. + apply: continuous_comp; first exact: continuous_g'. + exact: (@norm_continuous _ R^o) . +under eq_integral. + move=> x. + rewrite inE/= in_itv/= => xae. + rewrite /g' ballFE_le//; last exact: lerB. + rewrite -(normrN (x - e - a)) !opprB addrA. + have /normr_idP -> : 0 <= a + e - x by rewrite subr_ge0 ler_wpDr. + rewrite -(addrAC _ (- x)) addrK. + rewrite -(sqrrN (a - x)) opprB. + over. +by under [RHS]eq_integral do rewrite normal_pdfE. +Qed. + +Lemma gE_y a e : (0 <= e) -> + (\int[mu]_(x in `[a + e, +oo[) `|g' a e x|%:E = + \int[mu]_(x in `[a, +oo[) `|normal_pdf a s x|%:E)%E. +Proof. +move=> e0. +rewrite ge0_integration_by_substitution_shift_itvy => /=; first last. +- move=> ? _. + exact: normr_ge0. +- apply/continuous_subspaceT. + move=> x. + apply: continuous_comp; first exact: continuous_g'. + exact: (@norm_continuous _ R^o). +under eq_integral. + move=> x. + rewrite inE/= in_itv/= andbT => aex. + rewrite /g' ballFE_ge//; last exact: lerD. + have /normr_idP -> : 0 <= x + e - a by rewrite subr_ge0 ler_wpDr. + rewrite -(addrAC _ (- a)) addrK. + over. +by under [RHS]eq_integral do rewrite normal_pdfE. +Qed. + +Lemma normal_prob_continuous (V : set R) : measurable V -> + continuous (fun m => fine (normal_prob m s V)). +Proof. +move=> mV a. +near (0 : R)^'+ => e. +set g := g' a e. +have mg := mg' a e. +apply: (@continuity_under_integral _ _ _ mu _ _ _ _ (a - e) (a + e) _ _ _ g) => //=. +- rewrite in_itv/=. + by rewrite ltrDl gtrBl andbb. +- move=> x _. + apply: (integrableS measurableT) => //. + exact: integrable_normal_pdf. +- apply/aeW => y _. + move=> x. + under [X in _ _ X]eq_fun do rewrite normal_pdfE -(sqrrN (y - _)) opprB. + exact: continuous_normal_pdf0. +- apply: (integrableS measurableT) => //=. + apply/integrableP; split. + exact/measurable_EFinP. + rewrite -(setUv (ball a e)). + rewrite ge0_integral_setU//=; first last. + exact/disj_setPCl. + rewrite setUv. + apply/measurable_EFinP. + exact: measurableT_comp. + apply: measurableC. + exact: measurable_ball. + exact: measurable_ball. + apply: lte_add_pinfty. + under eq_integral. + move=> x xae. + rewrite /g/g' xae. + rewrite expr0n/= oppr0 mul0r expR0 mulr1. + over. + rewrite integral_cst/=. + apply: lte_mul_pinfty => //. + rewrite ball_itv lebesgue_measure_itv/= ifT -?EFinD ?ltry// lte_fin. + by rewrite ltrBlDr -addrA -ltrBlDl subrr -mulr2n mulrn_wgt0. + exact: measurable_ball. + have -> : (\int[mu]_(x in ~` ball a e) `|g x|%:E = \int[mu]_x `|normal_pdf a s x|%:E)%E. + rewrite ball_itv setCitv ge0_integral_setU//=; first last. + apply/disj_setPRL. + rewrite setCitvl. + apply: subset_itvr; rewrite bnd_simp. + rewrite -{2}(opprK e); apply: ler_ltB => //. + exact: gtrN. + apply: measurable_funTS. + apply/measurable_EFinP. + exact: measurableT_comp. + rewrite gE_Ny// gE_y//. + rewrite -integral_itv_obnd_cbnd; last first. + apply: measurableT_comp => //. + apply: measurable_funTS. + exact: measurable_normal_pdf. + rewrite -ge0_integral_setU/= ?measurable_itv//; first last. + apply/disj_setPRL. + by rewrite setCitvl. + rewrite -setCitvl setUv. + apply/measurable_EFinP. + apply: measurableT_comp => //. + exact: measurable_normal_pdf. + by rewrite -setCitvl setUv. + under eq_integral => x do rewrite (_:_%:E = `|(normal_pdf a s x)%:E|%E)//. + apply/abse_integralP; rewrite //=. + apply/measurable_EFinP. + exact: measurable_normal_pdf. + by rewrite integral_normal_pdf ltry. +move=> x ax. +apply/aeW => y Vy. +have /normr_idP -> : 0 <= normal_pdf x s y by apply: normal_pdf_ge0. +rewrite normal_pdfE /g/g'. +case: ifP. + move=> _. + rewrite ler_pM//. + rewrite expr0n/= oppr0 mul0r ler_expR. + apply: mulr_le0_ge0. + rewrite oppr_le0. + exact: sqr_ge0. + by rewrite invr_ge0 mulrn_wge0// sqr_ge0. +move/negP/negP; rewrite notin_setE/= ball_itv/= in_itv/= => Hy. +rewrite ler_pM//. +rewrite ler_expR. +rewrite !mulNr lerN2. +rewrite ler_pM => //. + exact: sqr_ge0. + by rewrite invr_ge0 mulrn_wge0// sqr_ge0. +move: Hy; move/negP/nandP; rewrite -!leNgt; case => [yae|aey]. + rewrite -normrN opprB. + have /normr_idP -> : 0 <= a - y. + rewrite subr_ge0. + apply: (le_trans yae). + by rewrite gerBl. + rewrite -[leRHS]sqrrN opprB. + rewrite ler_sqr ?nnegrE; first last. + rewrite subr_ge0. + apply/ltW; apply: (le_lt_trans yae). + by move: ax; rewrite in_itv/= => /andP[]. + by rewrite addrAC subr_ge0. + rewrite addrAC lerD2r. + by apply/ltW; move: ax; rewrite in_itv/= => /andP[]. +have /normr_idP -> : 0 <= y - a. + rewrite subr_ge0; apply: le_trans aey. + by rewrite lerDl. +rewrite ler_sqr ?nnegrE; first last. + rewrite subr_ge0. + apply: le_trans aey. + apply/ltW. + by move: ax; rewrite in_itv/= => /andP[]. + by rewrite -addrA -opprD subr_ge0. +rewrite -addrA -opprD. +rewrite lerD2l lerN2. +by rewrite ltW//; move: ax; rewrite in_itv/= => /andP[]. +Unshelve. end_near. + +(* (note for 3.) +have mlimf : measurable_fun Ys (fun x0 => limn [eta f x x0])%E. + admit. +have H1 : {ae mu, forall x0, Ys x0 -> f x x0 x1 @[x1 --> \oo] --> (limn [eta f x x0])%E}. + admit. + have e x : \bar R := (expR (- x ^+ 2 / (s^+2 *+ 2)))%:E. +have : forall M x0, exists x, (`| f x x0 1 | > M)%E. +move=> M x0. + have H2 : mu.-integrable Ys e. + admit. + have H3 : \forall N \near \oo, {ae mu, forall x0 n, (N <= n)%N -> Ys x0 -> (`|f x x0 n| <= e x0)%E}. + admit. + have [Hdc1 Hdc2 Hdc3] := dominated_convergence mYs mf mlimf H1 H2 H3. + exact: Hdc3. +*) +Admitted. + +(* (note for 1.) + admit. + under [X in limn X]eq_fun do rewrite (fineK (integral_f_fin_num _ _ _)). + over. +rewrite /=. +apply: (emeasurable_fun_cvg (fun n : nat => fun x => (\int[mu]_(x0 in Ys) f x x0 n)%E)). + move=> m/=. + under [X in _ _ X]eq_fun do rewrite -(fineK (integral_f_fin_num _ _ _)); apply/measurable_EFinP => /=. + apply: continuous_measurable_fun. + apply: continuousT_integral => /=. + - admit. + - apply: aeW => /= x. + admit. + - exists (cst (Num.sqrt (s^+2 * pi *+ 2))^-1 * 1). + admit. +move=> x _. +rewrite /f. +rewrite -near_monotone_convergence//; first last. +- admit. +- admit. +- move=> n; apply: measurable_funTS; exact: (measurable_f_second x n). +apply: cvg_near_monotone_convergence => //. +- admit. +- admit. +move=> /= y _. +have := exp_coeff2_near_in_increasing (- (y - x) ^+ 2 / (s ^+ 2 *+ 2)). +move=> [] N _ Hnear. +exists N => //. +move=> k/= Nk n' m'; rewrite !inE/= => kn' km' n'm'. +rewrite lee_fin. +rewrite ler_pM2l; last first. + rewrite invr_gt0. + rewrite sqrtr_gt0. + rewrite pmulrn_lgt0//. + rewrite mulr_gt0//. + by rewrite exprn_even_gt0. + exact: pi_gt0. +by apply: (Hnear k) => //; rewrite !inE/=. +Admitted. + +(* +under [X in _ _ X]eq_fun do rewrite -(fineK (normal_prob_fin_num _ Ys)). +apply/measurable_EFinP. +apply: (@measurability _ _ _ _ _ _ (@RGenOpens.G R)). + by rewrite /measurable/=/measurableR RGenOpens.measurableE. +move=> _ -[_ [a [b ->]] <-]. +rewrite setTI. +apply: is_interval_measurable. +move=> x y/=. +*) +move=> _ -[_ [r r01] [Ys mYs <-]] <-. +apply: emeasurable_fun_infty_o => //=. +under [X in _ _ X]eq_fun. + move=> x. + rewrite -(fineK (normal_prob_fin_num x Ys)). + rewrite normal_shift0//=. + rewrite /pushforward. + rewrite shift_preimage. + rewrite /normal_prob/=. + rewrite integration_by_substitution_shift/=. (* TODO *) + over. +apply: measurableT_comp => //=. +apply/measurable_EFinP => /=. +apply: continuous_measurable_fun => /=. +apply: (@continuousT_integral (fun x0 x1 : R => normal_pdf 0 s (x1 - x0)) Ys). +- move=> z. + exact: integrable_normal_pdf0. +- apply: aeW => /=. + rewrite /measurable/=/g_sigma_algebraType/ocitv_type. + move=> x. + apply: continuous_comp => //=. + apply: continuous_comp. + exact: (@opp_continuous _ R^o). + exact: (@add_continuous x). + exact: continuous_normal_pdf. +exists (cst ((Num.sqrt (s^+2 * pi *+ 2))^-1)). +move=> x _ y/=. +rewrite /normal_pdf; case: ifP => [|_]; first by move/negP: s0. + +admit. +Admitted. +*) + +(* +Lemma measurable_normal_prob2' : +forall U : set R, measurable U -> measurable_fun [set: R] (normal_prob2 ^~ U). +Proof. +move=> U mU. +apply: (@measurability _ _ _ _ _ _ + (@pset _ _ _ : set (set (pprobability _ R)))) => //. +move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. +rewrite /normal_kernel/=. + + +under [X in measurable_fun _ X]eq_fun. + move=> x. + rewrite (_: normal_kernel _ _ = (fine (normal_kernel x U))%:E); last first. + admit. + rewrite normal_shift0/=. + over. +apply: measurableT_comp; last by []. +apply: measurableT_comp; first exact: EFin_measurable. +rewrite /=. +apply: continuous_measurable_fun. +(* +apply: continuous_comp. +under [X in continuous X]eq_fun. + move=> x. + rewrite normal_shift0/=. + over. ++\*) +Admitted. +*) + +(* +HB.instance Definition _ := + isKernel.Build _ _ _ _ _ normal_prob2 measurable_normal_prob2. +*) + +Lemma measurable_normal_prob2 : + measurable_fun setT (normal_prob2 : R -> pprobability _ _). +Proof. +apply: (@measurability _ _ _ _ _ _ + (@pset _ _ _ : set (set (pprobability _ R)))) => //. +move=>_ -[_ [r r01] [Ys mYs <-]] <-. +apply: emeasurable_fun_infty_o => //=. +under [X in _ _ X]eq_fun. + move=> x. + rewrite -(@fineK _ (normal_prob x s Ys)); last first. +(* rewrite -(@fineK _ (\int[_]_(_ in Ys) _)%E); last exact: integral_f_fin_num. *) + rewrite ge0_fin_numE => //. + apply: (@le_lt_trans _ _ (normal_prob x s setT)). + by apply: le_measure => //; rewrite inE/=. + apply: (@le_lt_trans _ _ 1%E). + exact: probability_le1. + exact: ltey. + over. +apply/measurable_EFinP. +apply: continuous_measurable_fun. +exact: normal_prob_continuous. +Qed. + +End normal_kernel. + +Section normal_probability. +Context {R : realType}. + +Lemma measurable_normal_s_prob (s : R) : + measurable_fun setT + (fun m : R => normal_prob m s : pprobability _ _). +Proof. +apply: (@measurability _ _ _ _ _ _ + (@pset _ _ _ : set (set (pprobability _ R)))) => //. +move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. +Admitted. +(* +have := measurable_normal_prob2 (integral_normal^~ s) mYs. +Qed. +*) + +End normal_probability. + +Section dirac_delta. +Local Open Scope ereal_scope. +Context {R: realType}. +Local Notation mu := lebesgue_measure. + +Hypothesis integral_normal : forall (m : R) (s : {posnum R}), + (\int[@lebesgue_measure R]_x (normal_pdf m s%:num x)%:E = 1%E)%E. + +Definition dirac_delta (m x : R) : \bar R := lim ((normal_pdf m s x)%:E @[s --> 0^'+]). + +Lemma integralT_dirac_delta m : + \int[mu]_x dirac_delta m x = 1. +Proof. +rewrite [LHS](_:_= lim ((\int[mu]_x (normal_pdf m s x)%:E) @[s --> 0^'+])); last first. + rewrite /dirac_delta. + (* rewrite monotone_convergence *) + admit. +apply: lim_near_cst => //. +near=> s. +have [s0 s0s]: exists s' : {posnum R}, s'%:num = s. + have s0 : (0 < s)%R by []. + by exists (PosNum s0). +rewrite -s0s. +apply: integral_normal. +Abort. + +Lemma dirac_deltaE x A : + \int[mu]_(y in A) dirac_delta x y = \d_x A. +Proof. +rewrite diracE. +Abort. + +End dirac_delta. From cbcbe7108473df3e7a98561cb730725264a50f2a Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 12 Mar 2025 15:18:59 +0900 Subject: [PATCH 34/48] one admit in helloright --- theories/lang_syntax_hello_wip.v | 28 ++++++++++++++++++---------- 1 file changed, 18 insertions(+), 10 deletions(-) diff --git a/theories/lang_syntax_hello_wip.v b/theories/lang_syntax_hello_wip.v index 3c79472ce7..180d1a8a0b 100644 --- a/theories/lang_syntax_hello_wip.v +++ b/theories/lang_syntax_hello_wip.v @@ -518,15 +518,24 @@ transitivity (\int[mu]_(z in V) by move/mem_set. apply: eq_integral. move=> x _. -transitivity ((2 * pi * (Num.sqrt 2))^-1%:E * - \int[mu]_x0 ((expR (- (x - x0) ^+ 2 / 2))%R * expR (- (x0 - (y.1 / 2)) ^+ 2 ))%:E). - under eq_integral. - move=> z _. - rewrite !normal_pdfE//. - rewrite /normal_pdf0. - - over. - admit. +transitivity ((pi * (Num.sqrt 2))^-1%:E * + \int[mu]_z (expR (- (x - z) ^+ 2 / 2) * expR (- (z - y.1 / 2) ^+ 2))%:E). + rewrite -ge0_integralZl//=; last 3 first. + - apply/measurable_EFinP => //=; apply: measurable_funM => //=. + + apply: measurableT_comp => //=; apply: measurable_funM => //=. + apply/measurableT_comp => //; apply: measurable_funX. + exact: measurable_funB. + + do 2 apply: measurableT_comp => //=. + by apply: measurable_funX; exact: measurable_funD. + - by move=> z _; rewrite lee_fin mulr_ge0// expR_ge0. + - by rewrite lee_fin// invr_ge0// mulr_ge0// pi_ge0. + apply: eq_integral => /= z _. + rewrite /normal_pdf oner_eq0 gt_eqF; last by rewrite invr_gt0. + rewrite /normal_pdf0 expr1n mul1r exprVn sqr_sqrtr//. + rewrite -mulrnAr -(mulr_natr pi) -(mulr_natr 2^-1). + rewrite (mulrCA _ pi) mulVf ?mulr1 ?divr1// mulrACA; congr ((_ * _)%:E). + rewrite sqrtrM ?pi_ge0// !invfM// mulrAC; congr (_ / _). + by rewrite -[LHS]invfM -expr2 sqr_sqrtr ?pi_ge0. rewrite normal_pdfE// /normal_pdf0. evar (C : Real.sort R). transitivity (((2 * pi * Num.sqrt 2)^-1)%:E * @@ -536,7 +545,6 @@ transitivity (((2 * pi * Num.sqrt 2)^-1)%:E * admit. Admitted. - End helloRight_subproofs. Section helloRight_proof. From 1f6f747b1aa67395c06c645b867b52c2ced3af89 Mon Sep 17 00:00:00 2001 From: IshiguroYoshihiro Date: Wed, 12 Mar 2025 15:33:59 +0900 Subject: [PATCH 35/48] wip --- theories/lang_syntax_hello_wip.v | 40 ++++++++++++++++++++++++++++---- 1 file changed, 36 insertions(+), 4 deletions(-) diff --git a/theories/lang_syntax_hello_wip.v b/theories/lang_syntax_hello_wip.v index 180d1a8a0b..da54fd10fa 100644 --- a/theories/lang_syntax_hello_wip.v +++ b/theories/lang_syntax_hello_wip.v @@ -537,11 +537,43 @@ transitivity ((pi * (Num.sqrt 2))^-1%:E * rewrite sqrtrM ?pi_ge0// !invfM// mulrAC; congr (_ / _). by rewrite -[LHS]invfM -expr2 sqr_sqrtr ?pi_ge0. rewrite normal_pdfE// /normal_pdf0. -evar (C : Real.sort R). -transitivity (((2 * pi * Num.sqrt 2)^-1)%:E * - \int[mu]_x0 (expR (- (3 / 2)%R * (x0 - C) ^+ 2 - (x - y.1 / 2) ^+ 2 / (3 / 2 *+ 2)))%:E). - admit. +transitivity (((pi * Num.sqrt 2)^-1)%:E * + \int[mu]_x0 (expR (- (3 / 2)%R * (x0 - (x + y.1) / 3) ^+ 2 - (x - y.1 / 2) ^+ 2 / (3 / 2 *+ 2)))%:E). + congr *%E. + apply: eq_integral. + move=> z _. + rewrite -expRD. + congr EFin. + congr expR. + rewrite !sqrrD. + lra. (* gauss integral *) +under eq_integral do rewrite expRD EFinM. +rewrite ge0_integralZr//; first last. + rewrite lee_fin. + exact: expR_ge0. + move=> z _. + rewrite lee_fin. + exact: expR_ge0. + apply/measurable_EFinP. + apply: measurableT_comp => //. + apply: measurable_funM => //. + apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)%R) => //. + exact: measurable_funD. +rewrite [in RHS]sqr_sqrtr//. +rewrite -[in RHS]mulr_natr [in RHS]mulrAC divfK//. +rewrite mulNr EFinM muleA; congr *%E. +rewrite (_: \int[mu]_x0 (expR (- (3 / 2) * (x0 - (x + y.1)%E / 3) ^+ 2))%:E + = (Num.sqrt ((3 / 2) ^+ 2 * pi *+ 2))%:E * + \int[mu]_x0 (normal_pdf (x + y.1)%R (Num.sqrt (3 / 2)) x0)%:E); last first. + admit. +rewrite integral_normal_pdf mule1. +rewrite -EFinM; congr EFin. +rewrite -(mulr_natr (_ * pi)). +rewrite -mulrA. +rewrite sqrtrM//. +rewrite -mulrA. +xxx admit. Admitted. From 56c15764946c9531bdb3531a3573602aa04f2906 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 12 Mar 2025 16:53:05 +0900 Subject: [PATCH 36/48] minor progress - one admit - cleaning --- theories/lang_syntax.v | 2 +- theories/lang_syntax_examples.v | 48 +++ theories/lang_syntax_hello_wip.v | 31 +- theories/prob_lang_wip.v | 2 + theories/probability.v | 602 +++++++++++-------------------- 5 files changed, 271 insertions(+), 414 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index c8c9aad7d8..67da9417a3 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -954,7 +954,7 @@ Inductive evalD : forall g t, exp D g t -> e -D> r ; mr -> (exp_normal s0 e : exp D g _) -D> (fun x => @normal_prob _ (r x) s) ; - measurableT_comp (measurable_normal_prob2 s0 ) mr + measurableT_comp (measurable_normal_prob2 s0) mr | eval_normalize g t (e : exp P g t) k : e -P> k -> diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index 232f32c172..cec6569e52 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -404,6 +404,53 @@ Qed. End sample_binomial. +Section nondeterminism_and_weights. +Context {R : realType}. +Open Scope lang_scope. +Open Scope ring_scope. + +Definition binomial2p (p : R) : @exp R _ [::] _ := + [let "x" := Sample {exp_binomial 2 [{p}:R]} in return #{"x"}]. + +Definition return2 (p : R) : @exp R _ [::] _ := + [let "_" := Score {p ^+ 2}:R in return {2}:N]. + +Definition return1 (p : R) : @exp R _ [::] _ := + [let "_" := Score {p * `1-p *+ 2}:R in return {1}:N]. + +Definition return0 (p : R) : @exp R _ [::] _ := + [let "_" := Score {`1-p^+2}:R in return {0}:N]. + +Lemma exec_binomial2p (p : R) t U : 0 <= p <= 1 -> measurable U -> + execP (binomial2p p) t U = + execP (return2 p) t U + + execP (return1 p) t U + + execP (return0 p) t U. +Proof. +move=> /= /andP[p0 p1] mU. +(* simplify the lhs *) +rewrite [in LHS]execP_letin execP_sample/= execD_binomial/=. +rewrite execP_return/= !execD_real/= exp_var'E (execD_var_erefl "x")/=. +rewrite letin'E/= integral_binomial//=. +rewrite !big_ord_recr big_ord0//=. +rewrite !(bin0,bin1,bin2). +rewrite !(add0r,expr0,mul1r,mulr1,subn0,mulr1n,expr1). +(* simplify the rhs *) +rewrite /return2 /return1 /return0. +rewrite ![in RHS]execP_letin !execP_score/= !execD_real/=. +rewrite !execP_return/= !execD_nat/=. +rewrite !letin'E/=. +rewrite !ge0_integral_mscale//=. +rewrite !integral_dirac//=. +rewrite !diracT. +rewrite !(mul1e) ger0_norm ?sqr_ge0//. +rewrite ger0_norm ?mulrn_wge0 ?mulr_ge0 ?onem_ge0//. +rewrite ger0_norm ?sqr_ge0//. +by rewrite -addeA addeC -addeA addeCA addeA. +Qed. + +End nondeterminism_and_weights. + Section hard_constraint. Local Open Scope ring_scope. Local Open Scope lang_scope. @@ -976,3 +1023,4 @@ Example letinC_ground (g := [:: ("a", Unit); ("b", Bool)]) t1 t2 Proof. move=> U mU; exact: letinC. Qed. End letinC. + diff --git a/theories/lang_syntax_hello_wip.v b/theories/lang_syntax_hello_wip.v index da54fd10fa..d2a12b6db1 100644 --- a/theories/lang_syntax_hello_wip.v +++ b/theories/lang_syntax_hello_wip.v @@ -563,19 +563,24 @@ rewrite ge0_integralZr//; first last. rewrite [in RHS]sqr_sqrtr//. rewrite -[in RHS]mulr_natr [in RHS]mulrAC divfK//. rewrite mulNr EFinM muleA; congr *%E. -rewrite (_: \int[mu]_x0 (expR (- (3 / 2) * (x0 - (x + y.1)%E / 3) ^+ 2))%:E - = (Num.sqrt ((3 / 2) ^+ 2 * pi *+ 2))%:E * - \int[mu]_x0 (normal_pdf (x + y.1)%R (Num.sqrt (3 / 2)) x0)%:E); last first. - admit. -rewrite integral_normal_pdf mule1. -rewrite -EFinM; congr EFin. -rewrite -(mulr_natr (_ * pi)). -rewrite -mulrA. -rewrite sqrtrM//. -rewrite -mulrA. -xxx -admit. -Admitted. +rewrite [X in _ * X = _](_ : _ = (Num.sqrt ((1 / 3) * pi *+ 2))%:E * + \int[mu]_z (normal_pdf ((x + y.1) / 3) (Num.sqrt (1 / 3)) z)%:E); last first. + rewrite -ge0_integralZl//=; last 2 first. + by apply/measurable_EFinP; exact: measurable_normal_pdf. + by move=> /= z _; rewrite lee_fin normal_pdf_ge0. + apply: eq_integral => /= z _. + rewrite /normal_pdf gt_eqF// /normal_pdf0 sqr_sqrtr// -EFinM; congr EFin. + rewrite [RHS]mulrA -[LHS]mul1r; congr (_ * expR _)%R. + by rewrite divff// gt_eqF// sqrtr_gt0 pmulrn_rgt0// mulr_gt0// pi_gt0. + rewrite -(mulr_natl (1 / 3)) mul1r. + by rewrite !mulNr (mulrC (3 / 2)) invfM invrK (mulrC (2^-1)). +rewrite integral_normal_pdf mule1 -EFinM mul1r; congr EFin. +rewrite -(mulr_natr (_ * pi)) sqrtrM ?mulr_ge0 ?pi_ge0//. +rewrite invfM mulrCA -mulrA mulVf ?mulr1 ?gt_eqF//. +rewrite !sqrtrM// invfM sqrtrV// -mulrA; congr *%R. +rewrite -[X in _ / X]sqr_sqrtr ?pi_ge0//. +by rewrite -exprVn expr2 mulrCA divff ?mulr1// sqrtr_eq0 leNgt pi_gt0. +Qed. End helloRight_subproofs. diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 2b422b33cb..0bb9b0a6a0 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -11,6 +11,8 @@ From mathcomp Require Import gauss_integral. (**md**************************************************************************) (* wip waiting for the Poisson distribution *) +(* *) +(* Another example from Section 4.2 in [Equation (13), Staton, ESOP 2017]. *) (******************************************************************************) Set Implicit Arguments. diff --git a/theories/probability.v b/theories/probability.v index 49805ed32f..2b7027287a 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -3010,6 +3010,7 @@ Abort. End near_ereal_nondecreasing_is_cvgn. +(* TODO: move as another PR *) Section near_monotone_convergence. Local Open Scope ereal_scope. @@ -3210,36 +3211,46 @@ Section normal_density. Context {R : realType}. Local Open Scope ring_scope. Local Import Num.ExtraDef. +Implicit Types m s x : R. -(* NB: If s = 0 then normal_pdf0 is cst 0 for all m and x - * because divr is defined by x / 0 = 0. - *) -Definition normal_pdf0 (m s x : R) : R := - (sqrtr (s ^+2 * pi *+ 2))^-1 * expR (- (x - m) ^+ 2 / (s ^+ 2*+ 2)). +Definition normalexp m s x := expR (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)). -Lemma normal_pdf0_center m s : normal_pdf0 m s = normal_pdf0 0 s \o center m. +Lemma measurable_normalexp m s : measurable_fun setT (normalexp m s). Proof. -apply/funext => x/=. -by rewrite /normal_pdf0 subr0. +apply: measurableT_comp => //=; apply: measurable_funM => //=. +apply: measurableT_comp => //=; apply: measurable_funX => //=. +exact: measurable_funB. Qed. -Lemma normal_pdf0_ge0 m s x : 0 <= normal_pdf0 m s x. -Proof. by rewrite mulr_ge0 ?expR_ge0// invr_ge0 mulr_ge0. Qed. +Lemma normalexp_center m s : normalexp m s = normalexp 0 s \o center m. +Proof. by apply/funext => x/=; rewrite /normalexp/= subr0. Qed. -Lemma normal_pdf0_gt0 m s x : s != 0 -> 0 < normal_pdf0 m s x. +Definition normalpi s := (sqrtr (s ^+ 2 * pi *+ 2))^-1. + +Lemma normalpi_ge0 s : 0 <= normalpi s. Proof. by rewrite invr_ge0. Qed. + +Lemma normalpi_gt0 s : s != 0 -> 0 < normalpi s. Proof. -move=> s0. -rewrite mulr_gt0 ?expR_gt0// invr_gt0//. -by rewrite sqrtr_gt0 pmulrn_rgt0// pmulr_rgt0// ?pi_gt0// exprn_even_gt0/=. +move=> s0; rewrite invr_gt0// sqrtr_gt0. +by rewrite pmulrn_rgt0// mulr_gt0 ?pi_gt0// exprn_even_gt0/=. Qed. +(* NB: If s = 0 then normal_pdf0 is cst 0 for all m and x + * because divr is defined by x / 0 = 0. + *) +Definition normal_pdf0 m s x : R := normalpi s * normalexp m s x. + Lemma measurable_normal_pdf0 m s : measurable_fun setT (normal_pdf0 m s). -Proof. -apply: measurable_funM => //=; apply: measurableT_comp => //=. -apply: measurable_funM => //=; apply: measurableT_comp => //=. -apply: measurableT_comp (exprn_measurable _) _ => /=. -exact: measurable_funD. -Qed. +Proof. by apply: measurable_funM => //=; exact: measurable_normalexp. Qed. + +Lemma normal_pdf0_center m s : normal_pdf0 m s = normal_pdf0 0 s \o center m. +Proof. by rewrite /normal_pdf0 normalexp_center. Qed. + +Lemma normal_pdf0_ge0 m s x : 0 <= normal_pdf0 m s x. +Proof. by rewrite mulr_ge0 ?normalpi_ge0 ?expR_ge0. Qed. + +Lemma normal_pdf0_gt0 m s x : s != 0 -> 0 < normal_pdf0 m s x. +Proof. by move=> s0; rewrite mulr_gt0 ?expR_gt0// normalpi_gt0. Qed. Lemma continuous_normal_pdf0 m s : continuous (normal_pdf0 m s). Proof. @@ -3249,16 +3260,14 @@ apply: (@cvg_comp _ R^o _ _ _ _ (nbhs (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)))); last exact: continuous_expR. apply: cvgM; last exact: cvg_cst. apply: (@cvgN _ _ _ _ _ _ ((x - m) ^+ 2 : R^o)). -apply: (@cvg_comp _ _ _ _ (fun x : R => x ^+ 2)%R _ (nbhs (x - m))). - apply: (@cvgB _ R^o) => //. - exact: cvg_cst. +apply: (@cvg_comp _ _ _ _ (fun x => x ^+ 2)%R _ (nbhs (x - m))). + by apply: (@cvgB _ R^o) => //; exact: cvg_cst. exact: sqr_continuous. Qed. -Lemma normal_pdf0_ub m s x : normal_pdf0 m s x <= (sqrtr (s ^+ 2 * pi *+ 2))^-1. +Lemma normal_pdf0_ub m s x : normal_pdf0 m s x <= normalpi s. Proof. -rewrite -[leRHS]mulr1. -rewrite ler_wpM2l ?invr_ge0// ?sqrtr_ge0. +rewrite /normal_pdf0 ler_piMr ?normalpi_ge0//. rewrite -[leRHS]expR0 ler_expR mulNr oppr_le0 mulr_ge0// ?sqr_ge0//. by rewrite invr_ge0 mulrn_wge0// sqr_ge0. Qed. @@ -3270,20 +3279,8 @@ Variables (R : realType) (m sigma : R). Local Open Scope ring_scope. Notation mu := lebesgue_measure. -Definition normal_pdf_part x := (expR (- (x - m) ^+ 2 / (sigma ^+ 2 *+ 2))). - -Let measurable_normal_pdf_part : measurable_fun setT normal_pdf_part. -Proof. -apply: measurableT_comp => //=. -apply: measurable_funM => //=. -apply: measurableT_comp => //=. -apply: measurable_funX => //=. -exact: measurable_funD. -Qed. - -Let integral_normal_pdf_part : sigma != 0 -> - (\int[mu]_x (normal_pdf_part x)%:E = - (Num.sqrt (sigma ^+ 2 * pi *+ 2))%:E)%E. +Lemma integral_normalexp : sigma != 0 -> + (\int[mu]_x (normalexp m sigma x)%:E = (normalpi sigma)^-1%:E)%E. Proof. move=> sigma_gt0. pose F (x : R^o) := (x - m) / (`|sigma| * Num.sqrt 2). @@ -3291,7 +3288,8 @@ have F'E : F^`()%classic = cst (`|sigma| * Num.sqrt 2)^-1. apply/funext => x; rewrite /F derive1E deriveM// deriveD// derive_cst scaler0. by rewrite add0r derive_id derive_cst addr0 scaler1. have := @integralT_gauss R. -rewrite (@increasing_ge0_integration_by_substitutionT _ F gauss_fun)//=; first last. +rewrite (@increasing_ge0_integration_by_substitutionT _ F gauss_fun)//=; + first last. - by move=> x; rewrite gauss_fun_ge0. - exact: continuous_gauss_fun. - apply/gt0_cvgMly; last exact: cvg_addrr. @@ -3301,22 +3299,22 @@ rewrite (@increasing_ge0_integration_by_substitutionT _ F gauss_fun)//=; first l - by rewrite F'E; exact: is_cvg_cst. - by rewrite F'E; exact: is_cvg_cst. - by rewrite F'E => ?; exact: cvg_cst. -- by move=> x y; rewrite /F ltr_pM2r ?invr_gt0 ?mulr_gt0 ?normr_gt0// ltrBlDr subrK. +- move=> x y. + by rewrite /F ltr_pM2r ?invr_gt0 ?mulr_gt0 ?normr_gt0// ltrBlDr subrK. move=> /(congr1 (fun x => x * (`|sigma| * Num.sqrt 2)%:E)%E). -rewrite -ge0_integralZr//=; last first. - by move=> x _; rewrite lee_fin mulr_ge0//= ?gauss_fun_ge0// F'E/= invr_ge0. +rewrite -ge0_integralZr//=; last 2 first. rewrite F'E; apply/measurable_EFinP/measurable_funM => //. apply/measurableT_comp => //; first exact: measurable_gauss_fun. by apply: measurable_funM => //; exact: measurable_funD. -move=> int_gauss. -apply: eq_trans. + by move=> x _; rewrite lee_fin mulr_ge0//= ?gauss_fun_ge0// F'E/= invr_ge0. +move=> int_gauss; apply: eq_trans. apply: eq_trans; last exact: int_gauss. apply: eq_integral => /= x _. rewrite F'E !fctE/= EFinM -muleA -EFinM mulVf ?mulr1; last first. by rewrite gt_eqF ?mulr_gt0 ?normr_gt0. rewrite /gauss_fun /F exprMn exprVn -[in RHS]mulNr. by rewrite exprMn real_normK ?num_real// sqr_sqrtr// mulr_natr. -rewrite -mulrnAl sqrtrM ?mulrn_wge0 ?sqr_ge0//. +rewrite /normalpi invrK -mulrnAl sqrtrM ?mulrn_wge0 ?sqr_ge0//. by rewrite -[in RHS]mulr_natr sqrtrM ?sqr_ge0// sqrtr_sqr !EFinM muleC. Qed. @@ -3370,12 +3368,14 @@ Local Open Scope ring_scope. Notation mu := lebesgue_measure. +(* PR in progress *) Lemma derive_shift (v k : R) : 'D_v (shift k : R^o -> R^o) = cst v. Proof. apply/funext => x/=. by rewrite deriveD// derive_id derive_cst addr0. Qed. +(* PR in progress *) Lemma is_derive_shift x v (k : R) : is_derive x v (shift k : R^o -> R^o) v. Proof. split => //. @@ -3442,7 +3442,6 @@ Qed. End shift_properties. Section normal_kernel. - Variable R : realType. Variables s : R. Hypothesis s0 : s != 0. @@ -3549,7 +3548,9 @@ apply: eq_integral. move=> x _. Admitted. *) - +(*Local Definition normal_prob2 := + (fun m => normal_prob m s) : _ -> pprobability _ _. +*) (* Lemma normal_shift0 x : normal_prob2 x = @@ -3769,14 +3770,11 @@ Admitted. added by the integral of `normal_pdf a s x` on ]a - e, a + e[ *) -Let g' a e : R -> R := fun x => if (x \in (ball a e : set R^o)) then - (Num.sqrt (s ^+ 2 * pi *+ 2))^-1 * - expR (- 0%R ^+ 2 / (s ^+ 2 *+ 2)) -else - (Num.sqrt (s ^+ 2 * pi *+ 2))^-1 * - expR (- (`|x - a| - e) ^+ 2 / (s ^+ 2 *+ 2)). +Let g' a e : R -> R := fun x => if x \in (ball a e : set R^o) then + normalpi s else normal_pdf0 e s `|x - a|. -Let ballFE_le (a e x : R) : x <= (a - e)%R -> (x \in (ball a e : set R^o)) = false. +Let ballFE_le (a e x : R) : x <= (a - e)%R -> + (x \in (ball a e : set R^o)) = false. Proof. move=> xae. apply: memNset. @@ -3784,7 +3782,8 @@ rewrite ball_itv/= in_itv/=; apply/negP/andP/not_andP; left. by apply/negP; rewrite -leNgt. Qed. -Let ballFE_ge (a e x : R) : a + e <= x -> (x \in (ball a e : set R^o)) = false. +Let ballFE_ge (a e x : R) : a + e <= x -> + (x \in (ball a e : set R^o)) = false. Proof. move=> xae. apply: memNset. @@ -3794,13 +3793,12 @@ Qed. Lemma g'a0 (a : R) : g' a 0 = normal_pdf0 a s. Proof. -rewrite /g'. -apply/funext => x. +apply/funext => x; rewrite /g'. have /orP [x0|x0] := le_total x a. rewrite ballFE_le; last by rewrite subr0. - by rewrite subr0 real_normK// num_real. + by rewrite /normal_pdf0 /normalexp subr0 real_normK// num_real. rewrite ballFE_ge; last by rewrite addr0. -by rewrite subr0 real_normK// num_real. +by rewrite /normal_pdf0 /normalexp subr0 real_normK// num_real. Qed. Lemma mg' a e : measurable_fun setT (g' a e). @@ -3809,163 +3807,137 @@ apply: measurable_fun_if => //. apply: (measurable_fun_bool true) => /=. rewrite setTI preimage_mem_true. exact: measurable_ball. -apply: measurable_funTS => /=. -apply: measurableT_comp => //. -apply: measurableT_comp => //. -apply: measurable_funM => //. -apply: measurableT_comp => //. -apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)) => //. -apply: measurable_funB => //. -apply: measurableT_comp => //. -exact: measurable_funD. +apply: measurable_funTS => /=; apply: measurable_funM => //. +apply: measurableT_comp => //; first exact: measurable_normalexp. +by apply: measurableT_comp => //; exact: measurable_funD. Qed. Lemma g'_ge0 a e x : 0 <= g' a e x. Proof. -rewrite /g'; case: ifP => _. - rewrite -[in leRHS](@subr0 _ 0) -normal_pdfE ?subr0; apply: normal_pdf_ge0. -rewrite -normal_pdfE; exact: normal_pdf_ge0. +rewrite /g'; case: ifP => _; first by rewrite normalpi_ge0. +exact: normal_pdf0_ge0. Qed. -Lemma continuous_g' (a e : R) : (0 <= e) -> continuous (g' a e). +Lemma continuous_g' (a e : R) : 0 <= e -> continuous (g' a e). Proof. move=> e0. -have tmp k : k < a - e -> ((`|k - a| - e) ^+ 2 = (k - (a - e)) ^+ 2). - move=> kae. - rewrite -normrN opprB. - have /normr_idP -> : (0 <= a - k). - by rewrite subr_ge0 ltW// (lt_le_trans kae)// gerBl. - by rewrite -sqrrN !opprB addrCA. -have tmp2 k : a + e < k -> ((`|k - a| - e) ^+ 2 = (k - (a + e)) ^+ 2). +have aNe k : k < a - e -> (`|k - a| - e) ^+ 2 = (k - (a - e)) ^+ 2. + move=> kae; rewrite ler0_norm; first by rewrite -sqrrN !opprB addrCA. + by rewrite subr_le0 (le_trans (ltW kae))// lerBlDl lerDr. +have aDe k : a + e < k -> (`|k - a| - e) ^+ 2 = (k - (a + e)) ^+ 2. move=> kae; rewrite opprD addrA. - have /normr_idP -> // : (0 <= k - a). - by rewrite subr_ge0 ltW// (le_lt_trans _ kae)// lerDl. + by rewrite ger0_norm// subr_ge0 (le_trans _ (ltW kae))// lerDl. apply: (@in1TT R). rewrite -continuous_open_subspace; last exact: openT. -rewrite (_:[set: R] = - `]-oo, (a - e)%R] `|` `[(a - e)%R, a + e] `|` `[a + e, +oo[); last first. +rewrite (_ : [set: R] = + `]-oo, (a - e)%R] `|` `[(a - e)%R, a + e] `|` `[a + e, +oo[); last first. rewrite -setUitv1// -setUA setUAC setUA -itv_bndbnd_setU//; last first. by rewrite bnd_simp lerD// ge0_cp. rewrite -setUitv1// -setUA setUAC setUA -itv_bndbnd_setU//. by rewrite set_itvE !setTU. apply: withinU_continuous. - rewrite -setUitv1//. - rewrite -setUA setUCA. - rewrite -itv_bndbnd_setU//; last first. - by rewrite bnd_simp lerD// ge0_cp. - rewrite setUidr//. - move=> _/= ->; rewrite in_itv/=. - by rewrite lerD// ge0_cp. - by apply: interval_closed. - apply: withinU_continuous; first exact: interval_closed. - exact: interval_closed. - apply/continuous_within_itvNycP; split. - move=> x. +- rewrite -setUitv1// -setUA setUCA -itv_bndbnd_setU//; last first. + by rewrite bnd_simp lerD// ge0_cp. + by rewrite setUidr// sub1set inE/= in_itv/= lerD// ge0_cp. +- exact: interval_closed. +- apply: withinU_continuous; first exact: interval_closed. + + exact: interval_closed. + + apply/continuous_within_itvNycP; split. + * move=> x. rewrite in_itv/= => xae. - apply/(@cvgrPdist_le _ R^o R _ _ (g' a e) (g' a e x)). - rewrite /=. - move=> eps eps0. + apply/(@cvgrPdist_le _ R^o R _ _ (g' a e) (g' a e x)) => /= eps eps0. near=> t. - have tae : t < a - e. - near: t. - exact: lt_nbhsl. + have tae : t < a - e by near: t; exact: lt_nbhsl. rewrite /g'. rewrite !ballFE_le ?(@ltW _ _ _ (a - e))//. - rewrite !tmp//. + rewrite /normal_pdf0 /normalexp !aNe//. + rewrite -!/(normalexp _ _ _) -!/(normal_pdf0 _ _ _). move=> {tae}; near: t. move: eps eps0. - apply/(@cvgrPdist_le _ _ _ (nbhs x)). + apply/(@cvgrPdist_le R^o). exact: continuous_normal_pdf0. - apply/(@cvgrPdist_lt _ R^o). - move=> eps eps0. - near=> t. - rewrite /g' !ballFE_le//. - rewrite -addrAC subrr sub0r normrN; have /normr_idP -> := e0. - rewrite subrr -(subrr (a - e)) tmp//. - near: t; move: eps eps0. - apply/(@cvgrPdist_lt _ R^o). - apply: cvg_at_left_filter. - exact: continuous_normal_pdf0. - move: e0; rewrite le_eqVlt => /predU1P; case => [<- | e0]. + * apply/(@cvgrPdist_lt _ R^o) => eps eps0. + near=> t. + rewrite /g' !ballFE_le//. + rewrite -addrAC subrr sub0r normrN (ger0_norm e0)//. + rewrite /normal_pdf0 /normalexp subrr -(subrr (a - e)) aNe//. + near: t; move: eps eps0. + apply/(@cvgrPdist_lt _ R^o). + apply: cvg_at_left_filter. + exact: continuous_normal_pdf0. + move: e0; rewrite le_eqVlt => /predU1P[<-|e0]. rewrite g'a0. apply: continuous_subspaceT. exact: continuous_normal_pdf0. apply/continuous_within_itvP; first by rewrite -(opprK e) ler_ltB// opprK gtrN. split. - move=> x xae. - rewrite /continuous_at. - rewrite /g' ifT; last by rewrite ball_itv inE/=. - apply/(@cvgrPdist_le _ R^o). - move=> eps eps0. - near=> t. - rewrite ifT; first by rewrite subrr normr0 ltW. - rewrite ball_itv inE/= in_itv/=; apply/andP; split. - near: t. - apply: lt_nbhsr. - by move: xae; rewrite in_itv/= => /andP[]. + + move=> x xae. + rewrite /continuous_at. + rewrite /g' ifT; last by rewrite ball_itv inE/=. + apply/(@cvgrPdist_le _ R^o) => eps eps0. + near=> t. + rewrite ifT; first by rewrite subrr normr0 ltW. + rewrite ball_itv inE/= in_itv/=; apply/andP; split. near: t. - apply: lt_nbhsl. + apply: lt_nbhsr. by move: xae; rewrite in_itv/= => /andP[]. - apply/(@cvgrPdist_le _ R^o). - move=> eps eps0. + near: t. + apply: lt_nbhsl. + by move: xae; rewrite in_itv/= => /andP[]. + + apply/(@cvgrPdist_le _ R^o) => eps eps0. near=> t. rewrite /g' ballFE_le// ifT; last first. rewrite ball_itv inE/= in_itv/=; apply/andP => []; split => //. near: t. apply: nbhs_right_lt. by rewrite -(opprK e) ler_ltB// opprK gtrN. - rewrite addrAC subrr sub0r normrN; have /ltW/normr_idP -> := e0. - by rewrite !subrr normr0 ltW. - apply/(@cvgrPdist_le _ R^o). - move=> eps eps0. - near=> t. - rewrite /g' ballFE_ge// ifT; last first. - rewrite ball_itv inE/= in_itv/=; apply/andP => []; split => //. - near: t. - apply: nbhs_left_gt. - by rewrite -(opprK e) ler_ltB// opprK gtrN. - rewrite addrAC subrr add0r; have /ltW/normr_idP -> := e0. - by rewrite !subrr normr0 ltW. -apply/continuous_within_itvcyP; split. - move=> x. - rewrite in_itv/= andbT => aex. - apply/(@cvgrPdist_le _ R^o). - rewrite /=. - move=> eps eps0. - near=> t. - have tae : a + e < t. + rewrite /normal_pdf0. + rewrite addrAC subrr sub0r normrN /normalexp (gtr0_norm e0) subrr. + by rewrite expr0n/= oppr0 mul0r expR0 mulr1 subrr normr0 ltW. + + apply/(@cvgrPdist_le _ R^o) => eps eps0. + near=> t. + rewrite /g' ballFE_ge// ifT; last first. + rewrite ball_itv inE/= in_itv/=; apply/andP => []; split => //. + near: t. + apply: nbhs_left_gt. + by rewrite -(opprK e) ler_ltB// opprK gtrN. + rewrite /normal_pdf0. + rewrite addrAC subrr add0r /normalexp (gtr0_norm e0) subrr expr0n oppr0. + by rewrite mul0r expR0 mulr1 subrr normr0 ltW. +- apply/continuous_within_itvcyP; split. + + move=> x. + rewrite in_itv/= andbT => aex. + apply/(@cvgrPdist_le _ R^o) => /= eps eps0. + near=> t. + have tae : a + e < t by near: t; exact: lt_nbhsr. + rewrite /g'. + rewrite !ballFE_ge ?(@ltW _ _ (a + e)%E)//. + rewrite /normal_pdf0 /normalexp !aDe// ?(@ltW _ _ (a + e)). + move=> {tae}; near: t. + move: eps eps0. + apply/(@cvgrPdist_le _ R^o). + exact: continuous_normal_pdf0. + + apply/(@cvgrPdist_le _ R^o) => eps eps0. + near=> t. + rewrite /g' !ballFE_ge//. + rewrite addrAC subrr add0r (ger0_norm e0)//. + rewrite /normal_pdf0 /normalexp subrr -(subrr (a + e)). + rewrite aDe//. near: t. - exact: lt_nbhsr. - rewrite /g'. - rewrite !ballFE_ge ?(@ltW _ _ (a + e)%E)//. - rewrite !tmp2// ?(@ltW _ _ (a + e)). - move=> {tae}; near: t. - move: eps eps0. - apply/(@cvgrPdist_le _ _ _ (nbhs x)). - exact: continuous_normal_pdf0. -apply/(@cvgrPdist_le _ R^o). -move=> eps eps0. -near=> t. -rewrite /g' !ballFE_ge//. -rewrite addrAC subrr add0r; have /normr_idP -> := e0. -rewrite subrr -(subrr (a + e)). -rewrite tmp2//. -near: t. -move: eps eps0. -apply/cvgrPdist_le. -apply: cvg_at_right_filter. -apply: continuous_normal_pdf0. + move: eps eps0. + apply/cvgrPdist_le. + apply: cvg_at_right_filter. + exact: continuous_normal_pdf0. Unshelve. all: end_near. Qed. -Lemma gE_Ny a e : (0 <= e) -> +Lemma gE_Ny a e : 0 <= e -> (\int[mu]_(x in `]-oo, (a - e)%R]) `|g' a e x|%:E = - \int[mu]_(x in `]-oo, a]) `|normal_pdf a s x|%:E)%E. + \int[mu]_(x in `]-oo, a]) `|normal_pdf a s x|%:E)%E. Proof. move=> e0. rewrite ge0_integration_by_substitution_shift_itvNy => /=; first last. -- move=> ? _; exact: normr_ge0. -- apply/continuous_subspaceT. - move=> x. +- by move=> ? _; exact: normr_ge0. +- apply/continuous_subspaceT => x. apply: continuous_comp; first exact: continuous_g'. exact: (@norm_continuous _ R^o) . under eq_integral. @@ -3974,22 +3946,21 @@ under eq_integral. rewrite /g' ballFE_le//; last exact: lerB. rewrite -(normrN (x - e - a)) !opprB addrA. have /normr_idP -> : 0 <= a + e - x by rewrite subr_ge0 ler_wpDr. + rewrite /normal_pdf0 /normalexp. rewrite -(addrAC _ (- x)) addrK. rewrite -(sqrrN (a - x)) opprB. over. -by under [RHS]eq_integral do rewrite normal_pdfE. +by apply: eq_integral => /= x xay; rewrite /normal_pdf (negbTE s0). Qed. -Lemma gE_y a e : (0 <= e) -> +Lemma gE_y a e : 0 <= e -> (\int[mu]_(x in `[a + e, +oo[) `|g' a e x|%:E = - \int[mu]_(x in `[a, +oo[) `|normal_pdf a s x|%:E)%E. + \int[mu]_(x in `[a, +oo[) `|normal_pdf a s x|%:E)%E. Proof. move=> e0. rewrite ge0_integration_by_substitution_shift_itvy => /=; first last. -- move=> ? _. - exact: normr_ge0. -- apply/continuous_subspaceT. - move=> x. +- by move=> ? _; exact: normr_ge0. +- apply/continuous_subspaceT => x. apply: continuous_comp; first exact: continuous_g'. exact: (@norm_continuous _ R^o). under eq_integral. @@ -3997,13 +3968,14 @@ under eq_integral. rewrite inE/= in_itv/= andbT => aex. rewrite /g' ballFE_ge//; last exact: lerD. have /normr_idP -> : 0 <= x + e - a by rewrite subr_ge0 ler_wpDr. - rewrite -(addrAC _ (- a)) addrK. + rewrite /normal_pdf0 /normalexp -(addrAC _ (- a)) addrK. over. -by under [RHS]eq_integral do rewrite normal_pdfE. +rewrite /=. +by apply: eq_integral => /= x xay; rewrite /normal_pdf (negbTE s0). Qed. Lemma normal_prob_continuous (V : set R) : measurable V -> - continuous (fun m => fine (normal_prob m s V)). + continuous (fun m => fine (normal_prob m s V)). Proof. move=> mV a. near (0 : R)^'+ => e. @@ -4017,241 +3989,75 @@ apply: (@continuity_under_integral _ _ _ mu _ _ _ _ (a - e) (a + e) _ _ _ g) => exact: integrable_normal_pdf. - apply/aeW => y _. move=> x. - under [X in _ _ X]eq_fun do rewrite normal_pdfE -(sqrrN (y - _)) opprB. + under [X in _ _ X]eq_fun. + move=> x0. + rewrite normal_pdfE// /normal_pdf0 /normal_fun -(sqrrN (y - _)) opprB. + over. exact: continuous_normal_pdf0. - apply: (integrableS measurableT) => //=. - apply/integrableP; split. - exact/measurable_EFinP. - rewrite -(setUv (ball a e)). - rewrite ge0_integral_setU//=; first last. - exact/disj_setPCl. - rewrite setUv. - apply/measurable_EFinP. - exact: measurableT_comp. - apply: measurableC. - exact: measurable_ball. + apply/integrableP; split; first exact/measurable_EFinP. + rewrite -(setUv (ball a e)) ge0_integral_setU//=; last 4 first. exact: measurable_ball. + by apply: measurableC; exact: measurable_ball. + rewrite setUv. + by apply/measurable_EFinP; exact: measurableT_comp. + exact/disj_setPCl. apply: lte_add_pinfty. under eq_integral. move=> x xae. - rewrite /g/g' xae. - rewrite expr0n/= oppr0 mul0r expR0 mulr1. + rewrite /g /g' xae. over. rewrite integral_cst/=. apply: lte_mul_pinfty => //. rewrite ball_itv lebesgue_measure_itv/= ifT -?EFinD ?ltry// lte_fin. by rewrite ltrBlDr -addrA -ltrBlDl subrr -mulr2n mulrn_wgt0. exact: measurable_ball. - have -> : (\int[mu]_(x in ~` ball a e) `|g x|%:E = \int[mu]_x `|normal_pdf a s x|%:E)%E. + rewrite [ltLHS](_ : _ = \int[mu]_x `|normal_pdf a s x|%:E)%E; last first. rewrite ball_itv setCitv ge0_integral_setU//=; first last. apply/disj_setPRL. rewrite setCitvl. apply: subset_itvr; rewrite bnd_simp. - rewrite -{2}(opprK e); apply: ler_ltB => //. - exact: gtrN. - apply: measurable_funTS. - apply/measurable_EFinP. + by rewrite -{2}(opprK e) ler_ltB// gtrN. + apply: measurable_funTS; apply/measurable_EFinP. exact: measurableT_comp. - rewrite gE_Ny// gE_y//. - rewrite -integral_itv_obnd_cbnd; last first. - apply: measurableT_comp => //. - apply: measurable_funTS. + rewrite gE_Ny// gE_y// -integral_itv_obnd_cbnd; last first. + apply: measurableT_comp => //; apply: measurable_funTS. exact: measurable_normal_pdf. rewrite -ge0_integral_setU/= ?measurable_itv//; first last. - apply/disj_setPRL. - by rewrite setCitvl. + by apply/disj_setPRL; rewrite setCitvl. rewrite -setCitvl setUv. - apply/measurable_EFinP. - apply: measurableT_comp => //. + apply/measurable_EFinP; apply: measurableT_comp => //. exact: measurable_normal_pdf. by rewrite -setCitvl setUv. - under eq_integral => x do rewrite (_:_%:E = `|(normal_pdf a s x)%:E|%E)//. - apply/abse_integralP; rewrite //=. - apply/measurable_EFinP. - exact: measurable_normal_pdf. + under eq_integral do rewrite -abse_EFin. + apply/abse_integralP => //=. + by apply/measurable_EFinP; exact: measurable_normal_pdf. by rewrite integral_normal_pdf ltry. move=> x ax. apply/aeW => y Vy. -have /normr_idP -> : 0 <= normal_pdf x s y by apply: normal_pdf_ge0. -rewrite normal_pdfE /g/g'. -case: ifP. - move=> _. - rewrite ler_pM//. - rewrite expr0n/= oppr0 mul0r ler_expR. - apply: mulr_le0_ge0. - rewrite oppr_le0. - exact: sqr_ge0. - by rewrite invr_ge0 mulrn_wge0// sqr_ge0. -move/negP/negP; rewrite notin_setE/= ball_itv/= in_itv/= => Hy. -rewrite ler_pM//. -rewrite ler_expR. -rewrite !mulNr lerN2. -rewrite ler_pM => //. - exact: sqr_ge0. +rewrite ger0_norm; last exact: normal_pdf_ge0. +rewrite normal_pdfE// /g /g'. +case: ifPn => [_|]; first exact: normal_pdf0_ub. +rewrite notin_setE/= ball_itv/= in_itv/= => aey. +rewrite /normal_pdf0 ler_pM//. +rewrite ler_expR !mulNr lerN2 ler_pM //. + exact: sqr_ge0. by rewrite invr_ge0 mulrn_wge0// sqr_ge0. -move: Hy; move/negP/nandP; rewrite -!leNgt; case => [yae|aey]. - rewrite -normrN opprB. - have /normr_idP -> : 0 <= a - y. - rewrite subr_ge0. - apply: (le_trans yae). - by rewrite gerBl. - rewrite -[leRHS]sqrrN opprB. - rewrite ler_sqr ?nnegrE; first last. - rewrite subr_ge0. - apply/ltW; apply: (le_lt_trans yae). - by move: ax; rewrite in_itv/= => /andP[]. +move: ax; rewrite in_itv/= => /andP[aex xae]. +move: aey; move/negP/nandP; rewrite -!leNgt => -[yae|aey]. + rewrite -normrN opprB ger0_norm; last first. + by rewrite subr_ge0 (le_trans yae)// gerBl. + rewrite -[leRHS]sqrrN opprB ler_sqr ?nnegrE; first last. + rewrite subr_ge0 ltW// (le_lt_trans yae)//. by rewrite addrAC subr_ge0. - rewrite addrAC lerD2r. - by apply/ltW; move: ax; rewrite in_itv/= => /andP[]. -have /normr_idP -> : 0 <= y - a. - rewrite subr_ge0; apply: le_trans aey. - by rewrite lerDl. -rewrite ler_sqr ?nnegrE; first last. - rewrite subr_ge0. - apply: le_trans aey. - apply/ltW. - by move: ax; rewrite in_itv/= => /andP[]. + by rewrite addrAC lerD2r ltW. +rewrite ger0_norm; last first. + by rewrite subr_ge0 (le_trans _ aey)// lerDl. +rewrite ler_sqr ?nnegrE; last 2 first. by rewrite -addrA -opprD subr_ge0. -rewrite -addrA -opprD. -rewrite lerD2l lerN2. -by rewrite ltW//; move: ax; rewrite in_itv/= => /andP[]. -Unshelve. end_near. - -(* (note for 3.) -have mlimf : measurable_fun Ys (fun x0 => limn [eta f x x0])%E. - admit. -have H1 : {ae mu, forall x0, Ys x0 -> f x x0 x1 @[x1 --> \oo] --> (limn [eta f x x0])%E}. - admit. - have e x : \bar R := (expR (- x ^+ 2 / (s^+2 *+ 2)))%:E. -have : forall M x0, exists x, (`| f x x0 1 | > M)%E. -move=> M x0. - have H2 : mu.-integrable Ys e. - admit. - have H3 : \forall N \near \oo, {ae mu, forall x0 n, (N <= n)%N -> Ys x0 -> (`|f x x0 n| <= e x0)%E}. - admit. - have [Hdc1 Hdc2 Hdc3] := dominated_convergence mYs mf mlimf H1 H2 H3. - exact: Hdc3. -*) -Admitted. - -(* (note for 1.) - admit. - under [X in limn X]eq_fun do rewrite (fineK (integral_f_fin_num _ _ _)). - over. -rewrite /=. -apply: (emeasurable_fun_cvg (fun n : nat => fun x => (\int[mu]_(x0 in Ys) f x x0 n)%E)). - move=> m/=. - under [X in _ _ X]eq_fun do rewrite -(fineK (integral_f_fin_num _ _ _)); apply/measurable_EFinP => /=. - apply: continuous_measurable_fun. - apply: continuousT_integral => /=. - - admit. - - apply: aeW => /= x. - admit. - - exists (cst (Num.sqrt (s^+2 * pi *+ 2))^-1 * 1). - admit. -move=> x _. -rewrite /f. -rewrite -near_monotone_convergence//; first last. -- admit. -- admit. -- move=> n; apply: measurable_funTS; exact: (measurable_f_second x n). -apply: cvg_near_monotone_convergence => //. -- admit. -- admit. -move=> /= y _. -have := exp_coeff2_near_in_increasing (- (y - x) ^+ 2 / (s ^+ 2 *+ 2)). -move=> [] N _ Hnear. -exists N => //. -move=> k/= Nk n' m'; rewrite !inE/= => kn' km' n'm'. -rewrite lee_fin. -rewrite ler_pM2l; last first. - rewrite invr_gt0. - rewrite sqrtr_gt0. - rewrite pmulrn_lgt0//. - rewrite mulr_gt0//. - by rewrite exprn_even_gt0. - exact: pi_gt0. -by apply: (Hnear k) => //; rewrite !inE/=. -Admitted. - -(* -under [X in _ _ X]eq_fun do rewrite -(fineK (normal_prob_fin_num _ Ys)). -apply/measurable_EFinP. -apply: (@measurability _ _ _ _ _ _ (@RGenOpens.G R)). - by rewrite /measurable/=/measurableR RGenOpens.measurableE. -move=> _ -[_ [a [b ->]] <-]. -rewrite setTI. -apply: is_interval_measurable. -move=> x y/=. -*) -move=> _ -[_ [r r01] [Ys mYs <-]] <-. -apply: emeasurable_fun_infty_o => //=. -under [X in _ _ X]eq_fun. - move=> x. - rewrite -(fineK (normal_prob_fin_num x Ys)). - rewrite normal_shift0//=. - rewrite /pushforward. - rewrite shift_preimage. - rewrite /normal_prob/=. - rewrite integration_by_substitution_shift/=. (* TODO *) - over. -apply: measurableT_comp => //=. -apply/measurable_EFinP => /=. -apply: continuous_measurable_fun => /=. -apply: (@continuousT_integral (fun x0 x1 : R => normal_pdf 0 s (x1 - x0)) Ys). -- move=> z. - exact: integrable_normal_pdf0. -- apply: aeW => /=. - rewrite /measurable/=/g_sigma_algebraType/ocitv_type. - move=> x. - apply: continuous_comp => //=. - apply: continuous_comp. - exact: (@opp_continuous _ R^o). - exact: (@add_continuous x). - exact: continuous_normal_pdf. -exists (cst ((Num.sqrt (s^+2 * pi *+ 2))^-1)). -move=> x _ y/=. -rewrite /normal_pdf; case: ifP => [|_]; first by move/negP: s0. - -admit. -Admitted. -*) - -(* -Lemma measurable_normal_prob2' : -forall U : set R, measurable U -> measurable_fun [set: R] (normal_prob2 ^~ U). -Proof. -move=> U mU. -apply: (@measurability _ _ _ _ _ _ - (@pset _ _ _ : set (set (pprobability _ R)))) => //. -move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. -rewrite /normal_kernel/=. - - -under [X in measurable_fun _ X]eq_fun. - move=> x. - rewrite (_: normal_kernel _ _ = (fine (normal_kernel x U))%:E); last first. - admit. - rewrite normal_shift0/=. - over. -apply: measurableT_comp; last by []. -apply: measurableT_comp; first exact: EFin_measurable. -rewrite /=. -apply: continuous_measurable_fun. -(* -apply: continuous_comp. -under [X in continuous X]eq_fun. - move=> x. - rewrite normal_shift0/=. - over. -+\*) -Admitted. -*) - -(* -HB.instance Definition _ := - isKernel.Build _ _ _ _ _ normal_prob2 measurable_normal_prob2. -*) + by rewrite subr_ge0 (le_trans _ aey)// ltW. +by rewrite -addrA -opprD lerD2l lerN2 ltW. +Unshelve. end_near. Qed. Lemma measurable_normal_prob2 : measurable_fun setT (normal_prob2 : R -> pprobability _ _). @@ -4263,16 +4069,12 @@ apply: emeasurable_fun_infty_o => //=. under [X in _ _ X]eq_fun. move=> x. rewrite -(@fineK _ (normal_prob x s Ys)); last first. -(* rewrite -(@fineK _ (\int[_]_(_ in Ys) _)%E); last exact: integral_f_fin_num. *) rewrite ge0_fin_numE => //. apply: (@le_lt_trans _ _ (normal_prob x s setT)). - by apply: le_measure => //; rewrite inE/=. - apply: (@le_lt_trans _ _ 1%E). - exact: probability_le1. - exact: ltey. + by rewrite le_measure ?inE. + exact: (le_lt_trans (probability_le1 _ _) (ltey _)). over. -apply/measurable_EFinP. -apply: continuous_measurable_fun. +apply/measurable_EFinP; apply: continuous_measurable_fun. exact: normal_prob_continuous. Qed. @@ -4288,7 +4090,7 @@ Proof. apply: (@measurability _ _ _ _ _ _ (@pset _ _ _ : set (set (pprobability _ R)))) => //. move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. -Admitted. +Abort. (* have := measurable_normal_prob2 (integral_normal^~ s) mYs. Qed. From e3527f7bdd098fde5e10bd917bfdf29d925288f0 Mon Sep 17 00:00:00 2001 From: IshYosh <103252572+IshiguroYoshihiro@users.noreply.github.com> Date: Mon, 17 Mar 2025 09:50:26 +0900 Subject: [PATCH 37/48] Prob lang noisy (#29) --- theories/probability.v | 237 ++++++++++++++++++++--------------------- 1 file changed, 113 insertions(+), 124 deletions(-) diff --git a/theories/probability.v b/theories/probability.v index 2b7027287a..dfd92694c0 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -2974,39 +2974,33 @@ Section near_ereal_nondecreasing_is_cvgn. Let G N := ([set n | (N <= n)%N]). -Lemma near_ereal_nondecreasing_is_cvgn (R : realType) (u_ : (\bar R) ^nat) : - (\forall N \near \oo, {in G N &, nondecreasing_seq u_ }) -> cvgn u_. +Lemma ereal_shiftn_nondecreasing_cvgn (R : realType) (u_ : (\bar R)^nat) + (N : nat) : +(* \forall N \near \oo, {in G N &, nondecreasing_seq u_ } + -> u_ @ \oo --> ereal_sup (range (fun n => u_ (n + N))). +*) +{in G N &, nondecreasing_seq u_ } + -> u_ @ \oo --> ereal_sup (range (fun n => u_ (n + N))). Proof. -move=> [N _ H]. -apply/cvg_ex. -exists (ereal_sup (range (fun n => u_ (n + N)))). +move=> H. rewrite -(cvg_shiftn N). apply: ereal_nondecreasing_cvgn. -move=> n m nm. -apply: (H N); rewrite /G ?inE//=. +move=> k m km. +apply: H; rewrite /G ?inE//=. - exact: leq_addl. - exact: leq_addl. - exact: leq_add. Qed. -Lemma near_ereal_nondecreasing_cvgn (R : realType) (u_ : (\bar R)^nat) : -(* - (\forall N \near \oo, {in G N &, nondecreasing_seq u_ }) - -> u_ @ \oo --> limn u_. (* ereal_sup range ? *) -*) -\forall N \near \oo, {in G N &, nondecreasing_seq u_ } - -> u_ @ \oo --> ereal_sup (range (fun n => u_ (n + N))). +Lemma near_ereal_nondecreasing_is_cvgn (R : realType) (u_ : (\bar R) ^nat) : + (\forall N \near \oo, {in G N &, nondecreasing_seq u_ }) -> cvgn u_. Proof. -near=> N. -(* -move=> [N _ H]. +move=> [] N _ H. apply/cvg_ex. -exists (limn (fun n => u_ (n + N))). -rewrite -(cvg_shiftn N). -apply: ereal_nondecreasing_cvgn. -*) -Abort. - +exists (ereal_sup (range (fun n => u_ (n + N)))). +apply: ereal_shiftn_nondecreasing_cvgn. +by apply: (H N); rewrite /G ?inE/=. +Qed. End near_ereal_nondecreasing_is_cvgn. @@ -3213,61 +3207,34 @@ Local Open Scope ring_scope. Local Import Num.ExtraDef. Implicit Types m s x : R. -Definition normalexp m s x := expR (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)). - -Lemma measurable_normalexp m s : measurable_fun setT (normalexp m s). -Proof. -apply: measurableT_comp => //=; apply: measurable_funM => //=. -apply: measurableT_comp => //=; apply: measurable_funX => //=. -exact: measurable_funB. -Qed. - -Lemma normalexp_center m s : normalexp m s = normalexp 0 s \o center m. -Proof. by apply/funext => x/=; rewrite /normalexp/= subr0. Qed. - -Definition normalpi s := (sqrtr (s ^+ 2 * pi *+ 2))^-1. - -Lemma normalpi_ge0 s : 0 <= normalpi s. Proof. by rewrite invr_ge0. Qed. - -Lemma normalpi_gt0 s : s != 0 -> 0 < normalpi s. -Proof. -move=> s0; rewrite invr_gt0// sqrtr_gt0. -by rewrite pmulrn_rgt0// mulr_gt0 ?pi_gt0// exprn_even_gt0/=. -Qed. - -(* NB: If s = 0 then normal_pdf0 is cst 0 for all m and x - * because divr is defined by x / 0 = 0. - *) -Definition normal_pdf0 m s x : R := normalpi s * normalexp m s x. - -Lemma measurable_normal_pdf0 m s : measurable_fun setT (normal_pdf0 m s). -Proof. by apply: measurable_funM => //=; exact: measurable_normalexp. Qed. +Definition normal_pdf0 m s x : R := normal_peak s * normal_fun m s x. Lemma normal_pdf0_center m s : normal_pdf0 m s = normal_pdf0 0 s \o center m. -Proof. by rewrite /normal_pdf0 normalexp_center. Qed. +Proof. by rewrite /normal_pdf0 normal_fun_center. Qed. Lemma normal_pdf0_ge0 m s x : 0 <= normal_pdf0 m s x. -Proof. by rewrite mulr_ge0 ?normalpi_ge0 ?expR_ge0. Qed. +Proof. by rewrite mulr_ge0 ?normal_peak_ge0 ?expR_ge0. Qed. Lemma normal_pdf0_gt0 m s x : s != 0 -> 0 < normal_pdf0 m s x. -Proof. by move=> s0; rewrite mulr_gt0 ?expR_gt0// normalpi_gt0. Qed. +Proof. by move=> s0; rewrite mulr_gt0 ?expR_gt0// normal_peak_gt0. Qed. + +Lemma measurable_normal_pdf0 m s : measurable_fun setT (normal_pdf0 m s). +Proof. by apply: measurable_funM => //=; exact: measurable_normal_fun. Qed. Lemma continuous_normal_pdf0 m s : continuous (normal_pdf0 m s). Proof. -move=> x. -apply: cvgM; first exact: cvg_cst. +move=> x; apply: cvgM; first exact: cvg_cst. apply: (@cvg_comp _ R^o _ _ _ _ (nbhs (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)))); last exact: continuous_expR. -apply: cvgM; last exact: cvg_cst. -apply: (@cvgN _ _ _ _ _ _ ((x - m) ^+ 2 : R^o)). -apply: (@cvg_comp _ _ _ _ (fun x => x ^+ 2)%R _ (nbhs (x - m))). - by apply: (@cvgB _ R^o) => //; exact: cvg_cst. +apply: cvgM; last exact: cvg_cst; apply: (@cvgN _ R^o). +apply: (@cvg_comp _ _ _ _ (@GRing.exp R^~ 2) _ (nbhs (x - m))). + apply: (@cvgB _ R^o) => //; exact: cvg_cst. exact: sqr_continuous. Qed. -Lemma normal_pdf0_ub m s x : normal_pdf0 m s x <= normalpi s. +Lemma normal_pdf0_ub m s x : normal_pdf0 m s x <= normal_peak s. Proof. -rewrite /normal_pdf0 ler_piMr ?normalpi_ge0//. +rewrite /normal_pdf0 ler_piMr ?normal_peak_ge0//. rewrite -[leRHS]expR0 ler_expR mulNr oppr_le0 mulr_ge0// ?sqr_ge0//. by rewrite invr_ge0 mulrn_wge0// sqr_ge0. Qed. @@ -3279,43 +3246,70 @@ Variables (R : realType) (m sigma : R). Local Open Scope ring_scope. Notation mu := lebesgue_measure. -Lemma integral_normalexp : sigma != 0 -> - (\int[mu]_x (normalexp m sigma x)%:E = (normalpi sigma)^-1%:E)%E. -Proof. -move=> sigma_gt0. -pose F (x : R^o) := (x - m) / (`|sigma| * Num.sqrt 2). -have F'E : F^`()%classic = cst (`|sigma| * Num.sqrt 2)^-1. - apply/funext => x; rewrite /F derive1E deriveM// deriveD// derive_cst scaler0. - by rewrite add0r derive_id derive_cst addr0 scaler1. -have := @integralT_gauss R. -rewrite (@increasing_ge0_integration_by_substitutionT _ F gauss_fun)//=; - first last. +Local Notation normal_peak := (normal_peak sigma). +Local Notation normal_fun := (normal_fun m sigma). + +Let measurable_normal_fun : measurable_fun setT normal_fun. +Proof. +apply: measurableT_comp => //=; apply: measurable_funM => //=. +apply: measurableT_comp => //=; apply: measurable_funX => //=. +exact: measurable_funD. +Qed. + +Let F (x : R^o) := (x - m) / (Num.sqrt (sigma ^+ 2 *+ 2)). +Lemma F'E : F^`()%classic = cst (Num.sqrt (sigma ^+ 2 *+ 2))^-1. +Proof. +apply/funext => x; rewrite /F derive1E deriveM// deriveD// derive_cst scaler0. +by rewrite add0r derive_id derive_cst addr0 scaler1. +Qed. + +Let int_gaussFE : sigma != 0 -> + (\int[mu]_x ((((gauss_fun \o F) * + (F^`())%classic) x)%:E * (Num.sqrt (sigma ^+ 2 *+ 2))%:E))%E = +(Num.sqrt (sigma ^+ 2 * pi *+ 2))%:E. +Proof. +move=> s0. +rewrite -mulrnAr -[in RHS]mulr_natr sqrtrM ?(sqrtrM 2) ?sqr_ge0 ?pi_ge0// !EFinM. +rewrite muleCA ge0_integralZr//=; last first. + by move=> x _; rewrite lee_fin mulr_ge0//= ?gauss_fun_ge0// F'E/= invr_ge0. + rewrite F'E; apply/measurable_EFinP/measurable_funM => //. + apply/measurableT_comp => //; first exact: measurable_gauss_fun. + by apply: measurable_funM => //; exact: measurable_funD. +congr *%E. +rewrite -increasing_ge0_integration_by_substitutionT//; first last. - by move=> x; rewrite gauss_fun_ge0. - exact: continuous_gauss_fun. - apply/gt0_cvgMly; last exact: cvg_addrr. - by rewrite invr_gt0// mulr_gt0// normr_gt0. + by rewrite invr_gt0// sqrtr_gt0 -mulr_natr mulr_gt0// exprn_even_gt0. - apply/gt0_cvgMlNy; last exact: cvg_addrr_Ny. - by rewrite invr_gt0// mulr_gt0// normr_gt0. + by rewrite invr_gt0// sqrtr_gt0 -mulr_natr mulr_gt0// exprn_even_gt0. - by rewrite F'E; exact: is_cvg_cst. - by rewrite F'E; exact: is_cvg_cst. - by rewrite F'E => ?; exact: cvg_cst. -- move=> x y. - by rewrite /F ltr_pM2r ?invr_gt0 ?mulr_gt0 ?normr_gt0// ltrBlDr subrK. -move=> /(congr1 (fun x => x * (`|sigma| * Num.sqrt 2)%:E)%E). -rewrite -ge0_integralZr//=; last 2 first. - rewrite F'E; apply/measurable_EFinP/measurable_funM => //. - apply/measurableT_comp => //; first exact: measurable_gauss_fun. - by apply: measurable_funM => //; exact: measurable_funD. - by move=> x _; rewrite lee_fin mulr_ge0//= ?gauss_fun_ge0// F'E/= invr_ge0. -move=> int_gauss; apply: eq_trans. - apply: eq_trans; last exact: int_gauss. - apply: eq_integral => /= x _. - rewrite F'E !fctE/= EFinM -muleA -EFinM mulVf ?mulr1; last first. - by rewrite gt_eqF ?mulr_gt0 ?normr_gt0. - rewrite /gauss_fun /F exprMn exprVn -[in RHS]mulNr. - by rewrite exprMn real_normK ?num_real// sqr_sqrtr// mulr_natr. -rewrite /normalpi invrK -mulrnAl sqrtrM ?mulrn_wge0 ?sqr_ge0//. -by rewrite -[in RHS]mulr_natr sqrtrM ?sqr_ge0// sqrtr_sqr !EFinM muleC. +- move=> x y xy; rewrite /F ltr_pM2r ?ltr_leB//. + rewrite invr_gt0 ?sqrtr_gt0 ?pmulrn_lgt0 ?exprn_even_gt0//. +apply: integralT_gauss. +by rewrite -(mulr_natr (_ ^+ 2)) sqrtrM ?sqr_ge0. +Qed. + +Let integral_normal_pdf_part : sigma != 0 -> + (\int[mu]_x (normal_fun x)%:E = + (Num.sqrt (sigma ^+ 2 * pi *+ 2))%:E)%E. +Proof. +move=> s0; rewrite -int_gaussFE//; apply: eq_integral => /= x _. +rewrite F'E !fctE/= EFinM -muleA -EFinM mulVf ?mulr1; last first. + by rewrite gt_eqF// sqrtr_gt0 pmulrn_lgt0// exprn_even_gt0. +congr (expR _)%:E; rewrite mulNr exprMn exprVn; congr (- (_ * _^-1)%R). +by rewrite sqr_sqrtr// pmulrn_lge0// sqr_ge0. +Qed. + +Let integrable_normal_pdf_part : sigma != 0 -> + mu.-integrable setT (EFin \o normal_fun). +Proof. +move=> s0; apply/integrableP; split. + by apply/measurable_EFinP; apply: measurable_normal_fun. +under eq_integral do rewrite /= ger0_norm ?expR_ge0//. +by rewrite /= integral_normal_pdf_part// ltry. Qed. Local Notation normal_pdf := (normal_pdf m sigma). @@ -3332,18 +3326,9 @@ Qed. Let normal_sigma_additive : semi_sigma_additive normal_prob. Proof. -move=> /= F mF tF mUF. +move=> /= A mA tA mUA. rewrite /normal_prob/= integral_bigcup//=; last first. - apply/integrableP; split. - apply/measurable_funTS/measurableT_comp => //. - exact: measurable_normal_pdf. - rewrite (_ : (fun x => _) = EFin \o normal_pdf); last first. - by apply/funext => x; rewrite gee0_abs// lee_fin normal_pdf_ge0 ?ltW. - apply: le_lt_trans. - apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. - by apply/measurable_EFinP; exact: measurable_normal_pdf. - by move=> ? _; rewrite lee_fin normal_pdf_ge0 ?ltW. - by rewrite integral_normal_pdf // ltey. + apply: (integrableS _ _ (@subsetT _ _)) => //; exact: integrable_normal_pdf. apply: is_cvg_ereal_nneg_natsum_cond => n _ _. by apply: integral_ge0 => /= x ?; rewrite lee_fin normal_pdf_ge0 ?ltW. Qed. @@ -3357,6 +3342,16 @@ Proof. by rewrite /normal_prob integral_normal_pdf. Qed. HB.instance Definition _ := @Measure_isProbability.Build _ _ R normal_prob normal_setT. +Lemma integrable_indic_itv (a b : R) (b0 b1 : bool) : a < b -> + mu.-integrable setT (EFin \o \1_[set` (Interval (BSide b0 a) (BSide b1 b))]). +Proof. +move=> ab. +apply/integrableP; split; first by apply/measurable_EFinP/measurable_indic. +apply/abse_integralP => //; first by apply/measurable_EFinP/measurable_indic. +rewrite integral_indic// setIT/= lebesgue_measure_itv/=. +by rewrite lte_fin ab -EFinD/= ltry. +Qed. + End normal_probability. (* TODO: PR *) @@ -3608,12 +3603,6 @@ apply: continuous_measurable_fun. Local Open Scope ereal_scope. -Lemma integral_normal_prob (f : R -> \bar R) (m : R) U : measurable U -> measurable_fun U f -> - \int[normal_prob2 m]_(x in U) `|f x| < +oo -> - \int[normal_prob2 m]_(x in U) f x = \int[mu]_(x in U) (f x * (normal_pdf m s x)%:E) :> \bar R. -Proof. -Abort. - Local Close Scope ereal_scope. (* @@ -3771,7 +3760,7 @@ Admitted. *) Let g' a e : R -> R := fun x => if x \in (ball a e : set R^o) then - normalpi s else normal_pdf0 e s `|x - a|. + normal_peak s else normal_pdf0 e s `|x - a|. Let ballFE_le (a e x : R) : x <= (a - e)%R -> (x \in (ball a e : set R^o)) = false. @@ -3796,9 +3785,9 @@ Proof. apply/funext => x; rewrite /g'. have /orP [x0|x0] := le_total x a. rewrite ballFE_le; last by rewrite subr0. - by rewrite /normal_pdf0 /normalexp subr0 real_normK// num_real. + by rewrite /normal_pdf0 /normal_fun subr0 real_normK// num_real. rewrite ballFE_ge; last by rewrite addr0. -by rewrite /normal_pdf0 /normalexp subr0 real_normK// num_real. +by rewrite /normal_pdf0 /normal_fun subr0 real_normK// num_real. Qed. Lemma mg' a e : measurable_fun setT (g' a e). @@ -3808,13 +3797,13 @@ apply: measurable_fun_if => //. rewrite setTI preimage_mem_true. exact: measurable_ball. apply: measurable_funTS => /=; apply: measurable_funM => //. -apply: measurableT_comp => //; first exact: measurable_normalexp. +apply: measurableT_comp => //; first exact: measurable_normal_fun. by apply: measurableT_comp => //; exact: measurable_funD. Qed. Lemma g'_ge0 a e x : 0 <= g' a e x. Proof. -rewrite /g'; case: ifP => _; first by rewrite normalpi_ge0. +rewrite /g'; case: ifP => _; first by rewrite normal_peak_ge0. exact: normal_pdf0_ge0. Qed. @@ -3850,8 +3839,8 @@ apply: withinU_continuous. have tae : t < a - e by near: t; exact: lt_nbhsl. rewrite /g'. rewrite !ballFE_le ?(@ltW _ _ _ (a - e))//. - rewrite /normal_pdf0 /normalexp !aNe//. - rewrite -!/(normalexp _ _ _) -!/(normal_pdf0 _ _ _). + rewrite /normal_pdf0 /normal_fun !aNe//. + rewrite -!/(normal_fun _ _ _) -!/(normal_pdf0 _ _ _). move=> {tae}; near: t. move: eps eps0. apply/(@cvgrPdist_le R^o). @@ -3860,7 +3849,7 @@ apply: withinU_continuous. near=> t. rewrite /g' !ballFE_le//. rewrite -addrAC subrr sub0r normrN (ger0_norm e0)//. - rewrite /normal_pdf0 /normalexp subrr -(subrr (a - e)) aNe//. + rewrite /normal_pdf0 /normal_fun subrr -(subrr (a - e)) aNe//. near: t; move: eps eps0. apply/(@cvgrPdist_lt _ R^o). apply: cvg_at_left_filter. @@ -3892,7 +3881,7 @@ apply: withinU_continuous. apply: nbhs_right_lt. by rewrite -(opprK e) ler_ltB// opprK gtrN. rewrite /normal_pdf0. - rewrite addrAC subrr sub0r normrN /normalexp (gtr0_norm e0) subrr. + rewrite addrAC subrr sub0r normrN /normal_fun (gtr0_norm e0) subrr. by rewrite expr0n/= oppr0 mul0r expR0 mulr1 subrr normr0 ltW. + apply/(@cvgrPdist_le _ R^o) => eps eps0. near=> t. @@ -3902,7 +3891,7 @@ apply: withinU_continuous. apply: nbhs_left_gt. by rewrite -(opprK e) ler_ltB// opprK gtrN. rewrite /normal_pdf0. - rewrite addrAC subrr add0r /normalexp (gtr0_norm e0) subrr expr0n oppr0. + rewrite addrAC subrr add0r /normal_fun (gtr0_norm e0) subrr expr0n oppr0. by rewrite mul0r expR0 mulr1 subrr normr0 ltW. - apply/continuous_within_itvcyP; split. + move=> x. @@ -3912,7 +3901,7 @@ apply: withinU_continuous. have tae : a + e < t by near: t; exact: lt_nbhsr. rewrite /g'. rewrite !ballFE_ge ?(@ltW _ _ (a + e)%E)//. - rewrite /normal_pdf0 /normalexp !aDe// ?(@ltW _ _ (a + e)). + rewrite /normal_pdf0 /normal_fun !aDe// ?(@ltW _ _ (a + e)). move=> {tae}; near: t. move: eps eps0. apply/(@cvgrPdist_le _ R^o). @@ -3921,7 +3910,7 @@ apply: withinU_continuous. near=> t. rewrite /g' !ballFE_ge//. rewrite addrAC subrr add0r (ger0_norm e0)//. - rewrite /normal_pdf0 /normalexp subrr -(subrr (a + e)). + rewrite /normal_pdf0 /normal_fun subrr -(subrr (a + e)). rewrite aDe//. near: t. move: eps eps0. @@ -3946,7 +3935,7 @@ under eq_integral. rewrite /g' ballFE_le//; last exact: lerB. rewrite -(normrN (x - e - a)) !opprB addrA. have /normr_idP -> : 0 <= a + e - x by rewrite subr_ge0 ler_wpDr. - rewrite /normal_pdf0 /normalexp. + rewrite /normal_pdf0 /normal_fun. rewrite -(addrAC _ (- x)) addrK. rewrite -(sqrrN (a - x)) opprB. over. @@ -3968,7 +3957,7 @@ under eq_integral. rewrite inE/= in_itv/= andbT => aex. rewrite /g' ballFE_ge//; last exact: lerD. have /normr_idP -> : 0 <= x + e - a by rewrite subr_ge0 ler_wpDr. - rewrite /normal_pdf0 /normalexp -(addrAC _ (- a)) addrK. + rewrite /normal_pdf0 /normal_fun -(addrAC _ (- a)) addrK. over. rewrite /=. by apply: eq_integral => /= x xay; rewrite /normal_pdf (negbTE s0). From fac4728d9836cce5009313936ca09fb5e8d4d145 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 18 Mar 2025 15:23:11 +0900 Subject: [PATCH 38/48] various fixes - gen reproductive - fix naming, doc - minor renaming --- _CoqProject | 1 + theories/Make | 1 + theories/lang_syntax.v | 34 - theories/lang_syntax_examples.v | 1 - theories/lang_syntax_hello_wip.v | 1180 ----------------------------- theories/lang_syntax_noisy.v | 819 ++++++++++++++++++++ theories/lang_syntax_table_game.v | 11 +- theories/prob_lang_wip.v | 62 ++ 8 files changed, 893 insertions(+), 1216 deletions(-) delete mode 100644 theories/lang_syntax_hello_wip.v create mode 100644 theories/lang_syntax_noisy.v diff --git a/_CoqProject b/_CoqProject index bda072280d..e05597f3c0 100644 --- a/_CoqProject +++ b/_CoqProject @@ -143,3 +143,4 @@ theories/lang_syntax_toy.v theories/lang_syntax.v theories/lang_syntax_examples.v theories/lang_syntax_table_game.v +theories/lang_syntax_noisy.v diff --git a/theories/Make b/theories/Make index b17786b9dc..2e259cc607 100644 --- a/theories/Make +++ b/theories/Make @@ -104,3 +104,4 @@ lang_syntax_toy.v lang_syntax.v lang_syntax_examples.v lang_syntax_table_game.v +lang_syntax_noisy.v diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 67da9417a3..ca98d0a4db 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -1671,38 +1671,4 @@ apply: eq_kernel => y V. exact: He. Qed. -Lemma congr_letinl_new {R : realType} g t1 t2 str (e1 e2 : @exp _ _ g t1) - (e : @exp _ _ (_ :: g) t2) x U : - (forall y V, measurable V -> execP e1 y V = execP e2 y V) -> - measurable U -> - @execP R g t2 [let str := e1 in e] x U = - @execP R g t2 [let str := e2 in e] x U. -Proof. -Abort. -(* by move=> + mU; move/eq_sfkernel => He; rewrite !execP_letin He. Qed. *) - -Lemma congr_letinr_new {R : realType} g t1 t2 str (e : @exp _ _ _ t1) - (e1 e2 : @exp _ _ (_ :: g) t2) x U : - (forall y V, measurable V -> execP e1 (y, x) V = execP e2 (y, x) V) -> - @execP R g t2 [let str := e in e1] x U = @execP R g t2 [let str := e in e2] x U. -Proof. -Abort. -(* - by move=> He; rewrite !execP_letin !letin'E; apply: eq_integral => ? _; exact: He. -Qed. -*) -Lemma congr_normalize_new {R : realType} g t (e1 e2 : @exp R _ g t) : - (forall x U, measurable U -> execP e1 x U = execP e2 x U) -> - execD [Normalize e1] = execD [Normalize e2]. -Proof. -Abort. -(* -move=> He; apply: eq_execD. -rewrite !execD_normalize_pt /=. -f_equal. -apply: eq_kernel => y V. -exact: He. -Qed. -*) - Local Close Scope lang_scope. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index cec6569e52..69b6fb3928 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -1023,4 +1023,3 @@ Example letinC_ground (g := [:: ("a", Unit); ("b", Bool)]) t1 t2 Proof. move=> U mU; exact: letinC. Qed. End letinC. - diff --git a/theories/lang_syntax_hello_wip.v b/theories/lang_syntax_hello_wip.v deleted file mode 100644 index d2a12b6db1..0000000000 --- a/theories/lang_syntax_hello_wip.v +++ /dev/null @@ -1,1180 +0,0 @@ -From Coq Require Import String. -From HB Require Import structures. -From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. -From mathcomp.classical Require Import mathcomp_extra boolp. -From mathcomp Require Import ring lra. -From mathcomp Require Import classical_sets functions cardinality fsbigop. -From mathcomp Require Import interval_inference reals ereal topology normedtype. -From mathcomp Require Import sequences esum measure lebesgue_measure numfun exp. -From mathcomp Require Import trigo realfun charge lebesgue_integral kernel. -From mathcomp Require Import probability prob_lang. -From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. - -(**md**************************************************************************) -(* # Inferring behavior from text-massage data example *) -(* *) -(* ref: *) -(* - Chung-chieh Shan, Equational reasoning for probabilistic programming, *) -(* POPL TutorialFest 2018 *) -(* https://homes.luddy.indiana.edu/ccshan/rational/equational-handout.pdf *) -(* - *) -(* *) -(******************************************************************************) - -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -Import Order.TTheory GRing.Theory Num.Def Num.ExtraDef Num.Theory. -Import numFieldTopology.Exports. - -Local Open Scope classical_set_scope. -Local Open Scope ring_scope. -Local Open Scope ereal_scope. - - -Definition neq01 {R : realType} := (@GRing.oner_neq0 R). -Definition exp_normal1 {R g} := (@exp_normal R g _ neq01). - -Definition neq0Vsqrt2 {R : realType} : ((@Num.sqrt R 2)^-1 != 0)%R. -Proof. exact: lt0r_neq0. Qed. -Definition exp_normal_Vsqrt2 {R g} := (@exp_normal R g _ neq0Vsqrt2). - -Definition neq0sqrt32 {R : realType} : ((@Num.sqrt R (3 / 2)) != 0)%R. -Proof. exact: lt0r_neq0. Qed. -Definition exp_normal_sqrt32 {R g} := (@exp_normal R g _ neq0sqrt32). - -(* -Section eq_lebesgue_measure_measurable. -Context {R : realType}. - -Local Notation mu := lebesgue_measure. - -Variable f : cumulative R. -Hypothesis intf : integrable mu setT (EFin \o f). - -Local Notation lsmf := (lebesgue_stieltjes_measure f). - -Let fin_lsmf : fin_num_fun lsmf. -Proof. -apply: lty_fin_num_fun. -rewrite /=/lsmf/measure_extension/mu_ext. -Admitted. - -HB.instance Definition _ := @Measure_isFinite.Build _ _ _ lsmf fin_lsmf. - -Lemma Radon_Nikodym_ae_unique : - ae_eq mu setT (EFin \o f) ('d lsmf '/d mu). -Proof. -apply: integral_ae_eq => //. - admit. -move=> E _ mE. -rewrite -Radon_Nikodym_integral//; last first. - admit. - - -Lemma eq_lebesgue_measure_measurable (m' : measure _ R) : -sigma_additive m' -> -(forall U : set _, measurable U -> m' U = mu U) -> -m' = mu. -Proof. -move=> sm' H. -apply: eq_measure. -apply/funext => U. -have : setT U by []. -rewrite -(setUv measurable). -move=> [mU|nmU]; first exact: H. -pose e := m' U. -have : m' U = e by []. -Abort. - -End eq_lebesgue_measure_measurable. -*) - -(* TODO: PR *) -Section ge0_fin_measurable_probability_integrable. -Context d {T : measurableType d} {R : realType} {p : probability T R} - {f : T -> \bar R}. - -Lemma ge0_fin_measurable_probability_integrable : - (forall x, 0 <= f x) -> - (exists M : R, forall x, f x <= M%:E) -> - measurable_fun setT f -> - p.-integrable setT f. -Proof. -move=> f0 [M fleM] mf. -apply/integrableP; split => //. -apply (@le_lt_trans _ _ (\int[p]_x M%:E)). - apply: ge0_le_integral => //=. - exact: measurableT_comp. - move=> t _. - exact: (@le_trans _ _ (f t)). - move=> t _. - by rewrite gee0_abs. -by rewrite integral_cst// probability_setT mule1 ltry. -Qed. - -End ge0_fin_measurable_probability_integrable. - -(* TODO: PR *) -Section pkernel_probability_integrable. -Context d d' {T : measurableType d} {T' : measurableType d'} {R : realType} - {p : probability T R} {f : R.-pker T ~> T'}. - -Lemma pkernel_probability_integrable V : - measurable V -> - p.-integrable setT (fun x => f x V). -Proof. -move=> mV. -apply: ge0_fin_measurable_probability_integrable => //. - exists 1%R. - move=> x. - apply: (@le_trans _ _ (f x setT)). - by apply: le_measure; rewrite ?inE. - by rewrite prob_kernel. -exact: measurable_kernel. -Qed. - -End pkernel_probability_integrable. - -(* TODO: move to probability? *) -Section normal_prob_properties. -Context {R : realType}. -Local Notation mu := lebesgue_measure. - -Local Open Scope charge_scope. - -Lemma normal_pdf_uniq_ae (m s : R) (s0 : (0 < s)%R) : - ae_eq mu setT - ('d ((charge_of_finite_measure (@normal_prob R m s))) '/d mu) - (EFin \o (@normal_pdf R m s)). -Proof. -apply: integral_ae_eq => //. - apply: Radon_Nikodym_integrable => /=. - exact: normal_prob_dominates. - apply/measurable_EFinP. - exact: measurable_normal_pdf. -move=> /= E _ mE. -by rewrite -Radon_Nikodym_integral//=; apply: normal_prob_dominates. -Qed. - -End normal_prob_properties. - -(* -Section disintegration_measure. -Context {R : realType}. -(* define lebesgue syntax *) - -End disintegration_measure. - -Section disintegration_program. - -Variable l : ctx. - -Definition lhs : exp R _ l _ := - [let "t" := Sample {exp_normal ltr01 (exp_real 0)} in - let "p" := (* ? TODO *) in - return (#{"t"}, #{"p"})]. - -Definition rhs : exp R _ l _ := - [let "t" := Sample exp_lebesgue(* TODO *) in - let "p" := Score (exp_normal_pdf [#{"t"}])) (* TODO *) in - return (#{"t"}, #{"p"})]. - -Lemma disintegration_normal : -execD lhs = exec rhs. - -End disintegration_program. -*) - -Section normal_prob_lemmas. -Context {R: realType}. -Local Notation mu := lebesgue_measure. - -(* TODO: move *) -Lemma emeasurable_normal_prob (s : R) U : s != 0%R -> measurable U -> - measurable_fun setT (fun x => normal_prob x s U). -Proof. -move=> s0 mU. -under [X in _ _ X]eq_fun. - move=> x. - rewrite -(@fineK _ (_ x _ _)); last first. - rewrite ge0_fin_numE//. - apply: (@le_lt_trans _ _ (normal_prob x s setT)). - by apply: le_measure; rewrite ?inE. - by rewrite probability_setT ltry. - over. -apply/measurable_EFinP. -apply: (continuous_measurable_fun). -exact: normal_prob_continuous. -Qed. - -(* TODO: s != 0 *) -Lemma integral_normal_prob (m s : R) (s0 : (0 < s)%R) f U : - measurable U -> - (normal_prob m s).-integrable U f -> - \int[@normal_prob _ m s]_(x in U) f x = - \int[mu]_(x in U) (f x * (normal_pdf m s x)%:E). -Proof. -move=> mU intf. -move/integrableP : (intf) => [mf intf_lty]. -rewrite -(Radon_Nikodym_change_of_variables (normal_prob_dominates m s))//=. -apply: ae_eq_integral => //=. - apply: emeasurable_funM => //. - apply: measurable_funTS. - have : charge_of_finite_measure (normal_prob m s) `<< mu. - exact: normal_prob_dominates m s. - move/Radon_Nikodym_integrable. - by move/integrableP => []. - apply: emeasurable_funM => //. - apply/measurable_EFinP. - apply: measurable_funTS. - exact: measurable_normal_pdf. -apply: ae_eq_mul2l. -apply: (ae_eq_subset (@subsetT _ U)). -exact: (normal_pdf_uniq_ae m s0). -Qed. - -(* -Definition tail1 : @exp R _ [:: ("x", Real)] _ := - [let "z" := Sample {exp_normal ltr01 [#{"x"}]} in - return #{"z"}]. - -Lemma helloRight01 : -execP [let "x" := Sample {exp_normal ltr01 (exp_real 0)} in - let "_" := Score {expR 1} `^ - ({0}:R - (#{"y0"} - #{"x"}) ^+ {2} * {2^-1}:R) - * {(Num.sqrt 2 * pi)^-1}:R in - {exp_weak _ [::] [:: ("x", Real); ("y0", Real)] ("_", Unit) - (exp_weak _ [:: ("x", Real)] [::] ("y0", Real) tail1 _)}] = - [let "_" := Score {expR 1} `^ ({0}:R - #{"y0"} ^+ {2} * {4^-1}:R) * - {(Num.sqrt (4 * pi))^-1}:R in - let "x" := Sample {exp_normal ltr0Vsqrt2 [#{"y0"} * {2^-1}:R ]} in - {tail1}]. -*) - -Lemma normal_prob_integrable_dirac (m s : R) (V : set R): measurable V -> - (normal_prob m s).-integrable setT (fun x => \d_x V). -Proof. -move=> mV. -apply/integrableP; split; first exact: measurable_fun_dirac. -rewrite -(setUv V). -rewrite ge0_integral_setU => //; first last. - exact/disj_setPCl. - rewrite setUv. - apply: measurableT_comp => //. - exact: measurable_fun_dirac. - exact: measurableC. -under eq_integral. - move=> x Vx. - rewrite diracE Vx/= normr1. - over. -under [X in _ + X < _]eq_integral. - move=> x. - rewrite inE/= => nVx. - have {}nVx := (memNset nVx). - rewrite indicE nVx/= normr0. - over. -rewrite !integral_cst//=; last exact: measurableC. -rewrite mul1e mul0e adde0. -apply: (le_lt_trans (probability_le1 (normal_prob m s) mV)). -exact: ltey. -Qed. - -Lemma integral_normal_prob_dirac (s : R) (m : R) V : - (0 < s)%R -> - measurable V -> - \int[normal_prob m s]_x0 (\d_x0 V) = normal_prob m s V. -Proof. -move=> s0 mV. -rewrite integral_normal_prob => //; last exact: (normal_prob_integrable_dirac m s). -under eq_integral do rewrite diracE. -rewrite /normal_prob. -rewrite [in RHS]integral_mkcond. -under [in RHS]eq_integral do rewrite patchE. -rewrite /=. -apply: eq_integral => x _. -by case: ifP => xV/=; rewrite ?mul1e ?mul0e. -Qed. - -End normal_prob_lemmas. - -Section hello_programs. -Local Open Scope lang_scope. -Context {R : realType}. -Local Notation mu := lebesgue_measure. - -Definition guard_real {g} str (r : R) : - @exp R P [:: (str, _) ; g] _ := - [if #{str} ==R {r}:R then return TT else Score {0}:R]. - -Definition helloWrong (y0 : R) : @exp R _ [::] _ := - [Normalize - let "x" := Sample {exp_normal1 (exp_real 0)} in - let "y" := Sample {exp_normal1 [#{"x"}]} in - let "_" := {guard_real "y" y0} in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return #{"z"}]. - -Definition helloRight : @exp R _ [:: ("y0", Real)] _ := - [Normalize - let "x" := Sample {exp_normal1 (exp_real 0)} in - let "_" := Score ({expR 1} `^ - ({0}:R - (#{"y0"} - #{"x"}) ^+ {2} * {2^-1}:R)) - * {(Num.sqrt 2 * pi)^-1}:R in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return #{"z"}]. - -Definition helloJoint : @exp R _ [::] _ := - [Normalize - let "x" := Sample {exp_normal1 (exp_real 0)} in - let "y" := Sample {exp_normal1 [#{"x"}]} in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return (#{"y"}, #{"z"})]. - -End hello_programs. - -Section helloRight_subproofs. -Local Open Scope lang_scope. -Context {R : realType}. -Local Notation mu := lebesgue_measure. - -Local Definition int_normal_helloRightP := - (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - \int[normal_prob 0 1]_x (fun x0 => - `|expR (- ((y.1 - x0) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * normal_prob x0 1 V) x). - -Local Definition int_mu_helloRightP := - (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - \int[mu]_x - (`|expR (- ((y.1 - x) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * normal_prob x 1 V * - (normal_pdf 0 1 x)%:E)). - -Local Definition int_normal_helloRight1P := -(fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => -(\int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x -(((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))))%:num)%:E * - normal_prob x 1 V)%E)). - -Local Definition int_mu_helloRight1P := -(fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - \int[mu]_x - (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E * - (normal_prob x 1 V * (normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 x)%:E))). - -Local Definition int_normal_helloRight2P := -(fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - \int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x normal_prob x 1 V). - -Local Definition int_mu_helloRight2P := -(fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - normal_prob (y.1 / 2) (Num.sqrt (3 / 2)) V). - -Lemma int_change_helloRightP y V : measurable V -> -int_normal_helloRightP y V = int_mu_helloRightP y V. -Proof. -move=> mV; rewrite /int_normal_helloRightP. -rewrite integral_normal_prob//; first last. - apply: integrableMr => //. - apply: (@measurableT_comp _ _ _ _ _ _ (@normr _ R^o)) => //. - apply: measurable_funM => //. - apply: measurableT_comp => //. - apply: measurableT_comp => //. - apply: measurable_funM => //. - apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => (t ^+ 2%R)%R)) => //. - exact: measurable_funB. - exists (Num.sqrt (2 * pi))^-1%R; split; first exact: num_real. - move=> x x_gt. - move=> p _. - rewrite /= normr_id. - have /normr_idP -> : (0 <= expR (- ((y.1 - p) ^+ 2%R / 2)) / (Num.sqrt (2 * pi)))%R. - apply: mulr_ge0; first exact: expR_ge0. - by rewrite invr_ge0 sqrtr_ge0// mulr_ge0// pi_ge0. - apply/ltW. - apply: le_lt_trans x_gt. - rewrite -[leRHS]mul1r ler_pM//. - exact: expR_ge0. - by rewrite -expR0 ler_expR oppr_le0 mulr_ge0// ?sqr_ge0// expR0 invr_ge0. - apply/integrableP; split. - exact: emeasurable_normal_prob. - apply/abse_integralP => //. - exact: emeasurable_normal_prob. - have/gee0_abs -> : 0 <= \int[normal_prob 0 1]_x normal_prob x 1 V. - exact: integral_ge0. - apply: (@le_lt_trans _ _ (\int[normal_prob 0 1]_x (cst 1%R x)%:E)). - apply: ge0_le_integral => //. - exact: emeasurable_normal_prob. - apply/measurable_EFinP. - exact: measurable_cst. - move=> x _. - exact: probability_le1. - by rewrite /= integral_cst// mul1e probability_setT ltry. -Qed. - -Lemma int_change_helloRight1P y V : measurable V -> -int_normal_helloRight1P y V = int_mu_helloRight1P y V. -Proof. -move=> mV; rewrite /int_normal_helloRight1P. -rewrite integral_normal_prob//; first last. - apply/integrableP; split. - apply: measurable_funeM. - exact: emeasurable_normal_prob. - apply: (@le_lt_trans _ _ - (\int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x - (((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / - Num.sqrt (4 * pi))))%:num)%:E * (cst 1%R x)%:E))). - apply: ge0_le_integral => //. - apply: measurableT_comp => //. - apply: measurable_funeM. - exact: emeasurable_normal_prob. - by move=> x _; rewrite /= mule1. - rewrite /= mule1. - exact/measurable_EFinP. - move=> x _/=. - rewrite gee0_abs; last exact: mule_ge0. - rewrite lee_pmul//. - exact: probability_le1. - rewrite integralZr//; last exact: finite_measure_integrable_cst. - by rewrite integral_cst// mule1 probability_setT ltry. -by under eq_integral do rewrite -muleA. -Qed. - -Lemma int_change_helloRight2P y V : - measurable V -> - int_normal_helloRight2P y V = int_mu_helloRight2P y V. -Proof. -move=> mV. -rewrite /int_normal_helloRight2P. -rewrite /int_mu_helloRight2P. -(* rewrite by fubini *) -transitivity (\int[mu]_(z in V) - \int[mu]_x ( - (normal_pdf x 1 z * normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 x)%:E)). - rewrite integral_normal_prob//; last first. - apply: ge0_fin_measurable_probability_integrable => //. - exists 1%R => ?; exact: probability_le1. - exact: emeasurable_normal_prob. - under eq_integral. - move=> x _. - rewrite /normal_prob. - rewrite -integralZr//; last first. - apply: (integrableS measurableT) => //. - exact: integrable_normal_pdf. - under eq_integral. - move=> z _. - rewrite -EFinM. - rewrite [X in X%:E](_:_= - (fun xz : R * R => (normal_pdf xz.1 1 xz.2 * normal_pdf (y.1 / 2) - (Num.sqrt 2)^-1 xz.1)%R) (x, z)); last by []. - over. - rewrite integral_mkcond. - rewrite epatch_indic. - over. - rewrite /=. - rewrite (@fubini_tonelli _ _ _ _ _ mu mu -(EFin \o ((fun xz : R * R => (normal_pdf xz.1 1 xz.2 * - normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 xz.1)%R * \1_V xz.2)%R))); - first last. - move=> x. - rewrite lee_fin/=. - apply: mulr_ge0 => //. - apply: mulr_ge0; exact: normal_pdf_ge0. - apply/measurable_EFinP. - apply: measurable_funM; last first. - apply: measurable_indic. - rewrite -[X in measurable X]setTX. - exact: measurableX. - apply: measurable_funM. - under [X in measurable_fun _ X]eq_fun. - move=> x. - rewrite normal_pdfE//. - rewrite normal_pdf0_center/=. - over. - rewrite /=. - apply: measurableT_comp. - exact: measurable_normal_pdf0. - exact: measurable_funB. - apply: measurableT_comp => //. - exact: measurable_normal_pdf. - under eq_integral. - move=> x _. - rewrite /=. - under eq_integral do rewrite EFinM. - rewrite -epatch_indic. - over. - rewrite /=. - rewrite [RHS]integral_mkcond/=. - apply: eq_integral. - move=> /= x _. - rewrite patchE. - case: ifPn => xV. - apply: eq_integral => z _/=. - by rewrite ifT. - apply: integral0_eq => /= z _. - rewrite ifF//. - apply/negbTE. - rewrite notin_setE/=. - move/negP in xV. - by move/mem_set. -apply: eq_integral. -move=> x _. -transitivity ((pi * (Num.sqrt 2))^-1%:E * - \int[mu]_z (expR (- (x - z) ^+ 2 / 2) * expR (- (z - y.1 / 2) ^+ 2))%:E). - rewrite -ge0_integralZl//=; last 3 first. - - apply/measurable_EFinP => //=; apply: measurable_funM => //=. - + apply: measurableT_comp => //=; apply: measurable_funM => //=. - apply/measurableT_comp => //; apply: measurable_funX. - exact: measurable_funB. - + do 2 apply: measurableT_comp => //=. - by apply: measurable_funX; exact: measurable_funD. - - by move=> z _; rewrite lee_fin mulr_ge0// expR_ge0. - - by rewrite lee_fin// invr_ge0// mulr_ge0// pi_ge0. - apply: eq_integral => /= z _. - rewrite /normal_pdf oner_eq0 gt_eqF; last by rewrite invr_gt0. - rewrite /normal_pdf0 expr1n mul1r exprVn sqr_sqrtr//. - rewrite -mulrnAr -(mulr_natr pi) -(mulr_natr 2^-1). - rewrite (mulrCA _ pi) mulVf ?mulr1 ?divr1// mulrACA; congr ((_ * _)%:E). - rewrite sqrtrM ?pi_ge0// !invfM// mulrAC; congr (_ / _). - by rewrite -[LHS]invfM -expr2 sqr_sqrtr ?pi_ge0. -rewrite normal_pdfE// /normal_pdf0. -transitivity (((pi * Num.sqrt 2)^-1)%:E * - \int[mu]_x0 (expR (- (3 / 2)%R * (x0 - (x + y.1) / 3) ^+ 2 - (x - y.1 / 2) ^+ 2 / (3 / 2 *+ 2)))%:E). - congr *%E. - apply: eq_integral. - move=> z _. - rewrite -expRD. - congr EFin. - congr expR. - rewrite !sqrrD. - lra. -(* gauss integral *) -under eq_integral do rewrite expRD EFinM. -rewrite ge0_integralZr//; first last. - rewrite lee_fin. - exact: expR_ge0. - move=> z _. - rewrite lee_fin. - exact: expR_ge0. - apply/measurable_EFinP. - apply: measurableT_comp => //. - apply: measurable_funM => //. - apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)%R) => //. - exact: measurable_funD. -rewrite [in RHS]sqr_sqrtr//. -rewrite -[in RHS]mulr_natr [in RHS]mulrAC divfK//. -rewrite mulNr EFinM muleA; congr *%E. -rewrite [X in _ * X = _](_ : _ = (Num.sqrt ((1 / 3) * pi *+ 2))%:E * - \int[mu]_z (normal_pdf ((x + y.1) / 3) (Num.sqrt (1 / 3)) z)%:E); last first. - rewrite -ge0_integralZl//=; last 2 first. - by apply/measurable_EFinP; exact: measurable_normal_pdf. - by move=> /= z _; rewrite lee_fin normal_pdf_ge0. - apply: eq_integral => /= z _. - rewrite /normal_pdf gt_eqF// /normal_pdf0 sqr_sqrtr// -EFinM; congr EFin. - rewrite [RHS]mulrA -[LHS]mul1r; congr (_ * expR _)%R. - by rewrite divff// gt_eqF// sqrtr_gt0 pmulrn_rgt0// mulr_gt0// pi_gt0. - rewrite -(mulr_natl (1 / 3)) mul1r. - by rewrite !mulNr (mulrC (3 / 2)) invfM invrK (mulrC (2^-1)). -rewrite integral_normal_pdf mule1 -EFinM mul1r; congr EFin. -rewrite -(mulr_natr (_ * pi)) sqrtrM ?mulr_ge0 ?pi_ge0//. -rewrite invfM mulrCA -mulrA mulVf ?mulr1 ?gt_eqF//. -rewrite !sqrtrM// invfM sqrtrV// -mulrA; congr *%R. -rewrite -[X in _ / X]sqr_sqrtr ?pi_ge0//. -by rewrite -exprVn expr2 mulrCA divff ?mulr1// sqrtr_eq0 leNgt pi_gt0. -Qed. - -End helloRight_subproofs. - -Section helloRight_proof. -Local Open Scope lang_scope. -Context {R : realType}. - -Local Notation mu := lebesgue_measure. - -(* TODO: remove P, helloRight -> helloRight0 *) -(* NB: exp_powR level setting is mistaken? *) -(* ((_ `^ _) * _) cannot write as (_ `^ _ * _) *) -Definition helloRightP : @exp R P [:: ("y0", Real)] Real := -[let "x" := Sample {exp_normal1 (exp_real 0)} in - let "_" := Score ({expR 1} `^ - ({0}:R - (#{"y0"} - #{"x"}) ^+ {2} * {2^-1}:R)) - * {(Num.sqrt (2 * pi))^-1}:R in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return #{"z"}]. - -(* rename *) -Definition helloRight1P : @exp R _ [:: ("y0", Real)] Real := - [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2} * {4^-1}:R)) * - {(Num.sqrt (4 * pi))^-1}:R in - let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R ]} in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return #{"z"}]. - -Definition helloRight1 : @exp R _ [:: ("y0", Real)] _ := - [Normalize {helloRight1P}]. - -Definition helloRight2P : @exp R _ [:: ("y0", Real)] _ := - [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2} * {4^-1}:R)) * - {(Num.sqrt (4 * pi))^-1}:R in - Sample {exp_normal_sqrt32 [#{"y0"} * {2^-1}:R]}]. (* ? *) - -Definition helloRight2 : @exp R _ [:: ("y0", Real)] _ := - [Normalize {helloRight2P}]. - -(* TODO: separate program and proof each line -Local Definition helloRightP_line1 : - forall _, @exp R P [:: ("y0", Real)] Real := -(fun (t : @exp R P [:: ("x", Real); ("y0", Real)] Real) => -[let "x" := Sample {exp_normal1 (exp_real 0)} in - {t}]). -*) - -Lemma execP_helloRight12 y V : measurable V -> - @execP R [:: ("_", Unit); ("y0", Real)] _ - [let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R ]} in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return #{"z"}] y V = - @execP R [:: ("_", Unit); ("y0", Real)] _ - [Sample {exp_normal_sqrt32 [#{"y0"} * {2^-1}:R]}] y V. -Proof. -move=> mV. -rewrite !execP_letin. -rewrite !execP_sample !execD_normal/=. -rewrite (@execD_bin _ _ binop_mult) execD_real/=. -rewrite execP_return. -rewrite !exp_var'E (execD_var_erefl "y0") (execD_var_erefl "x") (execD_var_erefl "z")/=. -rewrite !letin'E/=. -under eq_integral do rewrite letin'E/=. -rewrite /=. -(* *) -under eq_integral do rewrite integral_normal_prob_dirac//. -exact: (int_change_helloRight2P _ mV). -Qed. - -(* ref: https://homes.luddy.indiana.edu/ccshan/rational/equational-handout.pdf - * p.2, equation (7) - * change sampling order by bayes' theorem. - *) - -Local Definition executed_helloRightP := - (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - letin' - (sample (fun=> normal_prob 0 1) - (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) (kr 0))) - (letin' - (score - (measurable_funM - (measurableT_comp (measurable_powRr (expR 1)) - (measurable_funB (kr 0) - (measurable_funM - (measurable_funX 2%R - (measurable_funB (measurable_acc_typ [:: Real; Real] 1) - (measurable_acc_typ [:: Real; Real] 0))) - (kr 2^-1)))) (kr (Num.sqrt (2 * pi))^-1))) - (letin' - (sample - (fun - x : unit * - (g_sigma_algebraType R.-ocitv.-measurable * - (g_sigma_algebraType R.-ocitv.-measurable * unit)) => - normal_prob x.2.1 1) - (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) - (measurable_acc_typ [:: Unit; Real; Real] 1))) - (ret (measurable_acc_typ [:: Real; Unit; Real; Real] 0)))) y V : \bar R). - -Local Definition executed_helloRight1P := - (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - letin' - (score - (measurable_funM - (measurableT_comp (measurable_powRr (expR 1)) - (measurable_funB (kr 0) - (measurable_funM (measurable_funX 2%R (measurable_acc_typ [:: Real] 0)) (kr 4^-1)))) - (kr (Num.sqrt (4 * pi))^-1))) - (letin' - (sample - (fun x : unit * (g_sigma_algebraType R.-ocitv.-measurable * unit) => - probability_normal_prob__canonical__measure_Probability (x.2.1 / 2) (Num.sqrt 2)^-1) - (measurableT_comp (measurable_normal_prob2 neq0Vsqrt2) - (measurable_funM (measurable_acc_typ [:: Unit; Real] 1) (kr 2^-1)))) - (letin' - (sample - (fun - x : g_sigma_algebraType R.-ocitv.-measurable * - (unit * (g_sigma_algebraType R.-ocitv.-measurable * unit)) => - probability_normal_prob__canonical__measure_Probability x.1 1) - (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) - (measurable_acc_typ [:: Real; Unit; Real] 0))) - (ret (measurable_acc_typ [:: Real; Real; Unit; Real] 0)))) y V : \bar R). - -(* TODO: rename, this is step0 of lhs *) -Lemma lhs_execPE y V : - execP helloRightP y V = executed_helloRightP y V. -Proof. -rewrite !execP_letin. -rewrite execP_sample. -rewrite execD_normal/=. -rewrite execD_real/=. -rewrite execP_score. -rewrite (@execD_bin _ _ binop_mult)/=. -rewrite execD_pow_real/=. -rewrite (@execD_bin _ _ binop_minus)/=. -rewrite execD_real/=. -rewrite (@execD_bin _ _ binop_mult)/=. -rewrite execD_pow/=. -rewrite (@execD_bin _ _ binop_minus)/=. -rewrite exp_var'E/= (execD_var_erefl "y0")/=. -rewrite exp_var'E/= (execD_var_erefl "x")/=. -rewrite execD_real/=. -rewrite execD_real/=. -rewrite execP_sample. -rewrite execD_normal/=. -rewrite exp_var'E/= (execD_var_erefl "x")/=. -rewrite execP_return/=. -rewrite exp_var'E/= (execD_var_erefl "z")/=. -done. -Qed. - -Lemma rhs_execPE y V : - execP helloRight1P y V = executed_helloRight1P y V. -Proof. -rewrite execP_letin. -rewrite execP_score. -rewrite (@execD_bin _ _ binop_mult)/=. -rewrite execD_pow_real/=. -rewrite (@execD_bin _ _ binop_minus)/=. -rewrite execD_real/=. -rewrite (@execD_bin _ _ binop_mult)/=. -rewrite execD_pow/=. -rewrite exp_var'E/= (execD_var_erefl "y0")/=. -rewrite execD_real/=. -rewrite execD_real/=. -rewrite execP_letin. -rewrite execP_sample. -rewrite execD_normal/=. -rewrite (@execD_bin _ _ binop_mult)/=. -rewrite exp_var'E/= (execD_var_erefl "y0")/=. -rewrite execD_real/=. -rewrite execP_letin. -rewrite execP_sample. -rewrite execD_normal/=. -rewrite exp_var'E/= (execD_var_erefl "x")/=. -rewrite execP_return. -rewrite exp_var'E/= (execD_var_erefl "z")/=. -done. -Qed. - -Lemma helloRightP_intE y V : measurable V -> - executed_helloRightP y V = int_normal_helloRightP y V. -Proof. -move=> mV. -rewrite /executed_helloRightP. -rewrite [in LHS]letin'E/=. -under [in LHS]eq_integral. - move=> x _. - rewrite letin'E/=. - under eq_integral. - move=> u _. - rewrite letin'E/=. - rewrite integral_normal_prob_dirac//. - over. - rewrite /=. - rewrite ge0_integral_mscale/=; first last. - - by move=> ? _. - - by []. - - exact: measurableT. - rewrite integral_dirac; first last. - - by []. - - exact: measurableT. - rewrite diracT mul1e. - rewrite sub0r. - rewrite -expRM mul1r. - over. -rewrite /=. -done. -Qed. - -Lemma helloRight1P_intE y V : measurable V -> - executed_helloRight1P y V = int_normal_helloRight1P y V. -Proof. -move=> mV; rewrite /executed_helloRight1P. -rewrite letin'E/=. -under eq_integral. - move=> u _. - rewrite letin'E/=. - under eq_integral. - move=> x _. - rewrite letin'E/=. - rewrite integral_normal_prob_dirac//. - over. - over. -rewrite ge0_integral_mscale//; first last. - move=> ? _. - exact: integral_ge0. -rewrite integral_dirac//. -rewrite diracT mul1e. -rewrite sub0r -expRM mul1r. -rewrite /=. -rewrite -(@ge0_integralZl _ _ R (normal_prob _ _) _ measurableT _ _ (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E) )//; first last. -exact: emeasurable_normal_prob. -Qed. - -(* TODO: rename *) -Lemma execP_helloRight0_to_1 y V: -measurable V -> -execP helloRightP y V = execP helloRight1P y V. -Proof. -move=> mV. -(* lhs *) -rewrite lhs_execPE. -rewrite (helloRightP_intE _ mV). -rewrite (int_change_helloRightP _ mV). -(* rhs *) -rewrite rhs_execPE. -rewrite (helloRight1P_intE _ mV). -rewrite (int_change_helloRight1P _ mV). -(* eq_integral *) -apply: eq_integral. -move=> x _. -rewrite [in LHS]muleAC. -rewrite [in RHS](muleC (normal_prob x 1 V)) muleA. -congr (fun t => (t * normal_prob x 1 V)%E). -rewrite [in LHS]ger0_norm; last first. -- by rewrite mulr_ge0// expR_ge0. -rewrite [in RHS]ger0_norm; last first. -- by rewrite mulr_ge0// expR_ge0. -rewrite -!EFinM; congr EFin. -rewrite !normal_pdfE// /normal_pdf0. -rewrite mulrA. -rewrite mulrAC. -rewrite -(@sqrtrV _ 2)//. -rewrite sqr_sqrtr; last by rewrite invr_ge0. -rewrite subr0. -rewrite [X in X / _ = _]mulrC mulrA -expRD -mulrA -invfM. -rewrite [RHS]mulrAC [X in _ = _ * X / _]mulrC mulrA -mulrA -expRD -invfM. -congr *%R. - congr expR; lra. -congr GRing.inv. -rewrite expr1n mul1r -mulrnAr -(mulr_natl pi) mulrA mulVf// mul1r -expr2. -rewrite sqr_sqrtr; last by rewrite mulr_ge0// pi_ge0. -rewrite !sqrtrM// mulrCA -expr2 sqr_sqrtr; last exact: pi_ge0. -congr *%R. -have /normr_idP <- : (@GRing.zero R <= 2)%R by []. -rewrite -(sqrtr_sqr 2%R). -congr Num.sqrt. -lra. -Qed. - -Lemma helloRight0_to_2 y V : measurable V -> -execP helloRightP y V = execP helloRight2P y V. -Proof. -move=> mV. -rewrite execP_helloRight0_to_1//. -rewrite /helloRight1P/helloRight2P. -rewrite execP_letin/= [in RHS]execP_letin/=. -rewrite letin'E [RHS]letin'E. -by under eq_integral do rewrite execP_helloRight12//. -Qed. - -(* -Lemma helloRight0_to_1 : -execD helloRight = execD helloRight1. -Proof. -apply: congr_normalize => y V. -(* lhs *) -rewrite ![in LHS]execP_letin. -rewrite [in LHS]execP_sample. -rewrite [in LHS]execD_normal/=. -rewrite [in LHS]execD_real/=. -rewrite [in LHS]execP_score. -rewrite [in LHS]execD_pow_real/=. -rewrite [in LHS](@execD_bin _ _ binop_mult)/=. -rewrite [in LHS](@execD_bin _ _ binop_minus)/=. -rewrite [in LHS]execD_real/=. -rewrite [in LHS](@execD_bin _ _ binop_mult)/=. -rewrite [in LHS]execD_pow/=. -rewrite [in LHS](@execD_bin _ _ binop_minus)/=. -rewrite [in LHS]exp_var'E/= (execD_var_erefl "y0")/=. -rewrite [in LHS]exp_var'E/= (execD_var_erefl "x")/=. -rewrite [in LHS]execD_real/=. -rewrite [in LHS]execD_real/=. -rewrite [in LHS]execP_sample. -rewrite [in LHS]execD_normal/=. -rewrite [in LHS]exp_var'E/= (execD_var_erefl "x")/=. -rewrite [in LHS]execP_return/=. -rewrite [in LHS]exp_var'E/= (execD_var_erefl "z")/=. -(* rhs *) -rewrite [in RHS]execP_letin. -rewrite [in RHS]execP_score. -rewrite [in RHS]execD_pow_real/=. -rewrite [in RHS](@execD_bin _ _ binop_mult)/=. -rewrite [in RHS](@execD_bin _ _ binop_minus)/=. -rewrite [in RHS]execD_real/=. -rewrite [in RHS](@execD_bin _ _ binop_mult)/=. -rewrite [in RHS]execD_pow/=. -rewrite [in RHS]exp_var'E/= (execD_var_erefl "y0")/=. -rewrite [in RHS]execD_real/=. -rewrite [in RHS]execD_real/=. -rewrite [in RHS]execP_letin. -rewrite [in RHS]execP_sample. -rewrite [in RHS]execD_normal/=. -rewrite [in RHS](@execD_bin _ _ binop_mult)/=. -rewrite [in RHS]exp_var'E/= (execD_var_erefl "y0")/=. -rewrite [in RHS]execD_real/=. -rewrite [in RHS]execP_letin. -rewrite [in RHS]execP_sample. -rewrite [in RHS]execD_normal/=. -rewrite [in RHS]exp_var'E/= (execD_var_erefl "x")/=. -rewrite [in RHS]execP_return. -rewrite [in RHS]exp_var'E/= (execD_var_erefl "z")/=. -(* lhs *) -rewrite [in LHS]letin'E/=. -under [in LHS]eq_integral. - move=> x _. - rewrite letin'E/=. - under eq_integral. - move=> u _. - rewrite letin'E/=. - rewrite integral_normal_prob_dirac; last first. - admit. (* ? *) - over. - rewrite /=. - rewrite ge0_integral_mscale/=; first last. - - by move=> ? _. - - by []. - - exact: measurableT. - rewrite integral_dirac; first last. - - by []. - - exact: measurableT. - rewrite diracT mul1e. - rewrite sub0r. - rewrite -expRM mul1r. - over. -rewrite /=. -rewrite [in LHS]integral_normal_prob; first last. -- apply: integrableMr => //. - admit. - exists (Num.sqrt 2 * pi)^-1; split; first exact: num_real. - move=> x x_ge_Vsq2pi. - move=> p _. - rewrite /= normr_id. - have /normr_idP -> : 0 <= expR (- ((y.1 - p) ^+ 2%R / 2) / (Num.sqrt 2 * pi)). -red. - admit. -have := (@integrableZl _ _ _ (normal_prob 0 1) _ (@measurableT _ R) _ (fun x => normal_prob x 1 V)). -(* ? *) admit. -- apply: emeasurable_funM; last first. -(* TODO: Define instance of normal_prob is kernel *) - (* apply: measurable_kernel. *) -rewrite /=. - under [X in _ _ X]eq_fun. - move=> x. - rewrite -(@fineK _ (normal_prob _ _ _)); last first. - admit. - over. - apply/measurable_EFinP. - apply: continuous_measurable_fun. - apply: continuous_normal_prob. -have := (@continuous_measurable_fun _ (fun x => normal_prob x 1 V)). - have := measurable_normal_prob2. - apply/measurability - admit. - apply/measurable_EFinP. - admit. -- exact: measurableT. -- exact: ltr01. -(* rhs *) -rewrite [in RHS]letin'E/=. -under [in RHS]eq_integral. - move=> u _. - rewrite letin'E/=. - under eq_integral. - move=> x _. - rewrite letin'E/=. - rewrite integral_normal_prob_dirac. - over. - over. -rewrite ge0_integral_mscale; first last. -- move=> ? _. - exact: integral_ge0. -- exact: measurable_cst. -- exact: measurableT. -rewrite integral_dirac/=; first last. -- exact: measurable_cst. -- exact: measurableT. -rewrite diracT mul1e. -rewrite sub0r -expRM mul1r. -rewrite [in RHS]integral_normal_prob; first last. -- (* ok *) admit. -- (* ? *)admit. (* TODO1: ish *) -- exact: measurableT. -rewrite -ge0_integralZl; first last. -- by []. -- move=> ? _. - apply: mule_ge0 => //. - rewrite lee_fin. - exact: normal_pdf_ge0. -- apply: emeasurable_funM. - + (* TODO1: ish *) admit. - + (* ? *)admit. -- exact: measurableT. -(* eq_integral *) -apply: eq_integral. -move=> x _. -rewrite [in LHS]muleAC. -rewrite [in RHS](muleC (normal_prob1 x V)) muleA. -congr *%E. -rewrite [in LHS]ger0_norm; last first. -- (* ok *)admit. -rewrite [in RHS]ger0_norm; last first. -- (* ok *)admit. -rewrite /normal_pdf mul1r subr0 divr1. -(* ? *) (* TODO2: ish *) -rewrite -!EFinM. -congr EFin. -rewrite mulrA. -rewrite mulrAC. -rewrite -expRD. -rewrite [RHS]mulrA. -rewrite [RHS]mulrAC. -rewrite -expRD. -congr *%R. -- congr expR. - rewrite [in RHS]expr_div_n. - rewrite -[in RHS](@sqrtrV _ 2)//. - rewrite sqr_sqrtr//. - (* ? *) - admit. -admit. -Admitted. -*) - -(* ref : https://homes.luddy.indiana.edu/ccshan/rational/equational-handout.pdf - * p.2, equation (9) - * sum independent random variables that are normally distributed - *) -(* -Lemma helloRight1_to_2 : execD helloRight1 = execD helloRight2. -Proof. -apply: congr_normalize => y V. -(* TODO: split rewriting on LHS and RHS *) -(* lhs *) -rewrite [in LHS]execP_letin. -rewrite !execP_score. -rewrite (@execD_bin _ _ binop_mult)/=. -rewrite execD_pow_real/=. -rewrite (@execD_bin _ _ binop_minus)/=. -rewrite execD_real. -rewrite (@execD_bin _ _ binop_mult)/=. -rewrite execD_pow/=. -rewrite exp_var'E/= (execD_var_erefl "y0")/=. -rewrite !execD_real/=. -rewrite execP_letin. -rewrite execP_sample. -rewrite execD_normal/=. -rewrite !(@execD_bin _ _ binop_mult)/=. -rewrite exp_var'E/= (execD_var_erefl "y0")/= !execD_real/=. -rewrite execP_letin. -rewrite execP_sample. -rewrite execD_normal/=. -rewrite exp_var'E/= (execD_var_erefl "x")/=. -rewrite execP_return/=. -rewrite exp_var'E/= (execD_var_erefl "z")/=. -(* rhs *) -rewrite execP_letin. -rewrite execP_score/=. -rewrite (@execD_bin _ _ binop_mult)/=. -rewrite execD_real/=. -rewrite execD_pow_real/=. -rewrite (@execD_bin _ _ binop_minus)/= execD_real/=. -rewrite (@execD_bin _ _ binop_mult)/= execD_real/=. -rewrite exp_var'E/=. -rewrite execD_pow/=. -rewrite execP_sample. -rewrite execD_normal/=. -rewrite (@execD_bin _ _ binop_mult)/= execD_real/=. -rewrite exp_var'E/= 2!(execD_var_erefl "y0")/=. -(* lhs *) -rewrite letin'E/=. -under eq_integral. - move=> x _. - rewrite letin'E/=. - under eq_integral. - move=> z _. - rewrite letin'E/=. - rewrite integral_normal_prob_dirac//; . - admit. - over. - over. -rewrite ge0_integral_mscale//=; last first. - move=> ? _; exact: integral_ge0. -rewrite integral_dirac//. -rewrite diracT mul1e. -(* rhs *) -rewrite letin'E/=. -rewrite ge0_integral_mscale//=. -rewrite integral_dirac//= diracT mul1e. -Abort. -*) - -End helloRight_proof. - -(* TODO: move *) -Section exponential_pdf. -Context {R : realType}. -Notation mu := lebesgue_measure. -Variable (mean : R). -Hypothesis mean0 : (0 < mean)%R. - -Definition exponential_pdf' (x : R) := (mean^-1 * expR (- x / mean))%R. -Definition exponential_pdf := exponential_pdf' \_ `[0%R, +oo[. - -Lemma exponential_pdf_ge0 (x : R) : (0 <= exponential_pdf x)%R. -Proof. -apply: restrict_ge0 => {}x _. -apply: mulr_ge0; last exact: expR_ge0. -by rewrite invr_ge0 ltW. -Qed. - -End exponential_pdf. - -Definition exponential {R : realType} (m : R) - of \int[@lebesgue_measure R]_x (exponential_pdf m x)%:E = 1%E : set _ -> \bar R := - fun V => (\int[lebesgue_measure]_(x in V) (exponential_pdf m x)%:E)%E. - -Section exponential. -Context {R : realType}. -Local Open Scope ring_scope. - -Notation mu := lebesgue_measure. -Variable (mean : R). -Hypothesis mean0 : (0 < mean)%R. - -Hypothesis integrable_exponential : forall (m : R), (0 < m)%R -> - mu.-integrable [set: _] (EFin \o (exponential_pdf m)). - -Hypothesis integral_exponential_pdf : forall (m : R), (0 < m)%R -> - (\int[mu]_x (exponential_pdf m x)%:E = 1)%E. - -Local Notation exponential := (exponential (integral_exponential_pdf mean0)). - -Let exponential0 : exponential set0 = 0%E. -Proof. by rewrite /exponential integral_set0. Qed. - -Let exponential_ge0 A : (0 <= exponential A)%E. -Proof. -by rewrite /exponential integral_ge0//= => x _; rewrite lee_fin exponential_pdf_ge0. -Qed. - -Let exponential_sigma_additive : semi_sigma_additive exponential. -Proof. -Admitted. - -HB.instance Definition _ := isMeasure.Build _ _ _ - exponential exponential0 exponential_ge0 exponential_sigma_additive. - -Let exponential_setT : exponential [set: _] = 1%E. -Proof. by rewrite /exponential integral_exponential_pdf. Qed. - -HB.instance Definition _ := @Measure_isProbability.Build _ _ R exponential exponential_setT. - -End exponential. diff --git a/theories/lang_syntax_noisy.v b/theories/lang_syntax_noisy.v new file mode 100644 index 0000000000..4236fcd138 --- /dev/null +++ b/theories/lang_syntax_noisy.v @@ -0,0 +1,819 @@ +From Coq Require Import String. +From HB Require Import structures. +From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. +From mathcomp.classical Require Import mathcomp_extra boolp. +From mathcomp Require Import ring lra. +From mathcomp Require Import classical_sets functions cardinality fsbigop. +From mathcomp Require Import interval_inference reals ereal topology normedtype. +From mathcomp Require Import sequences esum measure lebesgue_measure numfun exp. +From mathcomp Require Import trigo realfun charge lebesgue_integral kernel. +From mathcomp Require Import probability prob_lang. +From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. + +(**md**************************************************************************) +(* # Observing a noisy draw from a normal distribution *) +(* *) +(* Formalization of the HelloRight example (Sect. 2.3 of [Shan, POPL 2018]). *) +(* *) +(* ref: *) +(* - Chung-chieh Shan, Equational reasoning for probabilistic programming, *) +(* POPL TutorialFest 2018 *) +(* https://homes.luddy.indiana.edu/ccshan/rational/equational-handout.pdf *) +(* - Praveen Narayanan. Verifiable and reusable conditioning. PhD thesis, *) +(* Indiana University, 2019. *) +(* *) +(* ``` *) +(* noisy0 == distribution of the next noisy measurement of a normally *) +(* distributed quantity *) +(* ``` *) +(* *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import Order.TTheory GRing.Theory Num.Def Num.ExtraDef Num.Theory. +Import numFieldTopology.Exports. + +Local Open Scope classical_set_scope. +Local Open Scope ring_scope. +Local Open Scope ereal_scope. +Local Open Scope string_scope. + +(* TODO: PR *) +Section ge0_fin_measurable_probability_integrable. +Context d {T : measurableType d} {R : realType} {p : probability T R} + {f : T -> \bar R}. + +Lemma ge0_fin_measurable_probability_integrable : + (forall x, 0 <= f x) -> + (exists M : R, forall x, f x <= M%:E) -> + measurable_fun setT f -> + p.-integrable setT f. +Proof. +move=> f0 [M fleM] mf. +apply/integrableP; split => //. +apply (@le_lt_trans _ _ (\int[p]_x M%:E)). + apply: ge0_le_integral => //=. + exact: measurableT_comp. + move=> t _. + exact: (@le_trans _ _ (f t)). + move=> t _. + by rewrite gee0_abs. +by rewrite integral_cst// probability_setT mule1 ltry. +Qed. + +End ge0_fin_measurable_probability_integrable. + +(* TODO: PR *) +Section pkernel_probability_integrable. +Context d d' {T : measurableType d} {T' : measurableType d'} {R : realType} + {p : probability T R} {f : R.-pker T ~> T'}. + +Lemma pkernel_probability_integrable V : measurable V -> + p.-integrable setT (fun x => f x V). +Proof. +move=> mV. +apply: ge0_fin_measurable_probability_integrable => //. + exists 1%R. + move=> x. + apply: (@le_trans _ _ (f x setT)). + by apply: le_measure; rewrite ?inE. + by rewrite prob_kernel. +exact: measurable_kernel. +Qed. + +End pkernel_probability_integrable. + +(* TODO: move to probability.v *) +Section normal_prob_properties. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +Local Open Scope charge_scope. + +Lemma normal_pdf_uniq_ae (m s : R) (s0 : (0 < s)%R) : + ae_eq mu setT + ('d ((charge_of_finite_measure (@normal_prob R m s))) '/d mu) + (EFin \o (@normal_pdf R m s)). +Proof. +apply: integral_ae_eq => //. + apply: Radon_Nikodym_integrable => /=. + exact: normal_prob_dominates. + apply/measurable_EFinP. + exact: measurable_normal_pdf. +move=> /= E _ mE. +by rewrite -Radon_Nikodym_integral//=; apply: normal_prob_dominates. +Qed. + +End normal_prob_properties. + +(* TODO: move *) +Section normal_prob_lemmas. +Context {R: realType}. +Local Notation mu := lebesgue_measure. + +Lemma emeasurable_normal_prob (s : R) U : s != 0%R -> measurable U -> + measurable_fun setT (fun x => normal_prob x s U). +Proof. +move=> s0 mU. +under [X in _ _ X]eq_fun. + move=> x. + rewrite -(@fineK _ (_ x _ _)); last first. + rewrite ge0_fin_numE//. + apply: (@le_lt_trans _ _ (normal_prob x s setT)). + by apply: le_measure; rewrite ?inE. + by rewrite probability_setT ltry. + over. +apply/measurable_EFinP. +apply: (continuous_measurable_fun). +exact: normal_prob_continuous. +Qed. + +(* TODO: s != 0 *) +Lemma integral_normal_prob (m s : R) (s0 : (0 < s)%R) f U : + measurable U -> + (normal_prob m s).-integrable U f -> + \int[@normal_prob _ m s]_(x in U) f x = + \int[mu]_(x in U) (f x * (normal_pdf m s x)%:E). +Proof. +move=> mU intf. +move/integrableP : (intf) => [mf intf_lty]. +rewrite -(Radon_Nikodym_change_of_variables (normal_prob_dominates m s))//=. +apply: ae_eq_integral => //=. + apply: emeasurable_funM => //. + apply: measurable_funTS. + have : charge_of_finite_measure (normal_prob m s) `<< mu. + exact: normal_prob_dominates m s. + move/Radon_Nikodym_integrable. + by move/integrableP => []. + apply: emeasurable_funM => //. + apply/measurable_EFinP. + apply: measurable_funTS. + exact: measurable_normal_pdf. +apply: ae_eqe_mul2l. +apply: (ae_eq_subset (@subsetT _ U)). +exact: (normal_pdf_uniq_ae m s0). +Qed. + +Lemma normal_prob_integrable_dirac (m s : R) (V : set R): measurable V -> + (normal_prob m s).-integrable setT (fun x => \d_x V). +Proof. +move=> mV. +apply/integrableP; split; first exact: measurable_fun_dirac. +rewrite -(setUv V). +rewrite ge0_integral_setU => //; first last. + exact/disj_setPCl. + rewrite setUv. + apply: measurableT_comp => //. + exact: measurable_fun_dirac. + exact: measurableC. +under eq_integral. + move=> x Vx. + rewrite diracE Vx/= normr1. + over. +under [X in _ + X < _]eq_integral. + move=> x. + rewrite inE/= => nVx. + have {}nVx := (memNset nVx). + rewrite indicE nVx/= normr0. + over. +rewrite !integral_cst//=; last exact: measurableC. +rewrite mul1e mul0e adde0. +apply: (le_lt_trans (probability_le1 (normal_prob m s) mV)). +exact: ltey. +Qed. + +Lemma integral_normal_prob_dirac (s : R) (m : R) V : + (0 < s)%R -> + measurable V -> + \int[normal_prob m s]_x0 (\d_x0 V) = normal_prob m s V. +Proof. +move=> s0 mV. +rewrite integral_normal_prob => //; last first. + exact: (normal_prob_integrable_dirac m s). +under eq_integral do rewrite diracE. +rewrite /normal_prob. +rewrite [in RHS]integral_mkcond. +under [in RHS]eq_integral do rewrite patchE. +rewrite /=. +apply: eq_integral => x _. +by case: ifP => xV/=; rewrite ?mul1e ?mul0e. +Qed. + +End normal_prob_lemmas. + +Section noisy_programs. +Local Open Scope lang_scope. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +Definition exp_normal1 {R g} := @exp_normal R g _ (@oner_neq0 R). + +(* NB: exp_powR level setting is mistaken? *) +(* ((_ `^ _) * _) cannot write as (_ `^ _ * _) *) +Definition noisy0' : @exp R P [:: ("y0", Real)] Real := + [let "x" := Sample {exp_normal1 [{0}:R]} in + let "_" := Score ({expR 1} `^ + ({0}:R - (#{"y0"} - #{"x"}) ^+ {2%R} * {2^-1}:R)) + * {(Num.sqrt (2 * pi))^-1}:R in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. + +Definition noisy0 : @exp R _ [:: ("y0", Real)] _ := [Normalize {noisy0'}]. + +(* other programs from Sect. 2.3 of [Shan, POPL 2018], + nothing proved about them yet, just for the sake of completeness *) +Definition guard_real {g} str (r : R) : + @exp R P [:: (str, _) ; g] _ := + [if #{str} ==R {r}:R then return TT else Score {0}:R]. + +Definition helloWrong (y0 : R) : @exp R _ [::] _ := + [Normalize + let "x" := Sample {exp_normal1 (exp_real 0)} in + let "y" := Sample {exp_normal1 [#{"x"}]} in + let "_" := {guard_real "y" y0} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. + +Definition helloJoint : @exp R _ [::] _ := + [Normalize + let "x" := Sample {exp_normal1 (exp_real 0)} in + let "y" := Sample {exp_normal1 [#{"x"}]} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return (#{"y"}, #{"z"})]. + +End noisy_programs. + +(* The following section contains the mathematical facts that are used + to verify the noisy program. They are proved beforehand to + optimize the time spent by the Qed command of Rocq. *) +Section noisy_subproofs. +Local Open Scope lang_scope. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +Local Definition int_normal_noisy0 + (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := + \int[normal_prob 0 1]_x (fun z => + `|expR (- ((y.1 - z) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * + normal_prob z 1 V) x. + +Local Definition int_mu_noisy0 + (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := + \int[mu]_x + (`|expR (- ((y.1 - x) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * + normal_prob x 1 V * + (normal_pdf 0 1 x)%:E). + +Local Definition int_normal_noisy1 + (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := + \int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x + ((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))))%:num%:E * + normal_prob x 1 V)%E. + +Local Definition int_mu_noisy1 + (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := + \int[mu]_x + (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E * + (normal_prob x 1 V * (normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 x)%:E)). + +Local Definition int_normal_noisy2 + (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := + \int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x normal_prob x 1 V. + +Local Definition int_mu_noisy2 + (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := + normal_prob (y.1 / 2) (Num.sqrt (3 / 2)) V. + +Lemma int_normal_mu_noisy0 y V : measurable V -> + int_normal_noisy0 y V = int_mu_noisy0 y V. +Proof. +move=> mV; rewrite /int_normal_noisy0. +rewrite integral_normal_prob//. +apply: integrableMr => //. + apply: (@measurableT_comp _ _ _ _ _ _ (@normr _ R^o)) => //. + apply: measurable_funM => //. + apply: measurableT_comp => //. + apply: measurableT_comp => //. + apply: measurable_funM => //. + apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => (t ^+ 2%R)%R)) => //. + exact: measurable_funB. + exists (Num.sqrt (2 * pi))^-1%R; split; first exact: num_real. + move=> x x_gt. + move=> p _. + rewrite /= normr_id. + have /normr_idP -> : (0 <= expR (- ((y.1 - p) ^+ 2%R / 2)) / (Num.sqrt (2 * pi)))%R. + apply: mulr_ge0; first exact: expR_ge0. + by rewrite invr_ge0 sqrtr_ge0// mulr_ge0// pi_ge0. + apply/ltW. + apply: le_lt_trans x_gt. + rewrite -[leRHS]mul1r ler_pM//. + by rewrite -expR0 ler_expR oppr_le0 mulr_ge0// ?sqr_ge0// expR0 invr_ge0. +apply/integrableP; split. + exact: emeasurable_normal_prob. +apply/abse_integralP => //. + exact: emeasurable_normal_prob. +have/gee0_abs -> : 0 <= \int[normal_prob 0 1]_x normal_prob x 1 V. + exact: integral_ge0. +apply: (@le_lt_trans _ _ (\int[normal_prob 0 1]_x (cst 1%R x)%:E)). + apply: ge0_le_integral => //. + exact: emeasurable_normal_prob. + apply/measurable_EFinP. + exact: measurable_cst. + move=> x _. + exact: probability_le1. +by rewrite /= integral_cst// mul1e probability_setT ltry. +Qed. + +Lemma int_normal_mu_noisy1 y V : measurable V -> + int_normal_noisy1 y V = int_mu_noisy1 y V. +Proof. +move=> mV; rewrite /int_normal_noisy1 integral_normal_prob//. + by under eq_integral do rewrite -muleA. +apply/integrableP; split. + apply: measurable_funeM. + exact: emeasurable_normal_prob. +apply: (@le_lt_trans _ _ (\int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x + ((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))))%:num%:E + * (cst 1%R x)%:E))). +apply: ge0_le_integral; [by []|by []| | | |]. +- apply: measurableT_comp => //. + apply: measurable_funeM. + exact: emeasurable_normal_prob. +- by move=> x _; rewrite /= mule1. +- rewrite /= mule1. + exact/measurable_EFinP. +- move=> x _/=. + rewrite gee0_abs; last exact: mule_ge0. + by rewrite lee_pmul// probability_le1. +- rewrite integralZr//; last exact: finite_measure_integrable_cst. + by rewrite integral_cst// mule1 probability_setT ltry. +Qed. + +(* TODO: generalize and move outside *) +Section conjugate_normal_property. + +Lemma conjugate_normal1 (m1 m2 s1 s2 : R) V : measurable V -> + (0 < s1)%R -> s2 != 0%R -> + \int[normal_prob m1 s1]_x normal_prob (m2 + x) s2 V = + \int[mu]_(y in V) \int[mu]_x (normal_pdf (m2 + x) s2 y * normal_pdf m1 s1 x)%:E. +Proof. +move=> mV s10 s20; rewrite integral_normal_prob//; last first. + apply: ge0_fin_measurable_probability_integrable => //=. + by exists 1%R => ?; exact: probability_le1. + apply: (@measurableT_comp _ _ _ _ _ _ + (fun x => normal_prob x s2 V) _ (fun x => m2 + x)). + exact: emeasurable_normal_prob. + exact: measurable_funD. +transitivity (\int[mu]_x + \int[mu]_y ((normal_pdf (m2 + x) s2 y * + normal_pdf m1 s1 x)%:E * (\1_V y)%:E)). + apply: eq_integral => y _. + rewrite /normal_prob. + rewrite -integralZr//; last first. + apply: (integrableS measurableT) => //. + exact: integrable_normal_pdf. + transitivity (\int[mu]_(x in V) (normal_pdf (m2 + y) s2 x * + normal_pdf m1 s1 y)%:E). + apply: eq_integral => z _. + by rewrite -EFinM. + rewrite integral_mkcond. + by rewrite epatch_indic. +rewrite (@fubini_tonelli _ _ _ _ _ mu mu (EFin \o + ((fun xz : R * R => (normal_pdf (m2 + xz.1) s2 xz.2 * + normal_pdf m1 s1 xz.1)%R + * \1_V xz.2)%R)))/=; last 2 first. + apply/measurable_EFinP; apply: measurable_funM => /=; last first. + apply: measurable_indic. + rewrite -[X in measurable X]setTX. + exact: measurableX. + apply: measurable_funM. + rewrite [X in measurable_fun _ X](_ : _ = (fun x0 => + normal_pdf0 0 s2 (x0.2 - (m2 + x0.1)%E))) /=; last first. + apply/funext=> x0. + rewrite /normal_pdf0. + rewrite normal_pdfE//. + by rewrite normal_fun_center. + apply: measurableT_comp. + exact: measurable_normal_pdf0. + rewrite /=. + under eq_fun do rewrite opprD. + apply: measurable_funD => //=. + exact: measurable_funB. + by apply: measurableT_comp => //; exact: measurable_normal_pdf. + move=> x/=. + by rewrite lee_fin !mulr_ge0 ?normal_pdf_ge0. +transitivity (\int[mu]_x + \int[mu]_x0 + ((fun x0 => + (normal_pdf (m2 + x0) s2 x * normal_pdf m1 s1 x0)%:E) \_ (fun=> V x)) x0). + apply: eq_integral => x0 _. + rewrite /=. + under eq_integral do rewrite EFinM. + by rewrite -epatch_indic. +rewrite [RHS]integral_mkcond/=. +apply: eq_integral => /= x0 _. +rewrite patchE. +case: ifPn => xV. + apply: eq_integral => z _/=. + rewrite patchE. + by rewrite ifT. +apply: integral0_eq => /= z _. +rewrite patchE ifF//; apply/negbTE. +rewrite notin_setE/=. +move/negP in xV. +by move/mem_set. +Qed. + +Lemma conjugate_normal2 (y m1 m2 s1 s2 : R) : s1 != 0%R -> (0 < s2)%R -> + \int[mu]_x (normal_pdf (m1 + x)%E s1 y * normal_pdf m2 s2 x)%:E = + (normal_peak s1 * normal_peak s2)%:E * + \int[mu]_z (normal_fun (m1 + z) s1 y * normal_fun m2 s2 z)%:E. +Proof. +move=> s10 s20. +rewrite -ge0_integralZl//=; last 3 first. +- apply/measurable_EFinP => //=; apply: measurable_funM => //=. + + apply: measurableT_comp => //=; apply: measurable_funM => //=. + apply/measurableT_comp => //; apply: measurable_funX. + under eq_fun do rewrite opprD. + apply: measurable_funD => //=. + exact: measurable_funB. + + apply: measurableT_comp => //=. + apply: measurable_funM => //. + apply: measurableT_comp => //. + by apply: measurable_funX; exact: measurable_funD. +- by move=> z _; rewrite lee_fin mulr_ge0// expR_ge0. +- by rewrite lee_fin mulr_ge0// ?normal_peak_ge0. +apply: eq_integral => /= z _. +rewrite /normal_pdf (negbTE s10) gt_eqF//. +rewrite /normal_pdf0. +rewrite mulrACA. +rewrite /normal_fun. +by congr EFin. +Qed. + +Lemma normal_peak1 : normal_peak 1 = (Num.sqrt (pi *+ 2))^-1%R :> R. +Proof. by rewrite /normal_peak expr1n mul1r. Qed. + +Lemma conjugate_normal (m2 y : R) V : measurable V -> + \int[normal_prob y (Num.sqrt 2)^-1]_x normal_prob (m2 + x) 1 V = + normal_prob (y + m2) (Num.sqrt (3 / 2)) V. +Proof. +move=> mV. +rewrite conjugate_normal1//; apply: eq_integral => x1 _. +clear V mV. +rewrite conjugate_normal2//. +rewrite normal_pdfE// /normal_pdf0. +transitivity ((pi * Num.sqrt 2)^-1%:E * + \int[mu]_x0 (expR (- (Num.sqrt (3 / 2) ^+ 2)%R * (x0 - (x1 + 2 * y - m2) / 3) ^+ 2 - (x1 - (y + m2)) ^+ 2 / ((Num.sqrt (3 / 2) ^+ 2) *+ 2)))%:E). + congr *%E. + rewrite normal_peak1. + rewrite /normal_peak. + congr EFin. + rewrite exprVn sqr_sqrtr//. + rewrite -mulr_natl mulrA divff// mul1r. + rewrite -(mulr_natl pi) sqrtrM// invfM. + rewrite invfM. + rewrite -mulrA mulrC. + rewrite -(invfM (Num.sqrt pi)) -expr2. + by rewrite sqr_sqrtr// pi_ge0. + apply: eq_integral. + move=> z _. + rewrite -expRD. + congr EFin. + congr expR. + rewrite !sqrrD. + rewrite exprVn sqr_sqrtr//. + rewrite sqr_sqrtr//. + lra. +(* gauss integral *) +under eq_integral do rewrite expRD EFinM. +rewrite ge0_integralZr//=; last first. + apply/measurable_EFinP. + apply: measurableT_comp => //. + apply: measurable_funM => //. + apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)%R) => //. + exact: measurable_funD. +rewrite /=. +rewrite /normal_peak /normal_fun. +rewrite [in RHS]sqr_sqrtr//. +rewrite -[in RHS]mulr_natr [in RHS]mulrAC divfK//. +rewrite mulNr EFinM muleA. +congr *%E; last first. + by rewrite sqr_sqrtr//. +rewrite [X in _ * X = _](_ : _ = (Num.sqrt ((1 / 3) * pi *+ 2))%:E * + \int[mu]_z (normal_pdf ((x1 + 2 * y - m2) / 3) (Num.sqrt (1 / 3)) z)%:E); last first. + rewrite -ge0_integralZl//=; last 2 first. + by apply/measurable_EFinP; exact: measurable_normal_pdf. + by move=> /= z _; rewrite lee_fin normal_pdf_ge0. + apply: eq_integral => /= z _. + rewrite /normal_pdf gt_eqF// /normal_pdf0 /normal_peak /normal_fun. + rewrite (@sqr_sqrtr _ (1 / 3))// -EFinM; congr EFin. + rewrite [RHS]mulrA -[LHS]mul1r; congr (_ * expR _)%R. + by rewrite divff// gt_eqF// sqrtr_gt0 pmulrn_rgt0// mulr_gt0// pi_gt0. + rewrite -(mulr_natl (1 / 3)) mul1r. + rewrite sqr_sqrtr//. + by rewrite !mulNr (mulrC (3 / 2)%R) invfM invrK (mulrC 2^-1%R). +rewrite integral_normal_pdf. +rewrite mule1 -EFinM mul1r; congr EFin. +rewrite -(mulr_natr (_ * pi)) sqrtrM ?mulr_ge0 ?pi_ge0//. +rewrite invfM mulrCA -mulrA mulVf ?mulr1 ?gt_eqF//. +rewrite !sqrtrM// invfM sqrtrV// -mulrA; congr *%R. +rewrite -[X in (_ / X)%R]sqr_sqrtr ?pi_ge0//. +by rewrite expr2 invfM mulrA divff ?div1r// gt_eqF// sqrtr_gt0 pi_gt0. +Qed. + +End conjugate_normal_property. + +Lemma int_normal_mu_noisy2 y V : measurable V -> + int_normal_noisy2 y V = int_mu_noisy2 y V. +Proof. +move=> mV. +have := @conjugate_normal 0 (y.1 / 2) _ mV. +under eq_integral do rewrite add0r. +rewrite /int_normal_noisy2. +rewrite addr0. +exact. +Qed. + +End noisy_subproofs. + +(* this section contains the verification of the noisy program per se *) +Section noisy_verification. +Local Open Scope lang_scope. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +(* definition of intermediate programs *) +Definition neq0Vsqrt2 : ((Num.sqrt 2)^-1 != 0 :> R)%R. +Proof. exact: lt0r_neq0. Qed. +Definition exp_normal_Vsqrt2 {g} := @exp_normal R g _ neq0Vsqrt2. + +Definition noisy1' : @exp R _ [:: ("y0", Real)] Real := + [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2%R} * {4^-1}:R)) * + {(Num.sqrt (4 * pi))^-1}:R in + let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R ]} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. + +Definition noisy1 : @exp R _ [:: ("y0", Real)] _ := [Normalize {noisy1'}]. + +Definition neq0sqrt32 : (Num.sqrt (3 / 2) != 0 :> R)%R. +Proof. exact: lt0r_neq0. Qed. +Definition exp_normal_sqrt32 {g} := @exp_normal R g _ neq0sqrt32. + +Definition noisy2' : @exp R _ [:: ("y0", Real)] _ := + [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2%R} * {4^-1}:R)) * + {(Num.sqrt (4 * pi))^-1}:R in + Sample {exp_normal_sqrt32 [#{"y0"} * {2^-1}:R]}]. + +Definition noisy2 : @exp R _ [:: ("y0", Real)] _ := [Normalize {noisy2'}]. + +(* from (7) to (9) in [Shan, POPL 2018] *) +Lemma execP_noisy'12 y V : measurable V -> + @execP R [:: ("_", Unit); ("y0", Real)] _ + [let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R ]} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}] y V = + @execP R [:: ("_", Unit); ("y0", Real)] _ + [Sample {exp_normal_sqrt32 [#{"y0"} * {2^-1}:R]}] y V. +Proof. +move=> mV. +(* execute syntax *) +rewrite !execP_letin. +rewrite !execP_sample !execD_normal/=. +rewrite (@execD_bin _ _ binop_mult) execD_real/=. +rewrite execP_return. +rewrite !exp_var'E (execD_var_erefl "y0") (execD_var_erefl "x") (execD_var_erefl "z")/=. +rewrite !letin'E/=. +under eq_integral do rewrite letin'E/=. +rewrite /=. +(* prove semantics *) +under eq_integral do rewrite integral_normal_prob_dirac//=. +exact: (int_normal_mu_noisy2 _ mV). +Qed. + +(* semantics of noisy0 (turned into a local definition as an attempt to optimize Qed) *) +Local Definition executed_noisy0' := + (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + letin' + (sample (fun=> normal_prob 0 1) + (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) (kr 0))) + (letin' + (score + (measurable_funM + (measurableT_comp (measurable_powRr (expR 1)) + (measurable_funB (kr 0) + (measurable_funM + (measurable_funX 2%R + (measurable_funB (measurable_acc_typ [:: Real; Real] 1) + (measurable_acc_typ [:: Real; Real] 0))) + (kr 2^-1)))) (kr (Num.sqrt (2 * pi))^-1))) + (letin' + (sample + (fun + x : unit * + (g_sigma_algebraType R.-ocitv.-measurable * + (g_sigma_algebraType R.-ocitv.-measurable * unit)) => + normal_prob x.2.1 1) + (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) + (measurable_acc_typ [:: Unit; Real; Real] 1))) + (ret (measurable_acc_typ [:: Real; Unit; Real; Real] 0)))) y V : \bar R). + +(* semantics of noisy1 (turned into a local definition as a attempt to optimized Qed) *) +Local Definition executed_noisy1' := + (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + letin' + (score + (measurable_funM + (measurableT_comp (measurable_powRr (expR 1)) + (measurable_funB (kr 0) + (measurable_funM + (measurable_funX 2%R (measurable_acc_typ [:: Real] 0)) (kr 4^-1)))) + (kr (Num.sqrt (4 * pi))^-1))) + (letin' + (sample + (fun x : unit * (g_sigma_algebraType R.-ocitv.-measurable * unit) => + normal_prob (x.2.1 / 2) (Num.sqrt 2)^-1) + (measurableT_comp (measurable_normal_prob2 neq0Vsqrt2) + (measurable_funM (measurable_acc_typ [:: Unit; Real] 1) (kr 2^-1)))) + (letin' + (sample + (fun + x : g_sigma_algebraType R.-ocitv.-measurable * + (unit * (g_sigma_algebraType R.-ocitv.-measurable * unit)) => + normal_prob x.1 1) + (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) + (measurable_acc_typ [:: Real; Unit; Real] 0))) + (ret (measurable_acc_typ [:: Real; Real; Unit; Real] 0)))) y V : \bar R). + +Lemma execP_noisy0'E y V : execP noisy0' y V = executed_noisy0' y V. +Proof. +rewrite !execP_letin. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite execD_real/=. +rewrite execP_score. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow_real/=. +rewrite (@execD_bin _ _ binop_minus)/=. +rewrite execD_real/=. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow/=. +rewrite (@execD_bin _ _ binop_minus)/=. +rewrite exp_var'E/= (execD_var_erefl "y0")/=. +rewrite exp_var'E/= (execD_var_erefl "x")/=. +rewrite execD_real/=. +rewrite execD_real/=. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite exp_var'E/= (execD_var_erefl "x")/=. +rewrite execP_return/=. +rewrite exp_var'E/= (execD_var_erefl "z")/=. +done. +Qed. + +Lemma executed_noisy0'_semantics y V : measurable V -> + executed_noisy0' y V = int_normal_noisy0 y V. +Proof. +move=> mV. +rewrite /executed_noisy0'. +rewrite [in LHS]letin'E/=. +under [in LHS]eq_integral. + move=> x _. + rewrite letin'E/=. + under eq_integral. + move=> u _. + rewrite letin'E/=. + rewrite integral_normal_prob_dirac//. + over. + rewrite /=. + rewrite ge0_integral_mscale/=; first last. + - by move=> ? _. + - by []. + - exact: measurableT. + rewrite integral_dirac; first last. + - by []. + - exact: measurableT. + rewrite diracT mul1e. + rewrite sub0r. + rewrite -expRM mul1r. + over. +by rewrite /=. +Qed. + +Lemma execP_noisy1'E y V : execP noisy1' y V = executed_noisy1' y V. +Proof. +rewrite execP_letin. +rewrite execP_score. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow_real/=. +rewrite (@execD_bin _ _ binop_minus)/=. +rewrite execD_real/=. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow/=. +rewrite exp_var'E/= (execD_var_erefl "y0")/=. +rewrite execD_real/=. +rewrite execD_real/=. +rewrite execP_letin. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite exp_var'E/= (execD_var_erefl "y0")/=. +rewrite execD_real/=. +rewrite execP_letin. +rewrite execP_sample. +rewrite execD_normal/=. +rewrite exp_var'E/= (execD_var_erefl "x")/=. +rewrite execP_return. +rewrite exp_var'E/= (execD_var_erefl "z")/=. +done. +Qed. + +Lemma executed_noisy1'_semantics y V : measurable V -> + executed_noisy1' y V = int_normal_noisy1 y V. +Proof. +move=> mV; rewrite /executed_noisy1'. +rewrite letin'E/=. +under eq_integral. + move=> u _. + rewrite letin'E/=. + under eq_integral. + move=> x _. + rewrite letin'E/=. + rewrite integral_normal_prob_dirac//. + over. + over. +rewrite ge0_integral_mscale//; first last. + move=> ? _. + exact: integral_ge0. +rewrite integral_dirac//. +rewrite diracT mul1e. +rewrite sub0r -expRM mul1r. +rewrite /=. +rewrite -(@ge0_integralZl _ _ R (normal_prob _ _) _ measurableT _ _ + (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E))//. +exact: emeasurable_normal_prob. +Qed. + +Lemma noisy'01 y V : measurable V -> execP noisy0' y V = execP noisy1' y V. +Proof. +move=> mV. +(* lhs *) +rewrite execP_noisy0'E. +rewrite (executed_noisy0'_semantics _ mV). +(* rhs *) +rewrite execP_noisy1'E. +rewrite (executed_noisy1'_semantics _ mV). +(* semantics *) +rewrite (int_normal_mu_noisy0 _ mV). +rewrite (int_normal_mu_noisy1 _ mV). +(* eq_integral *) +apply: eq_integral. +move=> x _. +rewrite [in LHS]muleAC. +rewrite [in RHS](muleC (normal_prob x 1 V)) muleA. +congr (fun t => t * normal_prob x 1 V)%E. +rewrite [in LHS]ger0_norm; last by rewrite mulr_ge0// expR_ge0. +rewrite [in RHS]ger0_norm; last by rewrite mulr_ge0// expR_ge0. +rewrite -!EFinM; congr EFin. +rewrite !normal_pdfE// /normal_pdf0. +rewrite mulrA. +rewrite mulrAC. +rewrite -(@sqrtrV _ 2)//. +rewrite /normal_peak. +rewrite sqr_sqrtr; last by rewrite invr_ge0. +rewrite /normal_fun. +rewrite subr0. +rewrite sqr_sqrtr; last by rewrite invr_ge0. +rewrite [X in (X / _)%R = _]mulrC mulrA -expRD -mulrA -invfM. +rewrite [RHS]mulrAC [X in _ = (_ * X / _)%R]mulrC mulrA -mulrA -expRD -invfM. +congr *%R. + congr expR. + lra. +congr GRing.inv. +rewrite expr1n mul1r -mulrnAr -(mulr_natl pi) mulrA mulVf// mul1r -expr2. +rewrite sqr_sqrtr; last by rewrite mulr_ge0// pi_ge0. +rewrite !sqrtrM// mulrCA -expr2 sqr_sqrtr; last exact: pi_ge0. +congr *%R. +rewrite -[LHS]gtr0_norm//. +rewrite -(sqrtr_sqr 2%R). +congr Num.sqrt. +lra. +Qed. + +Lemma noisy'12 y V : measurable V -> execP noisy1' y V = execP noisy2' y V. +Proof. +move=> mV. +rewrite /noisy1' /noisy2'. +rewrite execP_letin/= [in RHS]execP_letin/=. +rewrite letin'E [RHS]letin'E. +by under eq_integral do rewrite execP_noisy'12//. +Qed. + +Lemma noisy'02 y V : measurable V -> execP noisy0' y V = execP noisy2' y V. +Proof. by move=> mV; rewrite noisy'01// noisy'12. Qed. + +End noisy_verification. diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v index ce241d96f0..1ae367ceeb 100644 --- a/theories/lang_syntax_table_game.v +++ b/theories/lang_syntax_table_game.v @@ -13,6 +13,11 @@ From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. (**md**************************************************************************) (* # Eddy's table game example *) (* *) +(* Formalization of the Eddy's table game by equational reasoning. See *) +(* Sections 2.1, 3.2, 3.4, 3.5 of [Shan, POPL 2018]. The final statement of *) +(* equivalence is Lemma from_prog0_to_prog5. Intermediate steps are lemmas *) +(* named progij that turn the program progi into the program progj. *) +(* *) (* ref: *) (* - Chung-chieh Shan, Equational reasoning for probabilistic programming, *) (* POPL TutorialFest 2018 *) @@ -20,6 +25,10 @@ From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. (* - Sean R Eddy, What is Bayesian statistics?, Nature Biotechnology 22(9), *) (* 1177--1178 (2004) *) (* *) +(* ``` *) +(* prog0 == Eddy's table game represented as a probabilistic program *) +(* ``` *) +(* *) (******************************************************************************) Set Implicit Arguments. @@ -587,7 +596,7 @@ transitivity (\d_false U + \d_true U - bernoulli_prob (1 / 11) U : \bar R)%E; la rewrite -EFinD /bernoulli_pmf [X in X%:E](_ : _ = 1%R); last first. case: x => //; lra. over. - by rewrite /= dirac_bool. + by rewrite /= dirac_bool. rewrite -int_beta_prob_bernoulli. apply/esym/eqP; rewrite sube_eq//; last first. by rewrite ge0_adde_def// inE; exact: integral_ge0. diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 0bb9b0a6a0..4dac4d1814 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -114,3 +114,65 @@ Definition staton_counting : R.-sfker X ~> _ := (ret macc1of3)). End staton_counting. + +Section exponential_pdf. +Context {R : realType}. +Notation mu := lebesgue_measure. +Variable (mean : R). +Hypothesis mean0 : (0 < mean)%R. + +Definition exponential_pdf' (x : R) := (mean^-1 * expR (- x / mean))%R. +Definition exponential_pdf := exponential_pdf' \_ `[0%R, +oo[. + +Lemma exponential_pdf_ge0 (x : R) : (0 <= exponential_pdf x)%R. +Proof. +apply: restrict_ge0 => {}x _. +apply: mulr_ge0; last exact: expR_ge0. +by rewrite invr_ge0 ltW. +Qed. + +End exponential_pdf. + +Definition exponential {R : realType} (m : R) + of \int[@lebesgue_measure R]_x (exponential_pdf m x)%:E = 1%E : set _ -> \bar R := + fun V => (\int[lebesgue_measure]_(x in V) (exponential_pdf m x)%:E)%E. + +Section exponential. +Context {R : realType}. +Local Open Scope ring_scope. + +Notation mu := lebesgue_measure. +Variable (mean : R). +Hypothesis mean0 : (0 < mean)%R. + +Hypothesis integrable_exponential : forall (m : R), (0 < m)%R -> + mu.-integrable [set: _] (EFin \o (exponential_pdf m)). + +Hypothesis integral_exponential_pdf : forall (m : R), (0 < m)%R -> + (\int[mu]_x (exponential_pdf m x)%:E = 1)%E. + +Local Notation exponential := (exponential (integral_exponential_pdf mean0)). + +Let exponential0 : exponential set0 = 0%E. +Proof. by rewrite /exponential integral_set0. Qed. + +Let exponential_ge0 A : (0 <= exponential A)%E. +Proof. +rewrite /exponential integral_ge0//= => x _. +by rewrite lee_fin exponential_pdf_ge0. +Qed. + +Let exponential_sigma_additive : semi_sigma_additive exponential. +Proof. +Admitted. + +HB.instance Definition _ := isMeasure.Build _ _ _ + exponential exponential0 exponential_ge0 exponential_sigma_additive. + +Let exponential_setT : exponential [set: _] = 1%E. +Proof. by rewrite /exponential integral_exponential_pdf. Qed. + +HB.instance Definition _ := + @Measure_isProbability.Build _ _ R exponential exponential_setT. + +End exponential. From 76e09be74cb2e73fe56074cde9ed569c89bc941b Mon Sep 17 00:00:00 2001 From: IshiguroYoshihiro Date: Wed, 19 Mar 2025 15:14:48 +0900 Subject: [PATCH 39/48] normal_probD --- theories/lang_syntax_noisy.v | 297 ++++++++++++++++++++++++++++++++--- 1 file changed, 275 insertions(+), 22 deletions(-) diff --git a/theories/lang_syntax_noisy.v b/theories/lang_syntax_noisy.v index 4236fcd138..1e3eb73056 100644 --- a/theories/lang_syntax_noisy.v +++ b/theories/lang_syntax_noisy.v @@ -25,6 +25,7 @@ From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. (* ``` *) (* noisy0 == distribution of the next noisy measurement of a normally *) (* distributed quantity *) +(* noisy *) (* ``` *) (* *) (******************************************************************************) @@ -93,7 +94,7 @@ Local Notation mu := lebesgue_measure. Local Open Scope charge_scope. -Lemma normal_pdf_uniq_ae (m s : R) (s0 : (0 < s)%R) : +Lemma normal_pdf_uniq_ae (m s : R) (s0 : (s != 0)%R) : ae_eq mu setT ('d ((charge_of_finite_measure (@normal_prob R m s))) '/d mu) (EFin \o (@normal_pdf R m s)). @@ -131,8 +132,7 @@ apply: (continuous_measurable_fun). exact: normal_prob_continuous. Qed. -(* TODO: s != 0 *) -Lemma integral_normal_prob (m s : R) (s0 : (0 < s)%R) f U : +Lemma integral_normal_prob (m s : R) (s0 : (s != 0)%R) f U : measurable U -> (normal_prob m s).-integrable U f -> \int[@normal_prob _ m s]_(x in U) f x = @@ -186,7 +186,7 @@ exact: ltey. Qed. Lemma integral_normal_prob_dirac (s : R) (m : R) V : - (0 < s)%R -> + (s != 0)%R -> measurable V -> \int[normal_prob m s]_x0 (\d_x0 V) = normal_prob m s V. Proof. @@ -356,7 +356,7 @@ Qed. Section conjugate_normal_property. Lemma conjugate_normal1 (m1 m2 s1 s2 : R) V : measurable V -> - (0 < s1)%R -> s2 != 0%R -> + s1 != 0%R -> s2 != 0%R -> \int[normal_prob m1 s1]_x normal_prob (m2 + x) s2 V = \int[mu]_(y in V) \int[mu]_x (normal_pdf (m2 + x) s2 y * normal_pdf m1 s1 x)%:E. Proof. @@ -427,7 +427,7 @@ move/negP in xV. by move/mem_set. Qed. -Lemma conjugate_normal2 (y m1 m2 s1 s2 : R) : s1 != 0%R -> (0 < s2)%R -> +Lemma conjugate_normal2 (y m1 m2 s1 s2 : R) : s1 != 0%R -> s2 != 0%R -> \int[mu]_x (normal_pdf (m1 + x)%E s1 y * normal_pdf m2 s2 x)%:E = (normal_peak s1 * normal_peak s2)%:E * \int[mu]_z (normal_fun (m1 + z) s1 y * normal_fun m2 s2 z)%:E. @@ -447,7 +447,8 @@ rewrite -ge0_integralZl//=; last 3 first. - by move=> z _; rewrite lee_fin mulr_ge0// expR_ge0. - by rewrite lee_fin mulr_ge0// ?normal_peak_ge0. apply: eq_integral => /= z _. -rewrite /normal_pdf (negbTE s10) gt_eqF//. +rewrite /normal_pdf (negbTE s10). +rewrite (negbTE s20). rewrite /normal_pdf0. rewrite mulrACA. rewrite /normal_fun. @@ -457,17 +458,137 @@ Qed. Lemma normal_peak1 : normal_peak 1 = (Num.sqrt (pi *+ 2))^-1%R :> R. Proof. by rewrite /normal_peak expr1n mul1r. Qed. -Lemma conjugate_normal (m2 y : R) V : measurable V -> - \int[normal_prob y (Num.sqrt 2)^-1]_x normal_prob (m2 + x) 1 V = - normal_prob (y + m2) (Num.sqrt (3 / 2)) V. +(* Variable elimination and integration [Shan, Section 3.5, (9)], + * also known as the reproductive property of normal distribution. + *) +Lemma normal_probD (m1 s1 m2 s2: R) V : s1 != 0%R -> s2 != 0%R -> +measurable V -> + \int[normal_prob m1 s1]_x normal_prob (m2 + x) s2 V = + normal_prob (m1 + m2) (Num.sqrt (s1 ^+ 2 + s2 ^+ 2)) V. +Proof. +move=> s10 s20 mV. +rewrite conjugate_normal1//; apply: eq_integral => y _. +clear V mV. +rewrite conjugate_normal2//. +have s1s20 : (s1 ^+ 2 + s2 ^+ 2 != 0)%R. + by rewrite lt0r_neq0// addr_gt0// exprn_even_gt0. +have sqs1s20 : Num.sqrt (s1 ^+ 2 + s2 ^+ 2) != 0%R. + by rewrite lt0r_neq0// sqrtr_gt0 addr_gt0// exprn_even_gt0. +rewrite normal_pdfE /normal_pdf0//. +set S1 := (s1 ^+ 2)%R. +set S2 := (s2 ^+ 2)%R. +transitivity (((Num.sqrt S1 * Num.sqrt S2 * pi *+ 2)^-1)%:E * + \int[mu]_x (expR + (- (x - (y * s1 ^+ 2 + m1 * s2 ^+ 2 - m2 * s1 ^+ 2) + / (s1 ^+ 2 + s2 ^+ 2)%R ) ^+ 2 + / ((Num.sqrt ((s1 ^+ 2 * s2 ^+ 2) / (s1 ^+ 2 + s2 ^+ 2)%R) ^+ 2) *+ 2) + - (y - (m1 + m2)) ^+ 2 / ((s1 ^+ 2 + s2 ^+ 2) *+ 2)))%:E). + congr *%E. + rewrite /normal_peak. + congr EFin. + rewrite -2!(mulr_natr (_ * pi)). + rewrite !(sqrtrM 2) ?(@mulr_ge0 _ _ pi) ?sqr_ge0 ?pi_ge0//. + rewrite !(sqrtrM pi) ?sqr_ge0//. + rewrite ![in LHS]invfM. + rewrite mulrACA -(@sqrtrV _ 2)// -(expr2 (_ _^-1)). + rewrite (@sqr_sqrtr _ 2^-1) ?invr_ge0//. + rewrite mulrACA -(@sqrtrV _ pi) ?pi_ge0//. + rewrite -(expr2 (_ _^-1)) (@sqr_sqrtr _ pi^-1) ?invr_ge0// ?pi_ge0//. + rewrite -!invfM; congr GRing.inv. + by rewrite -[in RHS]mulr_natr (mulrC _ (Num.sqrt _)). + apply: eq_integral. + move=> x _. + rewrite -expRD. + congr EFin. + congr expR. + rewrite sqr_sqrtr; last first. + rewrite mulr_ge0 ?invr_ge0// ?addr_ge0 ?(@mulr_ge0 _ (_ ^+ 2))// ?sqr_ge0//. + field. + by apply/and3P; split. + +set DS12 := S1 + S2. +set MS12 := (S1 * S2)%R. +set C := (((y * s1 ^+ 2)%R + (m1 * s2 ^+ 2)%R)%E - m2 * s1 ^+ 2) / DS12. + +under eq_integral do rewrite expRD EFinM. +rewrite ge0_integralZr//=; last 3 first. + apply/measurable_EFinP. + apply: measurableT_comp => //. + apply: measurable_funM => //. + apply: measurableT_comp => //. + apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)%R) => //. + exact: measurable_funD. + by move=> z _; rewrite lee_fin ?expR_ge0. + by rewrite lee_fin expR_ge0. +rewrite /normal_peak /normal_fun. +rewrite [in RHS]EFinM. +rewrite [in RHS]sqr_sqrtr//; last first. + by rewrite addr_ge0// sqr_ge0. +rewrite muleA; congr *%E; last by rewrite -mulNr. + +(* gauss integral *) +have MS12DS12_gt0 : (0 < MS12 / DS12)%R. + rewrite divr_gt0//. + by rewrite mulr_gt0// exprn_even_gt0. + by rewrite addr_gt0// exprn_even_gt0. +transitivity (((Num.sqrt S1 * Num.sqrt S2 * pi *+ 2)^-1)%:E + * \int[mu]_x ((normal_peak (Num.sqrt (MS12 / DS12)))^-1%:E + * (normal_pdf C (Num.sqrt (MS12 / DS12)) x)%:E)). + congr *%E. + apply: eq_integral => x _. + rewrite -EFinM; congr EFin. + rewrite normal_pdfE; last first. + apply: lt0r_neq0. + by rewrite sqrtr_gt0. + rewrite mulrA mulVf// ?mul1r//. + rewrite lt0r_neq0// invr_gt0 sqrtr_gt0 pmulrn_lgt0// mulr_gt0// ?pi_gt0//. + rewrite exprn_even_gt0//=. + by rewrite lt0r_neq0// sqrtr_gt0. +rewrite ge0_integralZl//; last 3 first. +- apply/measurable_EFinP. + exact: measurable_normal_pdf. +- move=> x _. + rewrite lee_fin. + exact: normal_pdf_ge0. +- rewrite lee_fin invr_ge0. + exact: normal_peak_ge0. +rewrite integral_normal_pdf. +rewrite mule1 -EFinM; congr EFin. +rewrite -invfM; congr GRing.inv. +rewrite -sqrtrM ?sqr_ge0//. +rewrite /normal_peak sqr_sqrtr; last by rewrite ltW. +rewrite -3!mulrnAr. +rewrite (sqrtrM (pi *+ 2)); last by rewrite ltW. +rewrite invfM mulrCA. +rewrite -{1}(@sqr_sqrtr _ (pi *+ 2)); last by rewrite pmulrn_lge0 ?pi_ge0. +rewrite -2!(mulrA (Num.sqrt _)) divff// ?mulr1; last first. + by rewrite lt0r_neq0// sqrtr_gt0 pmulrn_lgt0 ?pi_gt0. +rewrite (sqrtrM (DS12^-1)); last by rewrite mulr_ge0 ?sqr_ge0. +rewrite sqrtrV; last by rewrite addr_ge0 ?sqr_ge0. +rewrite invfM invrK. +rewrite mulrAC mulrA mulVf ?mul1r; last first. + by rewrite lt0r_neq0// sqrtr_gt0 mulr_gt0 ?exprn_even_gt0. +rewrite sqrtrM; last by rewrite addr_ge0 ?sqr_ge0. +by rewrite mulrC. +Qed. + +(* +Lemma conjugate_normal (m1 m2 : R) V : measurable V -> + \int[normal_prob m1 (Num.sqrt 2)^-1]_x normal_prob (m2 + x) 1 V = + normal_prob (m1 + m2) (Num.sqrt (3 / 2)) V. Proof. move=> mV. rewrite conjugate_normal1//; apply: eq_integral => x1 _. clear V mV. rewrite conjugate_normal2//. rewrite normal_pdfE// /normal_pdf0. -transitivity ((pi * Num.sqrt 2)^-1%:E * - \int[mu]_x0 (expR (- (Num.sqrt (3 / 2) ^+ 2)%R * (x0 - (x1 + 2 * y - m2) / 3) ^+ 2 - (x1 - (y + m2)) ^+ 2 / ((Num.sqrt (3 / 2) ^+ 2) *+ 2)))%:E). +transitivity ((2 * pi * (Num.sqrt 2)^-1 * 1)^-1%:E * + \int[mu]_x0 (expR + (- (((Num.sqrt 2)^-1 ^+ 2 + 1 ^+ 2)%R / + (((Num.sqrt 2)^-1 ^+ 2 * 1 ^+ 2) *+ 2))%R * + (x0 - (x1 * (Num.sqrt 2)^-1 ^+ 2 + m1 * 1 ^+ 2 - m2 * (Num.sqrt 2)^-1 ^+ 2) + / ((Num.sqrt 2)^-1 ^+ 2 + 1 ^+ 2)%R ) ^+ 2 - + (x1 - (m1 + m2)) ^+ 2 / ((Num.sqrt (3 / 2) ^+ 2) *+ 2)))%:E). congr *%E. rewrite normal_peak1. rewrite /normal_peak. @@ -475,19 +596,24 @@ transitivity ((pi * Num.sqrt 2)^-1%:E * rewrite exprVn sqr_sqrtr//. rewrite -mulr_natl mulrA divff// mul1r. rewrite -(mulr_natl pi) sqrtrM// invfM. - rewrite invfM. - rewrite -mulrA mulrC. + rewrite invfM divr1. + rewrite -[LHS]mulrA. + rewrite -[in RHS]mulrA [in RHS]mulrC. rewrite -(invfM (Num.sqrt pi)) -expr2. - by rewrite sqr_sqrtr// pi_ge0. + rewrite sqr_sqrtr// ?pi_ge0//. + rewrite -{3}(@sqr_sqrtr _ 2%R)// expr2. + rewrite 2!invfM invrK. + by rewrite -mulf_div divff// mulr1 mulrC. apply: eq_integral. move=> z _. rewrite -expRD. congr EFin. congr expR. - rewrite !sqrrD. - rewrite exprVn sqr_sqrtr//. - rewrite sqr_sqrtr//. - lra. + rewrite (_ : (Num.sqrt (3 / 2) ^+ 2) = 3 / 2)%R; last by rewrite sqr_sqrtr. + rewrite (_ : (Num.sqrt 2)^-1 ^+ 2 = 1 / 2)%R; last first. + by rewrite exprVn sqr_sqrtr// div1r. + by field. + (* gauss integral *) under eq_integral do rewrite expRD EFinM. rewrite ge0_integralZr//=; last first. @@ -504,7 +630,7 @@ rewrite mulNr EFinM muleA. congr *%E; last first. by rewrite sqr_sqrtr//. rewrite [X in _ * X = _](_ : _ = (Num.sqrt ((1 / 3) * pi *+ 2))%:E * - \int[mu]_z (normal_pdf ((x1 + 2 * y - m2) / 3) (Num.sqrt (1 / 3)) z)%:E); last first. + \int[mu]_z (normal_pdf ((x1 + 2 * m1 - m2) / 3) (Num.sqrt (1 / 3)) z)%:E); last first. rewrite -ge0_integralZl//=; last 2 first. by apply/measurable_EFinP; exact: measurable_normal_pdf. by move=> /= z _; rewrite lee_fin normal_pdf_ge0. @@ -524,6 +650,7 @@ rewrite !sqrtrM// invfM sqrtrV// -mulrA; congr *%R. rewrite -[X in (_ / X)%R]sqr_sqrtr ?pi_ge0//. by rewrite expr2 invfM mulrA divff ?div1r// gt_eqF// sqrtr_gt0 pi_gt0. Qed. +*) End conjugate_normal_property. @@ -531,10 +658,13 @@ Lemma int_normal_mu_noisy2 y V : measurable V -> int_normal_noisy2 y V = int_mu_noisy2 y V. Proof. move=> mV. -have := @conjugate_normal 0 (y.1 / 2) _ mV. -under eq_integral do rewrite add0r. rewrite /int_normal_noisy2. +have := (@normal_probD (y.1 / 2) (Num.sqrt 2)^-1 0 1 _ _ _ mV). +under eq_integral do rewrite add0r. rewrite addr0. +rewrite (_: ((Num.sqrt 2)^-1 ^+ 2 + 1 ^+ 2 = 3 / 2)%R); last first. + rewrite exprVn sqr_sqrtr// expr1n -[in LHS]div1r -{3}(@divff _ 1%R)//. + rewrite addf_div// 2!mulr1 mul1r (_:1%R = 1%:R)// -natrD. exact. Qed. @@ -649,6 +779,33 @@ Local Definition executed_noisy1' := (measurable_acc_typ [:: Real; Unit; Real] 0))) (ret (measurable_acc_typ [:: Real; Real; Unit; Real] 0)))) y V : \bar R). +(* +Local Definition executed_noisy1'_alter := + (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => + letin' + (score + (measurable_funM + (measurableT_comp (measurable_powRr (expR 1)) + (measurable_funB (kr 0) + (measurable_funM + (measurable_funX 2%R (measurable_acc_typ [:: Real] 0)) + (kr 4^-1)))) (kr (Num.sqrt (4 * pi))^-1))) + (letin' + (sample + (fun x : unit * (g_sigma_algebraType R.-ocitv.-measurable * unit) => + probability_normal_prob__canonical__measure_Probability + (x.2.1 / 2) (Num.sqrt 2)^-1) + (measurableT_comp (measurable_normal_prob2 neq0Vsqrt2) + (measurable_funM (measurable_acc_typ [:: Unit; Real] 1) (kr 2^-1)))) + (letin' + (sample + (fun=> probability_normal_prob__canonical__measure_Probability 0 1) + (measurableT_comp (measurable_normal_prob2 neq01) (kr 0))) + (ret + (measurable_funD (measurable_acc_typ [:: Real; Real; Unit; Real] 1) + (measurable_acc_typ [:: Real; Real; Unit; Real] 0))))) y V : \bar R). +*) + Lemma execP_noisy0'E y V : execP noisy0' y V = executed_noisy0' y V. Proof. rewrite !execP_letin. @@ -817,3 +974,99 @@ Lemma noisy'02 y V : measurable V -> execP noisy0' y V = execP noisy2' y V. Proof. by move=> mV; rewrite noisy'01// noisy'12. Qed. End noisy_verification. + +(* Trying to shows why rewriting noisy1 to noisy2 is reproductive property *) +(* +Section rewrite noisy1'_to_variable_addition. + +Definition noisy1'_alter : @exp R _ [:: ("y0", Real)] Real := + [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2} * {4^-1}:R)) * + {(Num.sqrt (4 * pi))^-1}:R in + let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R]} in + let "x1" := Sample {exp_normal1 [{0}:R]} in + return #{"x"} + #{"x1"}]. + +Definition noisy1_alter : @exp R _ [:: ("y0", Real)] _ := + [Normalize {noisy1'_alter}]. + +Lemma execP_noisy1'_alterE y V : +@execP R [:: ("y0", Real)] Real noisy1'_alter y V = executed_noisy1'_alter y V. +Proof. +rewrite 3!execP_letin. +rewrite execP_return/=. +rewrite (@execD_bin _ _ binop_add)/=. +rewrite (exp_var'E "x") (exp_var'E "x1"). +rewrite (execD_var_erefl "x") (execD_var_erefl "x1")/=. +rewrite 2!execP_sample. +rewrite 2!execD_normal/=. +rewrite execD_real/=. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_real/=. +rewrite execP_score. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow_real/=. +rewrite (@execD_bin _ _ binop_minus)/=. +rewrite execD_real/=. +rewrite (@execD_bin _ _ binop_mult)/=. +rewrite execD_pow/=. +rewrite 2!execD_real/=. +rewrite 2!(exp_var'E "y0") 2!(execD_var_erefl "y0")/=. +rewrite /executed_noisy1'_alter. +rewrite /=. +Abort. + +Lemma noisy1_alterE y V : measurable V -> + @execP R [:: ("y0", Real)] Real noisy1' y V = + @execP R [:: ("y0", Real)] Real noisy1'_alter y V. +Proof. +move=> mV. +rewrite 3![in RHS]execP_letin. +rewrite [in RHS]execP_return/=. +rewrite [in RHS](@execD_bin _ _ binop_add)/=. +rewrite [in RHS](exp_var'E "x") (exp_var'E "x1"). +rewrite [in RHS](execD_var_erefl "x") (execD_var_erefl "x1")/=. +rewrite 2![in RHS]execP_sample. +rewrite 2![in RHS]execD_normal/=. +rewrite [in RHS]execD_real/=. +rewrite [in RHS](@execD_bin _ _ binop_mult)/=. +rewrite [in RHS]execD_real/=. +rewrite [in RHS]execP_score. +rewrite [in RHS](@execD_bin _ _ binop_mult)/=. +rewrite [in RHS]execD_pow_real/=. +rewrite [in RHS](@execD_bin _ _ binop_minus)/=. +rewrite [in RHS]execD_real/=. +rewrite [in RHS](@execD_bin _ _ binop_mult)/=. +rewrite [in RHS]execD_pow/=. +rewrite 2![in RHS]execD_real/=. +rewrite 2![in RHS](exp_var'E "y0") 2!(execD_var_erefl "y0")/=. +rewrite [in RHS]letin'E/=. +under [RHS]eq_integral. + move=> x _. + rewrite letin'E/=. + under eq_integral. + move=> z _. + rewrite letin'E/=. + rewrite integral_normal_prob_dirac//; last first. + admit. (* standard property of measurable sets *) + rewrite (_ : (fun x0 => V (z + x0)) = ((fun x0 => z + x0) @^-1` V)); last by []. + rewrite (_ : normal_prob 0 1 ((fun x0 => z + x0) @^-1` V) = normal_prob z 1 V); last first. + admit. (* general version of integration by substitution *) + over. + over. +rewrite ge0_integral_mscale//; last first. + move=> x _. + apply: integral_ge0. + by move=> z _. +rewrite integral_dirac// diracT mul1e. +rewrite -ge0_integralZl//; last first. + exact: emeasurable_normal_prob. +rewrite sub0r. +rewrite -expRM mul1r. + +rewrite execP_noisy1'E. +by rewrite executed_noisy1'_semantics. +Abort. + +End rewrite noisy1'_to_variable_addition. +*) + From c700162ee9d2d812c5720b68c6a130a6411031b2 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 20 Mar 2025 18:19:55 +0900 Subject: [PATCH 40/48] various fixes - rm intermediate defs - move beta to probability.v - Bernoulli, etc. syntax - order of args - derive shift now in master --- theories/lang_syntax.v | 212 +++------ theories/lang_syntax_examples.v | 64 +-- theories/lang_syntax_noisy.v | 705 ++++++++++-------------------- theories/lang_syntax_table_game.v | 142 +++--- theories/prob_lang.v | 24 +- theories/probability.v | 138 +++--- 6 files changed, 456 insertions(+), 829 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index ca98d0a4db..933b5f89ca 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -17,12 +17,6 @@ From mathcomp Require Import prob_lang lang_syntax_util. (* Probabilistic Programming Language in Coq using s-finite kernels in Coq. *) (* APLAS 2023 *) (* *) -(* beta distribution *) -(* ``` *) -(* beta_pdf == probability density function for beta *) -(* beta_prob == beta probability measure *) -(* ``` *) -(* *) (* ``` *) (* typ == syntax for types of data structures *) (* measurable_of_typ t == the measurable type corresponding to type t *) @@ -76,7 +70,6 @@ Reserved Notation "e -P> k" (at level 40). Local Open Scope classical_set_scope. Local Open Scope ring_scope. -Local Open Scope ereal_scope. (* In this module, we use our lemma continuous_FTC2 to compute the value of * integration of the indicator function over the interval [0, 1]. @@ -202,18 +195,7 @@ Qed. End integral_indicator_function. End integral_indicator_function. -(* TODO: naming *) -(* duplicated, cvg_compNP *) -(* -Lemma cvg_atNP {T : topologicalType} {R : numFieldType} (f : R -> T) (a : R) (l : T) : - f x @[x --> a] --> l <-> (f \o -%R) x @[x --> (- a)%R] --> l. -Proof. -rewrite nbhsN. -have <-// : f x @[x --> a] = fmap [eta f \o -%R] ((- x)%R @[x --> a]). -by apply/seteqP; split=> A; move=> [/= e e0 H]; exists e => //= B /H/=; rewrite opprK. -Qed. -*) - +(* TODO: move? *) Lemma derivable_oo_bnd_id {R : numFieldType} (a b : R) : derivable_oo_LRcontinuous (@id R^o) a b. Proof. @@ -235,13 +217,8 @@ split. - by apply/cvg_at_leftNP; rewrite -compA oppK. Qed. -Module increasing_change_of_variables_from_decreasing. -Section lt0. -Context {R : realType}. -Notation mu := lebesgue_measure. -Implicit Types (F G f : R -> R) (a b : R). - -Lemma continuous_withinN f a b : (a < b)%R -> +(* NB: not used *) +Lemma continuous_withinN {R : realType} (f : R -> R) a b : (a < b)%R -> {within `[(- b)%R, (- a)%R], continuous (f \o -%R)} -> {within `[a, b], continuous f}. Proof. @@ -262,83 +239,6 @@ move=> ab cf. by rewrite /= opprK in cf. Qed. -(* duplicated, integration_by_substitution_oppr *) -(* -Lemma oppr_change (f : R -> R) a b : (a < b)%R -> - {within `[a, b], continuous f} -> - \int[mu]_(x in `[a, b]) (f x)%:E = - \int[mu]_(x in `[-%R b, -%R a]) ((f \o -%R) x)%:E. -Proof. -move=> ab cf. -have dN : ((-%R : R -> R^o)^`() = cst (-1) :> (R -> R))%R. (* TODO: lemma? *) - by apply/funext => x/=; rewrite derive1E deriveN// derive_id. -rewrite integration_by_substitution_decreasing//. -- by apply: eq_integral => /= x _; rewrite dN/= opprK mulr1 -compA/= opprK. -- by move=> x y _ _ yx; rewrite ltrN2. -- by move=> y yab; rewrite dN; exact: cvg_cst. -- by rewrite dN; exact: is_cvg_cst. -- by rewrite dN; exact: is_cvg_cst. -- by apply: (@derivable_oo_bndN _ id) => //; exact: derivable_oo_bnd_id. -- apply: continuous_withinN. - + by rewrite ltrN2. - + rewrite -(_ : f = (f \o -%R) \o -%R)//; last first. - by apply/funext => y; rewrite /= opprK. - by rewrite !opprK. -Qed. -*) - -End lt0. -End increasing_change_of_variables_from_decreasing. - -(* duplicated? *) -Lemma decreasing_nonincreasing {R : realType} (F : R -> R) (J : interval R) : - {in J &, {homo F : x y /~ (x < y)%R}} -> - {in J &, {homo F : x y /~ (x <= y)%R}}. -Proof. -move=> dF. -move=> x y x01 y01. -by rewrite le_eqVlt => /predU1P[->//|/dF] => /(_ x01 y01)/ltW. -Qed. - -Local Close Scope ereal_scope. -Lemma cvg_comp_filter {R : realType} (f g : R -> R) (r l : R) : - continuous f -> - (f \o g) x @[x --> r] --> l -> - f x @[x --> g r] --> l. -Proof. -move=> cf fgrl. -apply/(@cvgrPdist_le _ R^o) => /= e e0. -have e20 : 0 < e / 2 by rewrite divr_gt0. -move/(@cvgrPdist_le _ R^o) : fgrl => /(_ _ e20) fgrl. -have := cf (g r). -move=> /(@cvgrPdist_le _ R^o) => /(_ _ e20)[x x0]H. -exists (minr x (e/2)). - by rewrite lt_min x0. -move=> z. -rewrite /ball_ /= => grze. -rewrite -[X in X - _](subrK (f (g r))). -rewrite -(addrA _ _ (- f z)). -apply: (le_trans (ler_normD _ _)). -rewrite (splitr e) lerD//. - case: fgrl => d /= d0 K. - apply: K. - by rewrite /ball_/= subrr normr0. -apply: H => /=. -by rewrite (lt_le_trans grze)// ge_min lexx. -Qed. -Local Open Scope ereal_scope. - -Lemma deriveX_idfun {R : realType} n x : - 'D_1 (@GRing.exp R^o ^~ n.+1) x = n.+1%:R *: (x ^+ n)%R. -Proof. by rewrite exp_derive /GRing.scale/= mulr1. Qed. - -Lemma decreasing_onem {R : numDomainType} : {homo (fun x : R => (1 - x)%R) : x y /~ (x < y)%R}. -Proof. -move=> b a ab. -by rewrite -ltrN2 !opprB ltr_leB. -Qed. - -(* TODO: move *) Lemma integral_exprn {R : realType} n : fine (\int[lebesgue_measure]_(x in `[0%R, 1%R]) (x ^+ n)%:E) = n.+1%:R^-1%R :> R. Proof. @@ -361,7 +261,7 @@ have dcF : derivable_oo_LRcontinuous F 0 1. have dFE : {in `]0%R, 1%R[, F^`() =1 (fun x : R => x ^+ n : R)%R}. move=> x x01. rewrite derive1Ml; last exact: exprn_derivable. - by rewrite derive1E deriveX_idfun mulrA mulVf// mul1r. + by rewrite derive1E exp_derive !mulrA mulVf// mulr1 mul1r. rewrite (@continuous_FTC2 _ (fun x : R => x ^+ n)%R F)//. by rewrite /F/= expr1n expr0n/= mulr1 mulr0 subr0. by apply: continuous_subspaceT; exact: exprn_continuous. @@ -605,7 +505,7 @@ Inductive exp : flag -> ctx -> typ -> Type := | exp_bool g : bool -> exp D g Bool | exp_nat g : nat -> exp D g Nat | exp_real g : R -> exp D g Real -| exp_pow g : nat -> exp D g Real -> exp D g Real (* TODO: nat should be at second *) +| exp_pow g : exp D g Real -> nat -> exp D g Real | exp_pow_real g : R (* base *) -> exp D g Real -> exp D g Real | exp_bin (b : binop) g : exp D g (type_of_binop b) -> exp D g (type_of_binop b) -> exp D g (type_of_binop b) @@ -620,10 +520,9 @@ Inductive exp : flag -> ctx -> typ -> Type := | exp_bernoulli g : exp D g Real -> exp D g (Prob Bool) | exp_binomial g (n : nat) : exp D g Real -> exp D g (Prob Nat) | exp_uniform g (a b : R) (ab : (a < b)%R) : exp D g (Prob Real) -| exp_beta g (a b : nat) (* NB: should be R *) : exp D g (Prob Real) +| exp_beta g (a b : nat) : exp D g (Prob Real) | exp_poisson g : nat -> exp D g Real -> exp D g Real -| exp_normal g (s : R) (s0 : (s != 0)%R) - : exp D g Real -> exp D g (Prob Real) (* NB: 0 < s *) +| exp_normal g : exp D g Real -> forall (s : R), (s != 0)%R -> exp D g (Prob Real) | exp_normalize g t : exp P g t -> exp D g (Prob t) | exp_letin g t1 t2 str : exp P g t1 -> exp P ((str, t1) :: g) t2 -> exp P g t2 @@ -661,7 +560,7 @@ Arguments exp_binomial {R g} &. Arguments exp_uniform {R g} &. Arguments exp_beta {R g} &. Arguments exp_poisson {R g}. -Arguments exp_normal {R g _} &. +Arguments exp_normal {R g} &. Arguments exp_normalize {R g _}. Arguments exp_letin {R g} & {_ _}. Arguments exp_sample {R g} & {t}. @@ -680,7 +579,7 @@ Notation "n ':N'" := (@exp_nat _ _ n%N) (in custom expr at level 1) : lang_scope. Notation "r ':R'" := (@exp_real _ _ r%R) (in custom expr at level 1, format "r :R") : lang_scope. -Notation "e ^+ n" := (exp_pow n e) +Notation "e ^+ n" := (exp_pow e n) (in custom expr at level 1) : lang_scope. Notation "e `^ r" := (exp_pow_real e r) (in custom expr at level 1) : lang_scope. @@ -732,6 +631,16 @@ Notation "'Score' e" := (exp_score e) (in custom expr at level 6) : lang_scope. Notation "'Normalize' e" := (exp_normalize e) (in custom expr at level 0) : lang_scope. +Notation "'Bernoulli' p" := (exp_bernoulli p) + (in custom expr at level 6) : lang_scope. +Notation "'Binomial' n k" := (exp_binomial n k) + (in custom expr at level 6) : lang_scope. +Notation "'Uniform' a b ab" := (exp_uniform a b ab) + (in custom expr at level 6) : lang_scope. +Notation "'Beta' a b" := (exp_beta a b) + (in custom expr at level 6) : lang_scope. +Notation "'Normal' m s s0" := (exp_normal m s s0) + (in custom expr at level 6) : lang_scope. Notation "'if' e1 'then' e2 'else' e3" := (exp_if e1 e2 e3) (in custom expr at level 7) : lang_scope. Notation "( e )" := e @@ -746,21 +655,21 @@ Fixpoint free_vars k g t (e : @exp R k g t) : seq string := | exp_bool _ _ => [::] | exp_nat _ _ => [::] | exp_real _ _ => [::] - | exp_pow _ _ e => free_vars e + | exp_pow _ e _ => free_vars e | exp_pow_real _ _ e => free_vars e - | exp_bin _ _ e1 e2 => free_vars e1 ++ free_vars e2 - | exp_rel _ _ e1 e2 => free_vars e1 ++ free_vars e2 - | exp_rel_real _ _ e1 e2 => free_vars e1 ++ free_vars e2 + | exp_bin _ _ e1 e2 => free_vars e1 ++ free_vars e2 + | exp_rel _ _ e1 e2 => free_vars e1 ++ free_vars e2 + | exp_rel_real _ _ e1 e2 => free_vars e1 ++ free_vars e2 | exp_pair _ _ _ e1 e2 => free_vars e1 ++ free_vars e2 | exp_proj1 _ _ _ e => free_vars e | exp_proj2 _ _ _ e => free_vars e | exp_var _ x _ _ => [:: x] - | exp_bernoulli _ e => free_vars e - | exp_binomial _ _ e => free_vars e + | exp_bernoulli _ e => free_vars e + | exp_binomial _ _ e => free_vars e | exp_uniform _ _ _ _ => [::] - | exp_beta _ _ _ => [::] + | exp_beta _ _ _ => [::] | exp_poisson _ _ e => free_vars e - | exp_normal _ _ _ e => free_vars e + | exp_normal _ e _ _ => free_vars e | exp_normalize _ _ e => free_vars e | exp_letin _ _ _ x e1 e2 => free_vars e1 ++ rem x (free_vars e2) | exp_sample _ _ _ => [::] @@ -775,11 +684,9 @@ End free_vars. Definition dval R g t := @mctx R g -> @mtyp R t. Definition pval R g t := R.-sfker @mctx R g ~> @mtyp R t. - Section weak. Context {R : realType}. Implicit Types (g h : ctx) (x : string * typ). -Local Open Scope ring_scope. Fixpoint mctx_strong g h x (f : @mctx R (g ++ x :: h)) : @mctx R (g ++ h) := match g as g0 return mctx (g0 ++ x :: h) -> mctx (g0 ++ h) with @@ -866,23 +773,6 @@ Context {R : realType}. Implicit Type (g : ctx) (str : string). Local Open Scope lang_scope. -Local Open Scope ring_scope. - -(* TODO: PR *) -Lemma measurable_powRr b : measurable_fun setT (@powR R b). -Proof. -rewrite /powR. -apply: measurable_fun_if => //. - rewrite preimage_true setTI/=. - case: (b == 0); rewrite ?set_true ?set_false. - apply: measurableT_comp => //. - exact: measurable_fun_eqr. - exact: measurable_fun_set0. -rewrite preimage_false setTI. -apply: measurableT_comp => //. -exact: mulrr_measurable. -Qed. - Inductive evalD : forall g t, exp D g t -> forall f : dval R g t, measurable_fun setT f -> Prop := | eval_unit g : ([TT] : exp D g _) -D> cst tt ; ktt @@ -894,10 +784,10 @@ Inductive evalD : forall g t, exp D g t -> | eval_real g r : ([r:R] : exp D g _) -D> cst r ; kr r | eval_pow g n (e : exp D g _) f mf : e -D> f ; mf -> - [e ^+ {n}] -D> (fun x => f x ^+ n) ; (measurable_funX n mf) + [e ^+ {n}] -D> (fun x => (f x ^+ n)%R) ; (measurable_funX n mf) | eval_pow_real g (e : exp D g _) r f mf : e -D> f ; mf -> - [{r} `^ e] -D> (fun x => r `^ (f x)) ; measurableT_comp (measurable_powRr r) mf + [{r} `^ e] -D> (fun x => (r `^ (f x))%R) ; measurableT_comp (measurable_powRr r) mf | eval_bin g bop (e1 : exp D g _) f1 mf1 e2 f2 mf2 : e1 -D> f1 ; mf1 -> e2 -D> f2 ; mf2 -> @@ -934,26 +824,24 @@ Inductive evalD : forall g t, exp D g t -> measurableT_comp measurable_bernoulli_prob mr | eval_binomial g n e r mr : - e -D> r ; mr -> (exp_binomial n e : exp D g _) -D> binomial_prob n \o r ; + e -D> r ; mr -> ([Binomial n e] : exp _ g _) -D> binomial_prob n \o r ; measurableT_comp (measurable_binomial_prob n) mr | eval_uniform g (a b : R) (ab : (a < b)%R) : - (exp_uniform a b ab : exp D g _) -D> cst (uniform_prob ab) ; + ([Uniform a b ab] : exp D g _) -D> cst (uniform_prob ab) ; measurable_cst _ | eval_beta g (a b : nat) : - (exp_beta a b : exp D g _) -D> cst (beta_prob a b) ; measurable_cst _ + ([Beta a b] : exp D g _) -D> cst (beta_prob a b) ; measurable_cst _ | eval_poisson g n (e : exp D g _) f mf : e -D> f ; mf -> exp_poisson n e -D> poisson_pmf ^~ n \o f ; measurableT_comp (measurable_poisson_pmf n measurableT) mf -(* TODO *) | eval_normal g s (s0 : (s != 0)%R) (e : exp D g _) r mr : e -D> r ; mr -> - (exp_normal s0 e : exp D g _) -D> - (fun x => @normal_prob _ (r x) s) ; + ([Normal e s s0] : exp D g _) -D> (fun x => @normal_prob _ (r x) s) ; measurableT_comp (measurable_normal_prob2 s0) mr | eval_normalize g t (e : exp P g t) k : @@ -1037,7 +925,7 @@ all: (rewrite {g t e u v mu mv hu}). - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. inj_ex H4; subst v. - inj_ex H2; subst e0. + inj_ex H0; subst e0. by move: H3 => /IH <-. - move=> g e r f mf ev IH {}v {}mv. inversion 1; subst g0 r0. @@ -1110,7 +998,7 @@ all: (rewrite {g t e u v mu mv hu}). by rewrite (IH _ _ H3). - move=> g s s0 e r mr ev IH {}v {}mv. inversion 1; subst g0 s1. - inj_ex H2; subst e0. + inj_ex H0; subst e0. inj_ex H3; subst v. by rewrite (IH _ _ H5). - move=> g t e k ev IH f mf. @@ -1201,7 +1089,7 @@ all: rewrite {g t e u v eu}. - move=> g n e f mf ev IH {}v {}mv. inversion 1; subst g0 n0. inj_ex H4; subst v. - inj_ex H2; subst e0. + inj_ex H0; subst e0. by move: H3 => /IH <-. - move=> g e b f mf ev IH {}v {}mv. inversion 1; subst g0 r. @@ -1276,7 +1164,7 @@ all: rewrite {g t e u v eu}. by rewrite (IH _ _ H3). - move=> g s s0 e r mr ev IH {}v {}mv. inversion 1; subst g0 s1. - inj_ex H2; subst e0. + inj_ex H0; subst e0. inj_ex H3; subst v. by rewrite (IH _ _ H5). - move=> g t e k ev IH {}v {}mv. @@ -1353,7 +1241,7 @@ all: rewrite {z g t}. - by do 2 eexists; exact: eval_bool. - by do 2 eexists; exact: eval_nat. - by do 2 eexists; exact: eval_real. -- move=> g n e [f [mf H]]. +- move=> g e [f [mf H]] n. by exists (fun x => (f x ^+ n)%R); eexists; exact: eval_pow. - move=> g r e [f [mf H]]. by exists (fun x => (r `^ (f x))%R); eexists; exact: eval_pow_real. @@ -1380,7 +1268,7 @@ all: rewrite {z g t}. - by eexists; eexists; exact: eval_beta. - move=> g h e [f [mf H]]. by exists (poisson_pmf ^~ h \o f); eexists; exact: eval_poisson. -- move=> g s s0 e [r [mr H]]. +- move=> g e [r [mr H]] s s0. exists (fun x => @normal_prob _ (r x) s : pprobability _ _). by eexists; exact: eval_normal. - move=> g t e [k ek]. @@ -1484,11 +1372,10 @@ Proof. exact/execD_evalD/eval_nat. Qed. Lemma execD_real g r : @execD g _ [r:R] = existT _ (cst r) (kr r). Proof. exact/execD_evalD/eval_real. Qed. -Local Open Scope ring_scope. Lemma execD_pow g (e : exp D g _) n : let f := projT1 (execD e) in let mf := projT2 (execD e) in - execD (exp_pow n e) = - @existT _ _ (fun x => f x ^+ n) (measurable_funX n mf). + execD (exp_pow e n) = + @existT _ _ (fun x => (f x ^+ n)%R) (measurable_funX n mf). Proof. by move=> f mf; apply/execD_evalD/eval_pow/evalD_execD. Qed. @@ -1496,7 +1383,7 @@ Qed. Lemma execD_pow_real g r (e : exp D g _) : let f := projT1 (execD e) in let mf := projT2 (execD e) in execD (exp_pow_real r e) = - @existT _ _ (fun x => r `^ f x) (measurableT_comp (measurable_powRr r) mf). + @existT _ _ (fun x => (r `^ f x)%R) (measurableT_comp (measurable_powRr r) mf). Proof. by move=> f mf; apply/execD_evalD/eval_pow_real/evalD_execD. Qed. @@ -1574,25 +1461,25 @@ Lemma execD_bernoulli g e : Proof. exact/execD_evalD/eval_bernoulli/evalD_execD. Qed. Lemma execD_binomial g n e : - @execD g _ (exp_binomial n e) = + @execD g _ [Binomial n e] = existT _ ((binomial_prob n : R -> pprobability nat R) \o projT1 (execD e)) (measurableT_comp (measurable_binomial_prob n) (projT2 (execD e))). Proof. exact/execD_evalD/eval_binomial/evalD_execD. Qed. Lemma execD_uniform g a b ab0 : - @execD g _ (exp_uniform a b ab0) = - existT _ (cst [the probability _ _ of uniform_prob ab0]) (measurable_cst _). + @execD g _ [Uniform a b ab0] = + existT _ (cst (uniform_prob ab0 : pprobability _ R)) (measurable_cst _). Proof. exact/execD_evalD/eval_uniform. Qed. Lemma execD_beta g a b : - @execD g _ (exp_beta a b) = - existT _ (cst [the probability _ _ of beta_prob a b]) (measurable_cst _). + @execD g _ [Beta a b] = + existT _ (cst (beta_prob a b : pprobability _ R)) (measurable_cst _). Proof. exact/execD_evalD/eval_beta. Qed. Lemma execD_normal g s s0 e : let f := projT1 (execD e) in let mf := projT2 (execD e) in - @execD g _ (@exp_normal _ _ s s0 e) = - existT _ (fun x => [the probability _ _ of @normal_prob _ (f x) s]) + @execD g _ [Normal e s s0] = + existT _ (fun x => @normal_prob _ (f x) s : pprobability _ R) (measurableT_comp (measurable_normal_prob2 s0) mf). Proof. exact/execD_evalD/eval_normal/evalD_execD. Qed. @@ -1670,5 +1557,4 @@ f_equal. apply: eq_kernel => y V. exact: He. Qed. - Local Close Scope lang_scope. diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index 69b6fb3928..c13f8ca378 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -12,7 +12,7 @@ From mathcomp Require Import probability. From mathcomp Require Import ring lra. (**md**************************************************************************) -(* # Examples using the Probabilistic Programming Language of lang_syntax.v *) +(* # Examples using the probabilistic Programming language of lang_syntax.v *) (* *) (* sample_pair1213 := normalize ( *) (* let x := sample (bernoulli 1/2) in *) @@ -131,8 +131,8 @@ Import Notations. Context {R : realType}. Definition sample_pair1213' : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli [{1 / 2}:R]} in - let "y" := Sample {exp_bernoulli [{1 / 3}:R]} in + [let "x" := Sample Bernoulli {1 / 2}:R in + let "y" := Sample Bernoulli {1 / 3}:R in return (#{"x"}, #{"y"})]. Definition sample_pair1213 : exp _ [::] _ := [Normalize {sample_pair1213'}]. @@ -193,8 +193,8 @@ Import Notations. Context {R : realType}. Definition sample_and1213' : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli [{1 / 2}:R]} in - let "y" := Sample {exp_bernoulli [{1 / 3}:R]} in + [let "x" := Sample Bernoulli {1 / 2}:R in + let "y" := Sample Bernoulli {1 / 3}:R in return #{"x"} && #{"y"}]. Lemma exec_sample_and1213' (A : set bool) : @@ -215,9 +215,9 @@ by congr (_%:E); congr (_ * _); field. Qed. Definition sample_and121212 : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli [{1 / 2}:R]} in - let "y" := Sample {exp_bernoulli [{1 / 2}:R]} in - let "z" := Sample {exp_bernoulli [{1 / 2}:R]} in + [let "x" := Sample Bernoulli {1 / 2}:R in + let "y" := Sample Bernoulli {1 / 2}:R in + let "z" := Sample Bernoulli {1 / 2}:R in return #{"x"} && #{"y"} && #{"z"}]. Lemma exec_sample_and121212 t U : @@ -244,13 +244,13 @@ Local Open Scope lang_scope. Import Notations. Context {R : realType}. -Definition bernoulli13_score := [Normalize - let "x" := Sample {@exp_bernoulli R [::] [{1 / 3}:R]} in +Definition bernoulli13_score : @exp R _ [::] _ := [Normalize + let "x" := Sample Bernoulli {1 / 3}:R in let "_" := if #{"x"} then Score {1 / 3}:R else Score {2 / 3}:R in return #{"x"}]. Lemma exec_bernoulli13_score : - execD bernoulli13_score = execD (exp_bernoulli [{1 / 5}:R]). + execD bernoulli13_score = execD [Bernoulli {1 / 5}:R]. Proof. apply: eq_execD. rewrite execD_bernoulli/= /bernoulli13_score execD_normalize_pt 2!execP_letin. @@ -285,13 +285,13 @@ by rewrite muleDl//; congr (_ + _)%E; rewrite !indicE /onem /=; case: (_ \in _); field. Qed. -Definition bernoulli12_score := [Normalize - let "x" := Sample {@exp_bernoulli R [::] [{1 / 2}:R]} in +Definition bernoulli12_score : @exp R _ [::] _ := [Normalize + let "x" := Sample Bernoulli {1 / 2}:R in let "r" := if #{"x"} then Score {1 / 3}:R else Score {2 / 3}:R in return #{"x"}]. Lemma exec_bernoulli12_score : - execD bernoulli12_score = execD (exp_bernoulli [{1 / 3}:R]). + execD bernoulli12_score = execD [Bernoulli {1 / 3}:R]. Proof. apply: eq_execD. rewrite execD_bernoulli/= /bernoulli12_score execD_normalize_pt 2!execP_letin. @@ -328,13 +328,13 @@ rewrite muleDl//; congr (_ + _)%E; Qed. (* https://dl.acm.org/doi/pdf/10.1145/2933575.2935313 (Sect. 4) *) -Definition bernoulli14_score := [Normalize - let "x" := Sample {@exp_bernoulli R [::] [{1 / 4}:R]} in +Definition bernoulli14_score : @exp R _ [::] _ := [Normalize + let "x" := Sample Bernoulli {1 / 4}:R in let "r" := if #{"x"} then Score {5}:R else Score {2}:R in return #{"x"}]. Lemma exec_bernoulli14_score : - execD bernoulli14_score = execD (exp_bernoulli [{5%:R / 11%:R}:R]). + execD bernoulli14_score = execD [Bernoulli {5%:R / 11%:R}:R]. Proof. apply: eq_execD. rewrite execD_bernoulli/= execD_normalize_pt 2!execP_letin. @@ -380,7 +380,7 @@ Open Scope lang_scope. Open Scope ring_scope. Definition sample_binomial3 : @exp R _ [::] _ := - [let "x" := Sample {exp_binomial 3 [{1 / 2}:R]} in + [let "x" := Sample Binomial {3} {1 / 2}:R in return #{"x"}]. Lemma exec_sample_binomial3 t U : measurable U -> @@ -410,7 +410,8 @@ Open Scope lang_scope. Open Scope ring_scope. Definition binomial2p (p : R) : @exp R _ [::] _ := - [let "x" := Sample {exp_binomial 2 [{p}:R]} in return #{"x"}]. + [let "x" := Sample Binomial {2} {p}:R in + return #{"x"}]. Definition return2 (p : R) : @exp R _ [::] _ := [let "_" := Score {p ^+ 2}:R in return {2}:N]. @@ -470,7 +471,7 @@ Qed. Lemma exec_score_fail (r : R) (r01 : (0 <= r <= 1)%R) : execP (g := [::]) [Score {r}:R] = - execP [let str := Sample {exp_bernoulli [{r}:R]} in + execP [let str := Sample Bernoulli {r}:R in if #str then return TT else {hard_constraint _}]. Proof. move: r01 => /andP[r0 r1]//. @@ -504,7 +505,7 @@ Local Open Scope lang_scope. Context (R : realType). Definition uniform_syntax : @exp R _ [::] _ := - [let "p" := Sample {exp_uniform 0 1 (@ltr01 R)} in + [let "p" := Sample Uniform {0} {1} {ltr01} in return #{"p"}]. Lemma exec_uniform_syntax t U : measurable U -> @@ -539,7 +540,7 @@ Local Open Scope lang_scope. Context (R : realType). Definition guard : @exp R _ [::] _ := [ - let "p" := Sample {exp_bernoulli [{1 / 3}:R]} in + let "p" := Sample Bernoulli {1 / 3}:R in let "_" := if #{"p"} then return TT else Score {0}:R in return #{"p"} ]. @@ -562,7 +563,7 @@ Local Open Scope lang_scope. Context (R : realType). Definition binomial_le : @exp R _ [::] Bool := - [let "a2" := Sample {exp_binomial 3 [{1 / 2}:R]} in + [let "a2" := Sample Binomial {3} {1 / 2}:R in return {1}:N <= #{"a2"}]. Lemma exec_binomial_le t U : @@ -586,7 +587,7 @@ lra. Qed. Definition binomial_guard : @exp R _ [::] Nat := - [let "a1" := Sample {exp_binomial 3 [{1 / 2}:R]} in + [let "a1" := Sample Binomial {3} {1 / 2}:R in let "_" := if #{"a1"} == {1}:N then return TT else Score {0}:R in return #{"a1"}]. @@ -650,9 +651,9 @@ by rewrite setTI//; exact: measurable_XMonemX. Qed. Lemma beta_bernoulli_bernoulli U : measurable U -> - @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in - Sample {exp_bernoulli [#{"p"}]}] tt U = - @execP R [::] _ [Sample {exp_bernoulli [{3 / 5}:R]}] tt U. + @execP R [::] _ [let "p" := Sample Beta {6} {4} in + Sample Bernoulli #{"p"}] tt U = + @execP R [::] _ [Sample Bernoulli {3 / 5}:R] tt U. Proof. move=> mU. rewrite execP_letin !execP_sample execD_beta !execD_bernoulli/=. @@ -749,7 +750,7 @@ Import Notations. Context {R : realType}. Definition staton_bus_syntax0 : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli [{2 / 7}:R]} in + [let "x" := Sample Bernoulli {2 / 7}:R in let "r" := if #{"x"} then return {3}:R else return {10}:R in let "_" := Score {exp_poisson 4 [#{"r"}]} in return #{"x"}]. @@ -843,7 +844,7 @@ Import Notations. Context {R : realType}. Definition staton_busA_syntax0 : @exp R _ [::] _ := - [let "x" := Sample {exp_bernoulli [{2 / 7}:R]} in + [let "x" := Sample Bernoulli {2 / 7}:R in let "_" := let "r" := if #{"x"} then return {3}:R else return {10}:R in Score {exp_poisson 4 [#{"r"}]} in @@ -856,7 +857,8 @@ Let sample_bern : R.-sfker munit ~> mbool := sample _ (measurableT_comp measurable_bernoulli_prob (measurable_cst (2 / 7 : R)%R)). Let ite_3_10 : R.-sfker mbool * munit ~> measurableTypeR R := - ite macc0of2 (@ret _ _ _ (measurableTypeR R) R _ (kr 3)) (@ret _ _ _ (measurableTypeR R) R _ (kr 10)). + ite macc0of2 (@ret _ _ _ (measurableTypeR R) R _ (kr 3)) + (@ret _ _ _ (measurableTypeR R) R _ (kr 10)). Let score_poisson4 : R.-sfker measurableTypeR R * (mbool * munit) ~> munit := score (measurableT_comp (measurable_poisson_pmf 4 measurableT) @@ -1020,6 +1022,6 @@ Example letinC_ground (g := [:: ("a", Unit); ("b", Bool)]) t1 t2 execP [let "y" := e2 in let "x" := e1 :+ {"y"} in return (#{"x"}, #{"y"})] ^~ U. -Proof. move=> U mU; exact: letinC. Qed. +Proof. by move=> U mU; exact: letinC. Qed. End letinC. diff --git a/theories/lang_syntax_noisy.v b/theories/lang_syntax_noisy.v index 1e3eb73056..d44bbe81d4 100644 --- a/theories/lang_syntax_noisy.v +++ b/theories/lang_syntax_noisy.v @@ -23,9 +23,8 @@ From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. (* Indiana University, 2019. *) (* *) (* ``` *) -(* noisy0 == distribution of the next noisy measurement of a normally *) +(* noisyA == distribution of the next noisy measurement of a normally *) (* distributed quantity *) -(* noisy *) (* ``` *) (* *) (******************************************************************************) @@ -204,157 +203,10 @@ Qed. End normal_prob_lemmas. -Section noisy_programs. -Local Open Scope lang_scope. -Context {R : realType}. -Local Notation mu := lebesgue_measure. - -Definition exp_normal1 {R g} := @exp_normal R g _ (@oner_neq0 R). - -(* NB: exp_powR level setting is mistaken? *) -(* ((_ `^ _) * _) cannot write as (_ `^ _ * _) *) -Definition noisy0' : @exp R P [:: ("y0", Real)] Real := - [let "x" := Sample {exp_normal1 [{0}:R]} in - let "_" := Score ({expR 1} `^ - ({0}:R - (#{"y0"} - #{"x"}) ^+ {2%R} * {2^-1}:R)) - * {(Num.sqrt (2 * pi))^-1}:R in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return #{"z"}]. - -Definition noisy0 : @exp R _ [:: ("y0", Real)] _ := [Normalize {noisy0'}]. - -(* other programs from Sect. 2.3 of [Shan, POPL 2018], - nothing proved about them yet, just for the sake of completeness *) -Definition guard_real {g} str (r : R) : - @exp R P [:: (str, _) ; g] _ := - [if #{str} ==R {r}:R then return TT else Score {0}:R]. - -Definition helloWrong (y0 : R) : @exp R _ [::] _ := - [Normalize - let "x" := Sample {exp_normal1 (exp_real 0)} in - let "y" := Sample {exp_normal1 [#{"x"}]} in - let "_" := {guard_real "y" y0} in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return #{"z"}]. - -Definition helloJoint : @exp R _ [::] _ := - [Normalize - let "x" := Sample {exp_normal1 (exp_real 0)} in - let "y" := Sample {exp_normal1 [#{"x"}]} in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return (#{"y"}, #{"z"})]. - -End noisy_programs. - -(* The following section contains the mathematical facts that are used - to verify the noisy program. They are proved beforehand to - optimize the time spent by the Qed command of Rocq. *) -Section noisy_subproofs. -Local Open Scope lang_scope. +Section conjugate_normal_property. Context {R : realType}. Local Notation mu := lebesgue_measure. -Local Definition int_normal_noisy0 - (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := - \int[normal_prob 0 1]_x (fun z => - `|expR (- ((y.1 - z) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * - normal_prob z 1 V) x. - -Local Definition int_mu_noisy0 - (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := - \int[mu]_x - (`|expR (- ((y.1 - x) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * - normal_prob x 1 V * - (normal_pdf 0 1 x)%:E). - -Local Definition int_normal_noisy1 - (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := - \int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x - ((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))))%:num%:E * - normal_prob x 1 V)%E. - -Local Definition int_mu_noisy1 - (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := - \int[mu]_x - (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E * - (normal_prob x 1 V * (normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 x)%:E)). - -Local Definition int_normal_noisy2 - (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := - \int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x normal_prob x 1 V. - -Local Definition int_mu_noisy2 - (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := - normal_prob (y.1 / 2) (Num.sqrt (3 / 2)) V. - -Lemma int_normal_mu_noisy0 y V : measurable V -> - int_normal_noisy0 y V = int_mu_noisy0 y V. -Proof. -move=> mV; rewrite /int_normal_noisy0. -rewrite integral_normal_prob//. -apply: integrableMr => //. - apply: (@measurableT_comp _ _ _ _ _ _ (@normr _ R^o)) => //. - apply: measurable_funM => //. - apply: measurableT_comp => //. - apply: measurableT_comp => //. - apply: measurable_funM => //. - apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => (t ^+ 2%R)%R)) => //. - exact: measurable_funB. - exists (Num.sqrt (2 * pi))^-1%R; split; first exact: num_real. - move=> x x_gt. - move=> p _. - rewrite /= normr_id. - have /normr_idP -> : (0 <= expR (- ((y.1 - p) ^+ 2%R / 2)) / (Num.sqrt (2 * pi)))%R. - apply: mulr_ge0; first exact: expR_ge0. - by rewrite invr_ge0 sqrtr_ge0// mulr_ge0// pi_ge0. - apply/ltW. - apply: le_lt_trans x_gt. - rewrite -[leRHS]mul1r ler_pM//. - by rewrite -expR0 ler_expR oppr_le0 mulr_ge0// ?sqr_ge0// expR0 invr_ge0. -apply/integrableP; split. - exact: emeasurable_normal_prob. -apply/abse_integralP => //. - exact: emeasurable_normal_prob. -have/gee0_abs -> : 0 <= \int[normal_prob 0 1]_x normal_prob x 1 V. - exact: integral_ge0. -apply: (@le_lt_trans _ _ (\int[normal_prob 0 1]_x (cst 1%R x)%:E)). - apply: ge0_le_integral => //. - exact: emeasurable_normal_prob. - apply/measurable_EFinP. - exact: measurable_cst. - move=> x _. - exact: probability_le1. -by rewrite /= integral_cst// mul1e probability_setT ltry. -Qed. - -Lemma int_normal_mu_noisy1 y V : measurable V -> - int_normal_noisy1 y V = int_mu_noisy1 y V. -Proof. -move=> mV; rewrite /int_normal_noisy1 integral_normal_prob//. - by under eq_integral do rewrite -muleA. -apply/integrableP; split. - apply: measurable_funeM. - exact: emeasurable_normal_prob. -apply: (@le_lt_trans _ _ (\int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x - ((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))))%:num%:E - * (cst 1%R x)%:E))). -apply: ge0_le_integral; [by []|by []| | | |]. -- apply: measurableT_comp => //. - apply: measurable_funeM. - exact: emeasurable_normal_prob. -- by move=> x _; rewrite /= mule1. -- rewrite /= mule1. - exact/measurable_EFinP. -- move=> x _/=. - rewrite gee0_abs; last exact: mule_ge0. - by rewrite lee_pmul// probability_le1. -- rewrite integralZr//; last exact: finite_measure_integrable_cst. - by rewrite integral_cst// mule1 probability_setT ltry. -Qed. - -(* TODO: generalize and move outside *) -Section conjugate_normal_property. - Lemma conjugate_normal1 (m1 m2 s1 s2 : R) V : measurable V -> s1 != 0%R -> s2 != 0%R -> \int[normal_prob m1 s1]_x normal_prob (m2 + x) s2 V = @@ -461,8 +313,8 @@ Proof. by rewrite /normal_peak expr1n mul1r. Qed. (* Variable elimination and integration [Shan, Section 3.5, (9)], * also known as the reproductive property of normal distribution. *) -Lemma normal_probD (m1 s1 m2 s2: R) V : s1 != 0%R -> s2 != 0%R -> -measurable V -> +Lemma normal_probD (m1 s1 m2 s2 : R) V : s1 != 0%R -> s2 != 0%R -> + measurable V -> \int[normal_prob m1 s1]_x normal_prob (m2 + x) s2 V = normal_prob (m1 + m2) (Num.sqrt (s1 ^+ 2 + s2 ^+ 2)) V. Proof. @@ -490,10 +342,10 @@ transitivity (((Num.sqrt S1 * Num.sqrt S2 * pi *+ 2)^-1)%:E * rewrite !(sqrtrM 2) ?(@mulr_ge0 _ _ pi) ?sqr_ge0 ?pi_ge0//. rewrite !(sqrtrM pi) ?sqr_ge0//. rewrite ![in LHS]invfM. - rewrite mulrACA -(@sqrtrV _ 2)// -(expr2 (_ _^-1)). + rewrite mulrACA -(@sqrtrV _ 2)// -(expr2 (_ _^-1)%R). rewrite (@sqr_sqrtr _ 2^-1) ?invr_ge0//. rewrite mulrACA -(@sqrtrV _ pi) ?pi_ge0//. - rewrite -(expr2 (_ _^-1)) (@sqr_sqrtr _ pi^-1) ?invr_ge0// ?pi_ge0//. + rewrite -(expr2 (_ _^-1)%R) (@sqr_sqrtr _ pi^-1) ?invr_ge0// ?pi_ge0//. rewrite -!invfM; congr GRing.inv. by rewrite -[in RHS]mulr_natr (mulrC _ (Num.sqrt _)). apply: eq_integral. @@ -508,18 +360,16 @@ transitivity (((Num.sqrt S1 * Num.sqrt S2 * pi *+ 2)^-1)%:E * set DS12 := S1 + S2. set MS12 := (S1 * S2)%R. -set C := (((y * s1 ^+ 2)%R + (m1 * s2 ^+ 2)%R)%E - m2 * s1 ^+ 2) / DS12. +set C := ((((y * s1 ^+ 2)%R + (m1 * s2 ^+ 2)%R)%E - m2 * s1 ^+ 2) / DS12)%R. under eq_integral do rewrite expRD EFinM. -rewrite ge0_integralZr//=; last 3 first. - apply/measurable_EFinP. - apply: measurableT_comp => //. - apply: measurable_funM => //. - apply: measurableT_comp => //. - apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)%R) => //. - exact: measurable_funD. - by move=> z _; rewrite lee_fin ?expR_ge0. - by rewrite lee_fin expR_ge0. +rewrite ge0_integralZr//=; last first. + apply/measurable_EFinP. + apply: measurableT_comp => //. + apply: measurable_funM => //. + apply: measurableT_comp => //. + apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)%R) => //. + exact: measurable_funD. rewrite /normal_peak /normal_fun. rewrite [in RHS]EFinM. rewrite [in RHS]sqr_sqrtr//; last first. @@ -572,97 +422,143 @@ rewrite sqrtrM; last by rewrite addr_ge0 ?sqr_ge0. by rewrite mulrC. Qed. -(* -Lemma conjugate_normal (m1 m2 : R) V : measurable V -> - \int[normal_prob m1 (Num.sqrt 2)^-1]_x normal_prob (m2 + x) 1 V = - normal_prob (m1 + m2) (Num.sqrt (3 / 2)) V. -Proof. -move=> mV. -rewrite conjugate_normal1//; apply: eq_integral => x1 _. -clear V mV. -rewrite conjugate_normal2//. -rewrite normal_pdfE// /normal_pdf0. -transitivity ((2 * pi * (Num.sqrt 2)^-1 * 1)^-1%:E * - \int[mu]_x0 (expR - (- (((Num.sqrt 2)^-1 ^+ 2 + 1 ^+ 2)%R / - (((Num.sqrt 2)^-1 ^+ 2 * 1 ^+ 2) *+ 2))%R * - (x0 - (x1 * (Num.sqrt 2)^-1 ^+ 2 + m1 * 1 ^+ 2 - m2 * (Num.sqrt 2)^-1 ^+ 2) - / ((Num.sqrt 2)^-1 ^+ 2 + 1 ^+ 2)%R ) ^+ 2 - - (x1 - (m1 + m2)) ^+ 2 / ((Num.sqrt (3 / 2) ^+ 2) *+ 2)))%:E). - congr *%E. - rewrite normal_peak1. - rewrite /normal_peak. - congr EFin. - rewrite exprVn sqr_sqrtr//. - rewrite -mulr_natl mulrA divff// mul1r. - rewrite -(mulr_natl pi) sqrtrM// invfM. - rewrite invfM divr1. - rewrite -[LHS]mulrA. - rewrite -[in RHS]mulrA [in RHS]mulrC. - rewrite -(invfM (Num.sqrt pi)) -expr2. - rewrite sqr_sqrtr// ?pi_ge0//. - rewrite -{3}(@sqr_sqrtr _ 2%R)// expr2. - rewrite 2!invfM invrK. - by rewrite -mulf_div divff// mulr1 mulrC. - apply: eq_integral. - move=> z _. - rewrite -expRD. - congr EFin. - congr expR. - rewrite (_ : (Num.sqrt (3 / 2) ^+ 2) = 3 / 2)%R; last by rewrite sqr_sqrtr. - rewrite (_ : (Num.sqrt 2)^-1 ^+ 2 = 1 / 2)%R; last first. - by rewrite exprVn sqr_sqrtr// div1r. - by field. +End conjugate_normal_property. -(* gauss integral *) -under eq_integral do rewrite expRD EFinM. -rewrite ge0_integralZr//=; last first. - apply/measurable_EFinP. - apply: measurableT_comp => //. - apply: measurable_funM => //. - apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)%R) => //. - exact: measurable_funD. -rewrite /=. -rewrite /normal_peak /normal_fun. -rewrite [in RHS]sqr_sqrtr//. -rewrite -[in RHS]mulr_natr [in RHS]mulrAC divfK//. -rewrite mulNr EFinM muleA. -congr *%E; last first. - by rewrite sqr_sqrtr//. -rewrite [X in _ * X = _](_ : _ = (Num.sqrt ((1 / 3) * pi *+ 2))%:E * - \int[mu]_z (normal_pdf ((x1 + 2 * m1 - m2) / 3) (Num.sqrt (1 / 3)) z)%:E); last first. - rewrite -ge0_integralZl//=; last 2 first. - by apply/measurable_EFinP; exact: measurable_normal_pdf. - by move=> /= z _; rewrite lee_fin normal_pdf_ge0. - apply: eq_integral => /= z _. - rewrite /normal_pdf gt_eqF// /normal_pdf0 /normal_peak /normal_fun. - rewrite (@sqr_sqrtr _ (1 / 3))// -EFinM; congr EFin. - rewrite [RHS]mulrA -[LHS]mul1r; congr (_ * expR _)%R. - by rewrite divff// gt_eqF// sqrtr_gt0 pmulrn_rgt0// mulr_gt0// pi_gt0. - rewrite -(mulr_natl (1 / 3)) mul1r. - rewrite sqr_sqrtr//. - by rewrite !mulNr (mulrC (3 / 2)%R) invfM invrK (mulrC 2^-1%R). -rewrite integral_normal_pdf. -rewrite mule1 -EFinM mul1r; congr EFin. -rewrite -(mulr_natr (_ * pi)) sqrtrM ?mulr_ge0 ?pi_ge0//. -rewrite invfM mulrCA -mulrA mulVf ?mulr1 ?gt_eqF//. -rewrite !sqrtrM// invfM sqrtrV// -mulrA; congr *%R. -rewrite -[X in (_ / X)%R]sqr_sqrtr ?pi_ge0//. -by rewrite expr2 invfM mulrA divff ?div1r// gt_eqF// sqrtr_gt0 pi_gt0. +Section noisy_programs. +Local Open Scope lang_scope. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +Definition exp_normal1 {g} (e : exp D g Real) := + [Normal e {1%R} {oner_neq0 R}]. + +(* NB: exp_powR level setting is mistaken? *) +(* ((_ `^ _) * _) cannot write as (_ `^ _ * _) *) +Definition noisyA' : @exp R P [:: ("y0", Real)] Real := + [let "x" := Sample {exp_normal1 [{0}:R]} in + let "_" := Score ({expR 1} `^ + ({0}:R - (#{"y0"} - #{"x"}) ^+ {2%R} * {2^-1}:R)) + * {(Num.sqrt (2 * pi))^-1}:R in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. + +Definition noisyA : @exp R _ [:: ("y0", Real)] _ := [Normalize {noisyA'}]. + +(* other programs from Sect. 2.3 of [Shan, POPL 2018], + nothing proved about them yet, just for the sake of completeness *) +Definition guard_real {g} str (r : R) : + @exp R P [:: (str, _) ; g] _ := + [if #{str} ==R {r}:R then return TT else Score {0}:R]. + +Definition helloWrong (y0 : R) : @exp R _ [::] _ := + [Normalize + let "x" := Sample {exp_normal1 (exp_real 0)} in + let "y" := Sample {exp_normal1 [#{"x"}]} in + let "_" := {guard_real "y" y0} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. + +Definition helloJoint : @exp R _ [::] _ := + [Normalize + let "x" := Sample {exp_normal1 (exp_real 0)} in + let "y" := Sample {exp_normal1 [#{"x"}]} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return (#{"y"}, #{"z"})]. + +End noisy_programs. + +(* The following section contains the mathematical facts that are used + to verify the noisy program. They are proved beforehand as an attempt + to optimize the time spent by the Qed command of Rocq. *) +Section noisy_subproofs. +Local Open Scope lang_scope. +Context {R : realType}. +Local Notation mu := lebesgue_measure. + +Local Definition noisyA_semantics_normal + (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := + \int[normal_prob 0 1]_x (fun z => + `|expR (- ((y.1 - z) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * + normal_prob z 1 V) x. + +Lemma noisyA_semantics_normalE y V : measurable V -> + noisyA_semantics_normal y V = + \int[mu]_x + (`|expR (- ((y.1 - x) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * + normal_prob x 1 V * + (normal_pdf 0 1 x)%:E). +Proof. +move=> mV; rewrite /noisyA_semantics_normal. +rewrite integral_normal_prob//. +apply: integrableMr => //. + apply: (@measurableT_comp _ _ _ _ _ _ (@normr _ R^o)) => //. + apply: measurable_funM => //. + apply: measurableT_comp => //. + apply: measurableT_comp => //. + apply: measurable_funM => //. + apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => (t ^+ 2%R)%R)) => //. + exact: measurable_funB. + exists (Num.sqrt (2 * pi))^-1%R; split; first exact: num_real. + move=> x x_gt. + move=> p _. + rewrite /= normr_id ger0_norm; last by rewrite mulr_ge0// expR_ge0. + apply/ltW; apply: le_lt_trans x_gt. + rewrite -[leRHS]mul1r ler_pM ?expR_ge0//. + by rewrite -expR0 ler_expR oppr_le0 mulr_ge0// ?sqr_ge0// expR0 invr_ge0. +apply/integrableP; split; first exact: emeasurable_normal_prob. +apply/abse_integralP => //; first exact: emeasurable_normal_prob. +rewrite gee0_abs//; last exact: integral_ge0. +apply: (@le_lt_trans _ _ (\int[normal_prob 0 1]_x (cst 1%R x)%:E)). + apply: ge0_le_integral => //; first exact: emeasurable_normal_prob. + by apply/measurable_EFinP; exact: measurable_cst. + by move=> x _; exact: probability_le1. +by rewrite /= integral_cst// mul1e probability_setT ltry. Qed. -*) -End conjugate_normal_property. +Local Definition noisyB_semantics_normal + (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := + \int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x + ((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / + Num.sqrt (4 * pi))))%:num%:E * + normal_prob x 1 V)%E. + +Lemma noisyB_semantics_normalE y V : measurable V -> + noisyB_semantics_normal y V = + \int[mu]_x + (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E * + (normal_prob x 1 V * (normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 x)%:E)). +Proof. +move=> mV; rewrite /noisyB_semantics_normal integral_normal_prob//. + by under eq_integral do rewrite -muleA. +apply/integrableP; split. + apply: measurable_funeM. + exact: emeasurable_normal_prob. +apply: (@le_lt_trans _ _ (\int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x + ((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))))%:num%:E + * (cst 1%R x)%:E))). +apply: ge0_le_integral; [by []|by []| | | |]. +- apply: measurableT_comp => //. + apply: measurable_funeM. + exact: emeasurable_normal_prob. +- by move=> x _; rewrite /= mule1. +- rewrite /= mule1. + exact/measurable_EFinP. +- move=> x _/=. + rewrite gee0_abs; last exact: mule_ge0. + by rewrite lee_pmul// probability_le1. +- rewrite integralZr//; last exact: finite_measure_integrable_cst. + by rewrite integral_cst// mule1 probability_setT ltry. +Qed. -Lemma int_normal_mu_noisy2 y V : measurable V -> - int_normal_noisy2 y V = int_mu_noisy2 y V. +Lemma noisyC_semanticsE (y : R) V : measurable V -> + \int[normal_prob y (Num.sqrt 2)^-1]_x normal_prob x 1 V = + normal_prob y (Num.sqrt (3 / 2)) V. Proof. move=> mV. -rewrite /int_normal_noisy2. -have := (@normal_probD (y.1 / 2) (Num.sqrt 2)^-1 0 1 _ _ _ mV). +have := @normal_probD R (y ) (Num.sqrt 2)^-1 0 1 _ _ _ mV. under eq_integral do rewrite add0r. rewrite addr0. -rewrite (_: ((Num.sqrt 2)^-1 ^+ 2 + 1 ^+ 2 = 3 / 2)%R); last first. +rewrite (_ : ((Num.sqrt 2)^-1 ^+ 2 + 1 ^+ 2 = 3 / 2)%R)//; last first. rewrite exprVn sqr_sqrtr// expr1n -[in LHS]div1r -{3}(@divff _ 1%R)//. rewrite addf_div// 2!mulr1 mul1r (_:1%R = 1%:R)// -natrD. exact. @@ -679,136 +575,38 @@ Local Notation mu := lebesgue_measure. (* definition of intermediate programs *) Definition neq0Vsqrt2 : ((Num.sqrt 2)^-1 != 0 :> R)%R. Proof. exact: lt0r_neq0. Qed. -Definition exp_normal_Vsqrt2 {g} := @exp_normal R g _ neq0Vsqrt2. -Definition noisy1' : @exp R _ [:: ("y0", Real)] Real := +Definition exp_normal_Vsqrt2 {g} (e : exp D g Real) := + [Normal e {(Num.sqrt 2)^-1%R} {neq0Vsqrt2}]. + +Definition tailB : @exp R _ [:: ("_", Unit); ("y0", Real)] Real := + [let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R]} in + let "z" := Sample {exp_normal1 [#{"x"}]} in + return #{"z"}]. +Definition noisyB' : @exp R _ [:: ("y0", Real)] Real := [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2%R} * {4^-1}:R)) * {(Num.sqrt (4 * pi))^-1}:R in - let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R ]} in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return #{"z"}]. - -Definition noisy1 : @exp R _ [:: ("y0", Real)] _ := [Normalize {noisy1'}]. + {tailB}]. +Definition noisyB : @exp R _ [:: ("y0", Real)] _ := [Normalize {noisyB'}]. Definition neq0sqrt32 : (Num.sqrt (3 / 2) != 0 :> R)%R. Proof. exact: lt0r_neq0. Qed. -Definition exp_normal_sqrt32 {g} := @exp_normal R g _ neq0sqrt32. +Definition exp_normal_sqrt32 {g} (e : exp D g Real) := + [Normal e {Num.sqrt (3 / 2)} {neq0sqrt32}]. -Definition noisy2' : @exp R _ [:: ("y0", Real)] _ := +Definition tailC : @exp R _ [:: ("_", Unit); ("y0", Real)] Real := + [Sample {exp_normal_sqrt32 [#{"y0"} * {2^-1}:R]}]. +Definition noisyC' : @exp R _ [:: ("y0", Real)] _ := [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2%R} * {4^-1}:R)) * {(Num.sqrt (4 * pi))^-1}:R in - Sample {exp_normal_sqrt32 [#{"y0"} * {2^-1}:R]}]. + {tailC}]. +Definition noisyC : @exp R _ [:: ("y0", Real)] _ := [Normalize {noisyC'}]. -Definition noisy2 : @exp R _ [:: ("y0", Real)] _ := [Normalize {noisy2'}]. - -(* from (7) to (9) in [Shan, POPL 2018] *) -Lemma execP_noisy'12 y V : measurable V -> - @execP R [:: ("_", Unit); ("y0", Real)] _ - [let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R ]} in - let "z" := Sample {exp_normal1 [#{"x"}]} in - return #{"z"}] y V = - @execP R [:: ("_", Unit); ("y0", Real)] _ - [Sample {exp_normal_sqrt32 [#{"y0"} * {2^-1}:R]}] y V. +Lemma noisyAB' y V : measurable V -> execP noisyA' y V = execP noisyB' y V. Proof. move=> mV. -(* execute syntax *) -rewrite !execP_letin. -rewrite !execP_sample !execD_normal/=. -rewrite (@execD_bin _ _ binop_mult) execD_real/=. -rewrite execP_return. -rewrite !exp_var'E (execD_var_erefl "y0") (execD_var_erefl "x") (execD_var_erefl "z")/=. -rewrite !letin'E/=. -under eq_integral do rewrite letin'E/=. -rewrite /=. -(* prove semantics *) -under eq_integral do rewrite integral_normal_prob_dirac//=. -exact: (int_normal_mu_noisy2 _ mV). -Qed. - -(* semantics of noisy0 (turned into a local definition as an attempt to optimize Qed) *) -Local Definition executed_noisy0' := - (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - letin' - (sample (fun=> normal_prob 0 1) - (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) (kr 0))) - (letin' - (score - (measurable_funM - (measurableT_comp (measurable_powRr (expR 1)) - (measurable_funB (kr 0) - (measurable_funM - (measurable_funX 2%R - (measurable_funB (measurable_acc_typ [:: Real; Real] 1) - (measurable_acc_typ [:: Real; Real] 0))) - (kr 2^-1)))) (kr (Num.sqrt (2 * pi))^-1))) - (letin' - (sample - (fun - x : unit * - (g_sigma_algebraType R.-ocitv.-measurable * - (g_sigma_algebraType R.-ocitv.-measurable * unit)) => - normal_prob x.2.1 1) - (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) - (measurable_acc_typ [:: Unit; Real; Real] 1))) - (ret (measurable_acc_typ [:: Real; Unit; Real; Real] 0)))) y V : \bar R). - -(* semantics of noisy1 (turned into a local definition as a attempt to optimized Qed) *) -Local Definition executed_noisy1' := - (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - letin' - (score - (measurable_funM - (measurableT_comp (measurable_powRr (expR 1)) - (measurable_funB (kr 0) - (measurable_funM - (measurable_funX 2%R (measurable_acc_typ [:: Real] 0)) (kr 4^-1)))) - (kr (Num.sqrt (4 * pi))^-1))) - (letin' - (sample - (fun x : unit * (g_sigma_algebraType R.-ocitv.-measurable * unit) => - normal_prob (x.2.1 / 2) (Num.sqrt 2)^-1) - (measurableT_comp (measurable_normal_prob2 neq0Vsqrt2) - (measurable_funM (measurable_acc_typ [:: Unit; Real] 1) (kr 2^-1)))) - (letin' - (sample - (fun - x : g_sigma_algebraType R.-ocitv.-measurable * - (unit * (g_sigma_algebraType R.-ocitv.-measurable * unit)) => - normal_prob x.1 1) - (measurableT_comp (measurable_normal_prob2 GRing.oner_neq0) - (measurable_acc_typ [:: Real; Unit; Real] 0))) - (ret (measurable_acc_typ [:: Real; Real; Unit; Real] 0)))) y V : \bar R). - -(* -Local Definition executed_noisy1'_alter := - (fun (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) => - letin' - (score - (measurable_funM - (measurableT_comp (measurable_powRr (expR 1)) - (measurable_funB (kr 0) - (measurable_funM - (measurable_funX 2%R (measurable_acc_typ [:: Real] 0)) - (kr 4^-1)))) (kr (Num.sqrt (4 * pi))^-1))) - (letin' - (sample - (fun x : unit * (g_sigma_algebraType R.-ocitv.-measurable * unit) => - probability_normal_prob__canonical__measure_Probability - (x.2.1 / 2) (Num.sqrt 2)^-1) - (measurableT_comp (measurable_normal_prob2 neq0Vsqrt2) - (measurable_funM (measurable_acc_typ [:: Unit; Real] 1) (kr 2^-1)))) - (letin' - (sample - (fun=> probability_normal_prob__canonical__measure_Probability 0 1) - (measurableT_comp (measurable_normal_prob2 neq01) (kr 0))) - (ret - (measurable_funD (measurable_acc_typ [:: Real; Real; Unit; Real] 1) - (measurable_acc_typ [:: Real; Real; Unit; Real] 0))))) y V : \bar R). -*) - -Lemma execP_noisy0'E y V : execP noisy0' y V = executed_noisy0' y V. -Proof. -rewrite !execP_letin. +(* reduce the lhs *) +rewrite 3![in LHS]execP_letin. rewrite execP_sample. rewrite execD_normal/=. rewrite execD_real/=. @@ -822,48 +620,14 @@ rewrite execD_pow/=. rewrite (@execD_bin _ _ binop_minus)/=. rewrite exp_var'E/= (execD_var_erefl "y0")/=. rewrite exp_var'E/= (execD_var_erefl "x")/=. -rewrite execD_real/=. -rewrite execD_real/=. -rewrite execP_sample. +rewrite 2!execD_real/=. +rewrite execP_sample/=. rewrite execD_normal/=. rewrite exp_var'E/= (execD_var_erefl "x")/=. rewrite execP_return/=. rewrite exp_var'E/= (execD_var_erefl "z")/=. -done. -Qed. - -Lemma executed_noisy0'_semantics y V : measurable V -> - executed_noisy0' y V = int_normal_noisy0 y V. -Proof. -move=> mV. -rewrite /executed_noisy0'. -rewrite [in LHS]letin'E/=. -under [in LHS]eq_integral. - move=> x _. - rewrite letin'E/=. - under eq_integral. - move=> u _. - rewrite letin'E/=. - rewrite integral_normal_prob_dirac//. - over. - rewrite /=. - rewrite ge0_integral_mscale/=; first last. - - by move=> ? _. - - by []. - - exact: measurableT. - rewrite integral_dirac; first last. - - by []. - - exact: measurableT. - rewrite diracT mul1e. - rewrite sub0r. - rewrite -expRM mul1r. - over. -by rewrite /=. -Qed. - -Lemma execP_noisy1'E y V : execP noisy1' y V = executed_noisy1' y V. -Proof. -rewrite execP_letin. +(* reduce the rhs *) +rewrite [in RHS]execP_letin. rewrite execP_score. rewrite (@execD_bin _ _ binop_mult)/=. rewrite execD_pow_real/=. @@ -886,50 +650,39 @@ rewrite execD_normal/=. rewrite exp_var'E/= (execD_var_erefl "x")/=. rewrite execP_return. rewrite exp_var'E/= (execD_var_erefl "z")/=. -done. -Qed. - -Lemma executed_noisy1'_semantics y V : measurable V -> - executed_noisy1' y V = int_normal_noisy1 y V. -Proof. -move=> mV; rewrite /executed_noisy1'. -rewrite letin'E/=. -under eq_integral. - move=> u _. +(* semantics *) +transitivity (noisyA_semantics_normal y V). + rewrite [in LHS]letin'E/=. + apply: eq_integral => x _. + rewrite letin'E/=. + under eq_integral. + move=> u _. + rewrite letin'E/= integral_normal_prob_dirac//. + over. + rewrite /=. + rewrite ge0_integral_mscale/=; [|by []..]. + rewrite integral_dirac; [|by []..]. + by rewrite diracT mul1e sub0r -expRM mul1r. +transitivity (noisyB_semantics_normal y V); last first. rewrite letin'E/=. under eq_integral. - move=> x _. + move=> u _. rewrite letin'E/=. - rewrite integral_normal_prob_dirac//. + under eq_integral. + move=> x _. + rewrite letin'E/=. + rewrite integral_normal_prob_dirac//. + over. over. - over. -rewrite ge0_integral_mscale//; first last. - move=> ? _. - exact: integral_ge0. -rewrite integral_dirac//. -rewrite diracT mul1e. -rewrite sub0r -expRM mul1r. -rewrite /=. -rewrite -(@ge0_integralZl _ _ R (normal_prob _ _) _ measurableT _ _ - (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E))//. -exact: emeasurable_normal_prob. -Qed. - -Lemma noisy'01 y V : measurable V -> execP noisy0' y V = execP noisy1' y V. -Proof. -move=> mV. -(* lhs *) -rewrite execP_noisy0'E. -rewrite (executed_noisy0'_semantics _ mV). -(* rhs *) -rewrite execP_noisy1'E. -rewrite (executed_noisy1'_semantics _ mV). -(* semantics *) -rewrite (int_normal_mu_noisy0 _ mV). -rewrite (int_normal_mu_noisy1 _ mV). -(* eq_integral *) -apply: eq_integral. -move=> x _. + rewrite /= ge0_integral_mscale//; first last. + by move=> ? _; exact: integral_ge0. + rewrite integral_dirac// diracT mul1e sub0r -expRM mul1r/=. + rewrite -(@ge0_integralZl _ _ R (normal_prob _ _) _ measurableT _ _ + (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E))//. + exact: emeasurable_normal_prob. +rewrite noisyA_semantics_normalE//. +rewrite noisyB_semantics_normalE//. +apply: eq_integral => /= x _. rewrite [in LHS]muleAC. rewrite [in RHS](muleC (normal_prob x 1 V)) muleA. congr (fun t => t * normal_prob x 1 V)%E. @@ -961,36 +714,63 @@ congr Num.sqrt. lra. Qed. -Lemma noisy'12 y V : measurable V -> execP noisy1' y V = execP noisy2' y V. +(* from (7) to (9) in [Shan, POPL 2018] *) +Lemma tailBC y V : measurable V -> + @execP R [:: ("_", Unit); ("y0", Real)] _ tailB y V = + @execP R [:: ("_", Unit); ("y0", Real)] _ tailC y V. Proof. move=> mV. -rewrite /noisy1' /noisy2'. +(* execute lhs *) +rewrite 2![in LHS]execP_letin. +rewrite 2![in LHS]execP_sample. +rewrite 2!execD_normal/=. +rewrite (@execD_bin _ _ binop_mult) execD_real/=. +rewrite execP_return. +rewrite exp_var'E (execD_var_erefl "y0")/=. +rewrite exp_var'E (execD_var_erefl "x")/=. +rewrite exp_var'E (execD_var_erefl "z")/=. +rewrite ![in LHS]letin'E/=. +under eq_integral do rewrite letin'E/=. +(* execute rhs *) +rewrite [in RHS]execP_sample/=. +rewrite execD_normal/=. +rewrite (@execD_bin _ _ binop_mult) execD_real/=. +rewrite exp_var'E (execD_var_erefl "y0")/=. +(* prove semantics *) +under eq_integral do rewrite integral_normal_prob_dirac//=. +by rewrite noisyC_semanticsE. +Qed. + +Lemma noisyBC' y V : measurable V -> execP noisyB' y V = execP noisyC' y V. +Proof. +move=> mV. +rewrite /noisyB' /noisyC'. rewrite execP_letin/= [in RHS]execP_letin/=. rewrite letin'E [RHS]letin'E. -by under eq_integral do rewrite execP_noisy'12//. +by under eq_integral do rewrite tailBC//. Qed. -Lemma noisy'02 y V : measurable V -> execP noisy0' y V = execP noisy2' y V. -Proof. by move=> mV; rewrite noisy'01// noisy'12. Qed. +Lemma noisyAC' y V : measurable V -> execP noisyA' y V = execP noisyC' y V. +Proof. by move=> mV; rewrite noisyAB'// noisyBC'. Qed. End noisy_verification. -(* Trying to shows why rewriting noisy1 to noisy2 is reproductive property *) +(* Trying to show why rewriting noisyB to noisyC is reproductive property *) (* -Section rewrite noisy1'_to_variable_addition. +Section rewrite noisyB'_to_variable_addition. -Definition noisy1'_alter : @exp R _ [:: ("y0", Real)] Real := +Definition noisyB'_alter : @exp R _ [:: ("y0", Real)] Real := [let "_" := Score ({expR 1} `^ ({0}:R - #{"y0"} ^+ {2} * {4^-1}:R)) * {(Num.sqrt (4 * pi))^-1}:R in let "x" := Sample {exp_normal_Vsqrt2 [#{"y0"} * {2^-1}:R]} in let "x1" := Sample {exp_normal1 [{0}:R]} in return #{"x"} + #{"x1"}]. -Definition noisy1_alter : @exp R _ [:: ("y0", Real)] _ := - [Normalize {noisy1'_alter}]. +Definition noisyB_alter : @exp R _ [:: ("y0", Real)] _ := + [Normalize {noisyB'_alter}]. -Lemma execP_noisy1'_alterE y V : -@execP R [:: ("y0", Real)] Real noisy1'_alter y V = executed_noisy1'_alter y V. +Lemma execP_noisyB'_alterE y V : +@execP R [:: ("y0", Real)] Real noisyB'_alter y V = executed_noisyB'_alter y V. Proof. rewrite 3!execP_letin. rewrite execP_return/=. @@ -1011,13 +791,13 @@ rewrite (@execD_bin _ _ binop_mult)/=. rewrite execD_pow/=. rewrite 2!execD_real/=. rewrite 2!(exp_var'E "y0") 2!(execD_var_erefl "y0")/=. -rewrite /executed_noisy1'_alter. +rewrite /executed_noisyB'_alter. rewrite /=. Abort. -Lemma noisy1_alterE y V : measurable V -> - @execP R [:: ("y0", Real)] Real noisy1' y V = - @execP R [:: ("y0", Real)] Real noisy1'_alter y V. +Lemma noisyB_alterE y V : measurable V -> + @execP R [:: ("y0", Real)] Real noisyB' y V = + @execP R [:: ("y0", Real)] Real noisyB'_alter y V. Proof. move=> mV. rewrite 3![in RHS]execP_letin. @@ -1063,10 +843,9 @@ rewrite -ge0_integralZl//; last first. rewrite sub0r. rewrite -expRM mul1r. -rewrite execP_noisy1'E. -by rewrite executed_noisy1'_semantics. +rewrite execP_noisyB'E. +by rewrite executed_noisyB'_semantics. Abort. -End rewrite noisy1'_to_variable_addition. +End rewrite noisyB'_to_variable_addition. *) - diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v index 1ae367ceeb..0ae45e9d59 100644 --- a/theories/lang_syntax_table_game.v +++ b/theories/lang_syntax_table_game.v @@ -26,7 +26,7 @@ From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. (* 1177--1178 (2004) *) (* *) (* ``` *) -(* prog0 == Eddy's table game represented as a probabilistic program *) +(* table0 == Eddy's table game represented as a probabilistic program *) (* ``` *) (* *) (******************************************************************************) @@ -45,7 +45,10 @@ Local Open Scope string_scope. Local Open Scope ereal_scope. Local Open Scope string_scope. -Lemma letin'_sample_uniform {R : realType} d d' (T : measurableType d) +Section execP_letin_uniform. +Local Open Scope ereal_scope. + +Let letin'_sample_uniform {R : realType} d d' (T : measurableType d) (T' : measurableType d') (a b : R) (ab : (a < b)%R) (u : R.-sfker [the measurableType _ of (_ * T)%type] ~> T') x y : measurable y -> @@ -69,8 +72,8 @@ Lemma execP_letin_uniform {R : realType} (forall (p : R) x U, (0 <= p <= 1)%R -> execP s0 (p, x) U = execP s1 (p, x) U) -> forall x U, measurable U -> - execP [let str := Sample {@exp_uniform _ g 0 1 (@ltr01 R)} in {s0}] x U = - execP [let str := Sample {@exp_uniform _ g 0 1 (@ltr01 R)} in {s1}] x U. + execP [let str := Sample Uniform {0%R} {1%R} {@ltr01 R} in {s0}] x U = + execP [let str := Sample Uniform {0%R} {1%R} {@ltr01 R} in {s1}] x U. Proof. move=> s01 x U mU. rewrite !execP_letin execP_sample execD_uniform/=. @@ -80,9 +83,10 @@ apply: eq_integral => p p01. apply: s01. by rewrite inE in p01. Qed. -Local Close Scope lang_scope. -Local Close Scope ereal_scope. +End execP_letin_uniform. + +(* NB: generic lemmas about bouned, to be moved *) Section bounded. Local Open Scope ring_scope. Local Open Scope lang_scope. @@ -157,9 +161,10 @@ Qed. End bounded. +(* TODO: move? *) Lemma measurable_bernoulli_expn {R : realType} U n : measurable_fun [set: g_sigma_algebraType R.-ocitv.-measurable] - (fun x : g_sigma_algebraType (R.-ocitv.-measurable) => bernoulli_prob (`1-x ^+ n) U). + (fun x : R => bernoulli_prob (`1-x ^+ n) U). Proof. apply: (measurableT_comp (measurable_bernoulli_prob2 _)) => //=. by apply: measurable_funX => //=; exact: measurable_funB. @@ -206,9 +211,8 @@ apply: ge0_le_integral => //=. - by move=> x _; rewrite lee_fin beta_pdf_le_beta_funV. Qed. -Section game_programs. +Section table_game_programs. Local Open Scope ring_scope. -Local Open Scope ereal_scope. Local Open Scope lang_scope. Context (R : realType). Local Notation mu := lebesgue_measure. @@ -216,74 +220,67 @@ Local Notation mu := lebesgue_measure. Definition guard {g} str n : @exp R P [:: (str, _) ; g] _ := [if #{str} == {n}:N then return TT else Score {0}:R]. -Definition prog0 : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_uniform 0 1 (@ltr01 R)} in - let "x" := Sample {exp_binomial 8 [#{"p"}]} in +Definition table0 : @exp R _ [::] _ := [Normalize + let "p" := Sample Uniform {0} {1} {ltr01} in + let "x" := Sample Binomial {8} #{"p"} in let "_" := {guard "x" 5} in - let "y" := Sample {exp_binomial 3 [#{"p"}]} in + let "y" := Sample Binomial {3} #{"p"} in return {1}:N <= #{"y"}]. Definition tail1 : @exp R _ [:: ("_", Unit); ("x", Nat) ; ("p", Real)] _ := - [Sample {exp_bernoulli [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%R}]}]. + [Sample Bernoulli {1}:R - {[{1}:R - #{"p"}]} ^+ {3}]. Definition tail2 : @exp R _ [:: ("_", Unit); ("p", Real)] _ := - [Sample {exp_bernoulli [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%R}]}]. + [Sample Bernoulli {1}:R - {[{1}:R - #{"p"}]} ^+ {3}]. Definition tail3 : @exp R _ [:: ("p", Real); ("_", Unit)] _ := - [Sample {exp_bernoulli [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%R}]}]. + [Sample Bernoulli {1}:R - {[{1}:R - #{"p"}]} ^+ {3}]. -Definition prog1 : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_uniform 0 1 (@ltr01 R)} in - let "x" := Sample {exp_binomial 8 [#{"p"}]} in +Definition table1 : @exp R _ [::] _ := [Normalize + let "p" := Sample Uniform {0} {1} {ltr01} in + let "x" := Sample Binomial {8} #{"p"} in let "_" := {guard "x" 5} in {tail1}]. -Definition prog2 : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_uniform 0 1 (@ltr01 R)} in +Definition table2 : @exp R _ [::] _ := [Normalize + let "p" := Sample Uniform {0} {1} {ltr01} in let "_" := Score {[{56}:R * #{"p"} ^+ {5%R} * {[{1}:R - #{"p"}]} ^+ {3%R}]} in {tail2}]. -Definition prog2' : @exp R _ [::] _ := - [Normalize - let "p" := Sample {exp_beta 1 1} in +Definition table2' : @exp R _ [::] _ := [Normalize + let "p" := Sample Beta {1} {1} in let "_" := Score {[{56}:R * #{"p"} ^+ {5%R} * {[{1}:R - #{"p"}]} ^+ {3%R}]} in {tail2}]. -Definition prog3 : @exp R _ [::] _ := - [Normalize +Definition table3 : @exp R _ [::] _ := [Normalize let "_" := Score {1 / 9}:R in - let "p" := Sample {exp_beta 6 4} in + let "p" := Sample Beta {6} {4} in {tail3}]. -Definition prog4 : @exp R _ [::] _ := - [Normalize +Definition table4 : @exp R _ [::] _ := [Normalize let "_" := Score {1 / 9}:R in - Sample {exp_bernoulli [{10 / 11}:R]}]. + Sample Bernoulli {10 / 11}:R]. -Definition prog5 : @exp R _ [::] _ := - [Normalize Sample {exp_bernoulli [{10 / 11}:R]}]. +Definition table5 : @exp R _ [::] _ := [Normalize Sample Bernoulli {10 / 11}:R]. -End game_programs. +End table_game_programs. Arguments tail1 {R}. Arguments tail2 {R}. Arguments guard {R g}. -Section from_prog0_to_prog1. +Section from_table0_to_table1. Local Open Scope ring_scope. Local Open Scope ereal_scope. Local Open Scope lang_scope. Context (R : realType). Local Notation mu := lebesgue_measure. -Let prog01_subproof +Let table01_subproof (x : mctx (untag (ctx_of (recurse Unit (recurse Nat (found "p" Real [::])))))) U : (0 <= x.2.2.1 <= 1)%R -> - execP [let "y" := Sample {exp_binomial 3 [#{"p"}]} in + execP [let "y" := Sample Binomial {3%R} #{"p"} in return {1}:N <= #{"y"}] x U = execP (@tail1 R) x U. Proof. @@ -302,26 +299,26 @@ rewrite (execD_var_erefl "p")/=. exact/integral_binomial_prob. Qed. -Lemma prog01 : execD (@prog0 R) = execD (@prog1 R). +Lemma table01 : execD (@table0 R) = execD (@table1 R). Proof. -rewrite /prog0 /prog1. +rewrite /table0 /table1. apply: congr_normalize => y A. apply: execP_letin_uniform => // p [] B p01. apply: congr_letinr => a1 V0. apply: congr_letinr => -[] V1. -exact: prog01_subproof. +exact: table01_subproof. Qed. -End from_prog0_to_prog1. +End from_table0_to_table1. -Section from_prog1_to_prog2. +Section from_table1_to_table2. Local Open Scope ring_scope. Local Open Scope ereal_scope. Local Open Scope lang_scope. Context (R : realType). Local Notation mu := lebesgue_measure. -Let prog12_subproof (y : @mctx R [::]) (V : set (@mtyp R Bool)) +Let table12_subproof (y : @mctx R [::]) (V : set (@mtyp R Bool)) (p : R) (x : projT2 (existT measurableType default_measure_display unit)) (U : set (mtyp Bool)) @@ -369,7 +366,7 @@ rewrite !(@execD_bin _ _ binop_minus)/=. by rewrite !execD_real/= !exp_var'E/= !(execD_var_erefl "p")/=. Qed. -Lemma prog12 : execD (@prog1 R) = execD (@prog2 R). +Lemma table12 : execD (@table1 R) = execD (@table2 R). Proof. apply: congr_normalize => y V. apply: execP_letin_uniform => // p x U /andP[p0 p1]. @@ -389,10 +386,10 @@ rewrite (@execD_bin _ _ binop_minus)/=/=. rewrite 2!execD_real/=. rewrite (execD_var_erefl "p")/=. rewrite -(mulrA 56%R). -exact: prog12_subproof. +exact: table12_subproof. Qed. -End from_prog1_to_prog2. +End from_table1_to_table2. Local Open Scope ereal_scope. @@ -438,14 +435,14 @@ do 2 apply: @measurableT_comp => //=. by apply: measurable_funM => //; exact: measurable_XMonemX. Qed. -Section from_prog2_to_prog3. +Section from_table2_to_table3. Local Open Scope ring_scope. Local Open Scope ereal_scope. Local Open Scope lang_scope. Context (R : realType). Local Notation mu := lebesgue_measure. -Lemma prog22' : execD (@prog2 R) = execD (@prog2' R). +Lemma table22' : execD (@table2 R) = execD (@table2' R). Proof. apply: congr_normalize => // x U. apply: congr_letinl => // y V. @@ -453,7 +450,7 @@ rewrite !execP_sample execD_uniform/= execD_beta/=. by rewrite beta_prob_uniform. Qed. -Lemma prog23 : execD (@prog2' R) = execD (@prog3 R). +Lemma table23 : execD (@table2' R) = execD (@table3 R). Proof. apply: congr_normalize => x U. (* reduce the LHS *) @@ -519,7 +516,7 @@ rewrite /= XMonemX00 mul1r [in LHS](mulrC 56%R) [in LHS]EFinM -[in LHS]muleA; co by rewrite !beta_funE/=; repeat rewrite !factE/=; rewrite -EFinM; congr EFin; lra. Qed. -End from_prog2_to_prog3. +End from_table2_to_table3. Local Open Scope ereal_scope. (* TODO: move? *) @@ -583,7 +580,8 @@ Qed. Lemma int_beta_prob_bernoulli_onem {R : realType} (U : set (@mtyp R Bool)) : \int[beta_prob 6 4]_y bernoulli_prob (`1-(`1-y ^+ 3)) U = bernoulli_prob (10 / 11) U :> \bar R. Proof. -transitivity (\d_false U + \d_true U - bernoulli_prob (1 / 11) U : \bar R)%E; last first. +transitivity + (\d_false U + \d_true U - bernoulli_prob (1 / 11) U : \bar R)%E; last first. rewrite /bernoulli_prob ifT; last lra. rewrite ifT; last lra. apply/eqP; rewrite sube_eq//; last first. @@ -635,16 +633,16 @@ Qed. Local Close Scope ereal_scope. -Section from_prog3_to_prog4. +Section from_table3_to_table4. Local Open Scope ereal_scope. Local Open Scope lang_scope. Context (R : realType). (* NB: not used *) -Lemma prog34' U : - @execP R [::] _ [let "p" := Sample {exp_beta 6 4} in - Sample {exp_bernoulli [{[{1}:R - #{"p"}]} ^+ {3%N}]}] tt U = - @execP R [::] _ [Sample {exp_bernoulli [{1 / 11}:R]}] tt U. +Lemma table34' U : + @execP R [::] _ [let "p" := Sample Beta {6%R} {4%R} in + Sample Bernoulli {[{1}:R - #{"p"}]} ^+ {3%N}] tt U = + @execP R [::] _ [Sample Bernoulli {1 / 11}:R] tt U. Proof. (* reduce the lhs *) rewrite execP_letin. @@ -659,10 +657,10 @@ rewrite letin'E/=. exact: int_beta_prob_bernoulli. Qed. -Lemma prog34 l u U : - @execP R l _ [let "p" := Sample {exp_beta 6 4} in - Sample {exp_bernoulli [{1}:R - {[{1}:R - #{"p"}]} ^+ {3%N}]}] u U = - @execP R l _ [Sample {exp_bernoulli [{10 / 11}:R]}] u U. +Lemma table34 l u U : + @execP R l _ [let "p" := Sample Beta {6%R} {4%R} in + Sample Bernoulli {1}:R - {[{1}:R - #{"p"}]} ^+ {3%N}] u U = + @execP R l _ [Sample Bernoulli {10 / 11}:R] u U. Proof. (* reduce the lhs *) rewrite execP_letin. @@ -678,9 +676,9 @@ rewrite letin'E/=. exact: int_beta_prob_bernoulli_onem. Qed. -End from_prog3_to_prog4. +End from_table3_to_table4. -Section from_prog4_to_prog5. +Section from_table4_to_table5. Local Open Scope ring_scope. Local Open Scope ereal_scope. Local Open Scope lang_scope. @@ -689,8 +687,8 @@ Local Notation mu := lebesgue_measure. Lemma normalize_score_bernoulli g p q (p0 : (0 < p)%R) (q01 : (0 <= q <= 1)%R) : @execD R g _ [Normalize let "_" := Score {p}:R in - Sample {exp_bernoulli [{q}:R]}] = - execD [Normalize Sample {exp_bernoulli [{q}:R]}]. + Sample Bernoulli {q}:R] = + execD [Normalize Sample Bernoulli {q}:R]. Proof. apply: eq_execD. rewrite !execD_normalize_pt/= !execP_letin !execP_score. @@ -714,16 +712,16 @@ rewrite integral_dirac//= diracT mul1e. by rewrite muleAC -EFinM divff ?gt_eqF// mul1e bernoulli_probE. Qed. -Lemma prog45 : execD (@prog4 R) = execD (@prog5 R). +Lemma table45 : execD (@table4 R) = execD (@table5 R). Proof. by rewrite normalize_score_bernoulli//; lra. Qed. -End from_prog4_to_prog5. +End from_table4_to_table5. -Lemma from_prog0_to_prog5 {R : realType} : execD (@prog0 R) = execD (@prog5 R). +Lemma from_table0_to_table5 {R : realType} : execD (@table0 R) = execD (@table5 R). Proof. -rewrite prog01 prog12 prog22' prog23. -rewrite -prog45. +rewrite table01 table12 table22' table23. +rewrite -table45. apply: congr_normalize => y V. apply: congr_letinr => x U. -by rewrite -prog34. +by rewrite -table34. Qed. diff --git a/theories/prob_lang.v b/theories/prob_lang.v index d0a9370569..c1893a8dc1 100644 --- a/theories/prob_lang.v +++ b/theories/prob_lang.v @@ -1575,7 +1575,7 @@ End letinC. (* examples *) -Lemma letin_sample_bernoulli d d' (T : measurableType d) +Lemma letin_sample_bernoulli_prob d d' (T : measurableType d) (T' : measurableType d') (R : realType) (r : R) (u : R.-sfker [the measurableType _ of (T * bool)%type] ~> T') x y : (0 <= r <= 1)%R -> @@ -1595,7 +1595,7 @@ Definition sample_and_return : R.-sfker T ~> _ := Lemma sample_and_returnE t U : sample_and_return t U = (2 / 7%:R)%:E * \d_true U + (5%:R / 7%:R)%:E * \d_false U. Proof. -rewrite /sample_and_return letin_sample_bernoulli; last lra. +rewrite /sample_and_return letin_sample_bernoulli_prob; last lra. by rewrite !retE onem27. Qed. @@ -1616,7 +1616,7 @@ Definition sample_and_branch : R.-sfker T ~> _ := Lemma sample_and_branchE t U : sample_and_branch t U = (2 / 7)%:E * \d_(3%R : R) U + (5 / 7)%:E * \d_(10%R : R) U. Proof. -rewrite /sample_and_branch letin_sample_bernoulli/=; last lra. +rewrite /sample_and_branch letin_sample_bernoulli_prob/=; last lra. by rewrite !iteE/= onem27. Qed. @@ -1635,9 +1635,9 @@ Lemma bernoulli_andE t U : bernoulli_and t U = sample_cst (bernoulli_prob (1 / 4)) t U. Proof. rewrite /bernoulli_and. -rewrite letin_sample_bernoulli; last lra. -rewrite (letin_sample_bernoulli (r := 1 / 2)); last lra. -rewrite (letin_sample_bernoulli (r := 1 / 2)); last lra. +rewrite letin_sample_bernoulli_prob; last lra. +rewrite (letin_sample_bernoulli_prob (r := 1 / 2)); last lra. +rewrite (letin_sample_bernoulli_prob (r := 1 / 2)); last lra. rewrite muleDr//= -muleDl//. rewrite !muleA -addeA -muleDl// -!EFinM !onem1S/= -splitr mulr1. have -> : (1 / 2 * (1 / 2) = 1 / 4%:R :> R)%R by rewrite mulf_div mulr1// -natrM. @@ -1682,7 +1682,7 @@ Let kstaton_bus_poissonE t U : kstaton_bus_poisson t U = (5 / 7)%:E * (poisson4 10)%:E * \d_false U. Proof. rewrite /kstaton_bus_poisson /kstaton_bus. -rewrite letin_sample_bernoulli; last lra. +rewrite letin_sample_bernoulli_prob; last lra. rewrite -!muleA; congr (_ * _ + _ * _). - rewrite letin_kret//. rewrite letin_iteT//. @@ -1732,7 +1732,7 @@ Let kstaton_bus_exponentialE t U : kstaton_bus_exponential t U = (5 / 7)%:E * (exp1560 10)%:E * \d_false U. Proof. rewrite /kstaton_bus. -rewrite letin_sample_bernoulli; last lra. +rewrite letin_sample_bernoulli_prob; last lra. rewrite -!muleA; congr (_ * _ + _ * _). - rewrite letin_kret//. rewrite letin_iteT//. @@ -1814,7 +1814,7 @@ Lemma trickE gamma X : trick gamma X = (r *+ (inl tt \in X) + q *+ ((inr true \in X) + (inr false \in X)))%:E. Proof. -have Dbernoulli : D =1 bernoulli_prob p by exact/eq_bernoulli/Dtrue. +have Dbernoulli : D =1 bernoulli_prob p by exact/eq_bernoulli_prob/Dtrue. have p_itv01 : (0 <= p <= 1)%R. by rewrite -2!lee_fin -Dtrue?measure_ge0 ?probability_le1. pose eqbern := eq_measure_integral _ (fun x _ _ => Dbernoulli x). @@ -1882,7 +1882,7 @@ HB.instance Definition _ := Kernel_isProbability.Build _ _ _ _ _ Theorem von_neumann_trickP gamma : von_neumann_trick gamma =1 bernoulli_prob 2^-1. -Proof. by apply: eq_bernoulli; rewrite von_neumann_trick_prob_kernel. Qed. +Proof. by apply: eq_bernoulli_prob; rewrite von_neumann_trick_prob_kernel. Qed. End von_neumann_trick_proof. @@ -2080,12 +2080,12 @@ Qed. End letin'A. -Lemma letin'_sample_bernoulli d d' (T : measurableType d) +Lemma letin'_sample_bernoulli_prob d d' (T : measurableType d) (T' : measurableType d') (R : realType) (r : R) (r01 : (0 <= r <= 1)%R) (u : R.-sfker bool * T ~> T') x y : letin' (sample_cst (bernoulli_prob r)) u x y = r%:E * u (true, x) y + (`1- r)%:E * u (false, x) y. -Proof. by rewrite letin'_letin letin_sample_bernoulli. Qed. +Proof. by rewrite letin'_letin letin_sample_bernoulli_prob. Qed. Section letin'_return. Context d d' d3 (X : measurableType d) (Y : measurableType d') diff --git a/theories/probability.v b/theories/probability.v index dfd92694c0..c40cedc975 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -1168,7 +1168,7 @@ Definition bernoulli_prob {R : realType} (p : R) : set bool -> \bar R := else \d_false A. -Section bernoulli. +Section bernoulli_prob. Context {R : realType} (p : R). Local Notation bernoulli := (bernoulli_prob p). @@ -1214,7 +1214,7 @@ Qed. HB.instance Definition _ := @Measure_isProbability.Build _ _ R bernoulli bernoulli_setT. -Lemma eq_bernoulli (P : probability bool R) : +Lemma eq_bernoulli_prob (P : probability bool R) : P [set true] = p%:E -> P =1 bernoulli. Proof. move=> Ptrue sb; rewrite /bernoulli /bernoulli_pmf. @@ -1231,7 +1231,7 @@ rewrite -[in LHS](eq_sb sb)/= measure_fin_bigcup//; last 2 first. - by apply: eq_fsbigr => /= -[]. Qed. -End bernoulli. +End bernoulli_prob. Section bernoulli_measure. Context {R : realType}. @@ -1263,15 +1263,15 @@ Arguments bernoulli_prob {R}. Lemma eq_bernoulliV2 {R : realType} (P : probability bool R) : P [set true] = P [set false] -> P =1 bernoulli_prob 2^-1. Proof. -move=> Ptrue_eq_false; apply/eq_bernoulli. -have : P [set: bool] = 1%E := probability_setT P. +move=> Ptrue_eq_false; apply/eq_bernoulli_prob. +have : P [set: bool] = 1%E := probability_setT _. rewrite setT_bool measureU//=; last first. by rewrite disjoints_subset => -[]//. rewrite Ptrue_eq_false -mule2n; move/esym/eqP. by rewrite -mule_natl -eqe_pdivrMl// mule1 => /eqP<-. Qed. -Section integral_bernoulli. +Section integral_bernoulli_prob. Context {R : realType}. Variables (p : R) (p01 : (0 <= p <= 1)%R). Local Open Scope ereal_scope. @@ -1288,9 +1288,9 @@ rewrite ge0_integral_measure_sum// 2!big_ord_recl/= big_ord0 adde0/=. by rewrite !ge0_integral_mscale//= !integral_dirac//= !diracT !mul1e. Qed. -End integral_bernoulli. +End integral_bernoulli_prob. -Section measurable_bernoulli. +Section measurable_bernoulli_prob. Local Open Scope ring_scope. Variable R : realType. Implicit Type p : R. @@ -1317,7 +1317,7 @@ move=> mU. exact: (measurable_kernel (kprobability measurable_bernoulli_prob)). Qed. -End measurable_bernoulli. +End measurable_bernoulli_prob. Arguments measurable_bernoulli_prob {R}. Section binomial_pmf. @@ -3264,9 +3264,9 @@ by rewrite add0r derive_id derive_cst addr0 scaler1. Qed. Let int_gaussFE : sigma != 0 -> - (\int[mu]_x ((((gauss_fun \o F) * - (F^`())%classic) x)%:E * (Num.sqrt (sigma ^+ 2 *+ 2))%:E))%E = -(Num.sqrt (sigma ^+ 2 * pi *+ 2))%:E. + (\int[mu]_x ((((gauss_fun \o F) * + (F^`())%classic) x)%:E * (Num.sqrt (sigma ^+ 2 *+ 2))%:E))%E = + (Num.sqrt (sigma ^+ 2 * pi *+ 2))%:E. Proof. move=> s0. rewrite -mulrnAr -[in RHS]mulr_natr sqrtrM ?(sqrtrM 2) ?sqr_ge0 ?pi_ge0// !EFinM. @@ -3354,35 +3354,19 @@ Qed. End normal_probability. -(* TODO: PR *) +(* TODO: move *) Section shift_properties. - Variable R : realType. Local Open Scope ring_scope. Notation mu := lebesgue_measure. -(* PR in progress *) -Lemma derive_shift (v k : R) : 'D_v (shift k : R^o -> R^o) = cst v. -Proof. -apply/funext => x/=. -by rewrite deriveD// derive_id derive_cst addr0. -Qed. - -(* PR in progress *) -Lemma is_derive_shift x v (k : R) : is_derive x v (shift k : R^o -> R^o) v. -Proof. -split => //. -by rewrite derive_val addr0. -Qed. - -(* TODO: In integration_by_substitution, (f : R -> R) => (f : R -> \bar R) *) Lemma ge0_integration_by_substitution_shift_itvNy (f : R -> R) (r e : R) : -{within `]-oo, r + e], continuous f} -> -{in `]-oo, r + e[, forall x : R, 0 <= f x} -> -(\int[mu]_(x in `]-oo, (r + e)%R]) (f x)%:E = -\int[mu]_(x in `]-oo, r]) ((f \o (shift e)) x)%:E)%E. + {within `]-oo, r + e], continuous f} -> + {in `]-oo, r + e[, forall x : R, 0 <= f x} -> + (\int[mu]_(x in `]-oo, (r + e)%R]) (f x)%:E = + \int[mu]_(x in `]-oo, r]) ((f \o (shift e)) x)%:E)%E. Proof. move=> cf f0. have := (derive_shift 1 e). @@ -3407,10 +3391,10 @@ by rewrite dshiftE mulr1/=. Qed. Lemma ge0_integration_by_substitution_shift_itvy (f : R -> R) (r e : R) : -{within `[r + e, +oo[, continuous f} -> -{in `]r + e, +oo[, forall x : R, 0 <= f x} -> -(\int[mu]_(x in `[r + e, +oo[) (f x)%:E = -\int[mu]_(x in `[r, +oo[) ((f \o (shift e)) x)%:E)%E. + {within `[r + e, +oo[, continuous f} -> + {in `]r + e, +oo[, forall x : R, 0 <= f x} -> + (\int[mu]_(x in `[r + e, +oo[) (f x)%:E = + \int[mu]_(x in `[r, +oo[) ((f \o (shift e)) x)%:E)%E. Proof. move=> cf f0. have := (derive_shift 1 e). @@ -4069,54 +4053,32 @@ Qed. End normal_kernel. -Section normal_probability. -Context {R : realType}. - -Lemma measurable_normal_s_prob (s : R) : - measurable_fun setT - (fun m : R => normal_prob m s : pprobability _ _). -Proof. -apply: (@measurability _ _ _ _ _ _ - (@pset _ _ _ : set (set (pprobability _ R)))) => //. -move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. -Abort. -(* -have := measurable_normal_prob2 (integral_normal^~ s) mYs. -Qed. -*) - -End normal_probability. - -Section dirac_delta. -Local Open Scope ereal_scope. -Context {R: realType}. -Local Notation mu := lebesgue_measure. - -Hypothesis integral_normal : forall (m : R) (s : {posnum R}), - (\int[@lebesgue_measure R]_x (normal_pdf m s%:num x)%:E = 1%E)%E. - -Definition dirac_delta (m x : R) : \bar R := lim ((normal_pdf m s x)%:E @[s --> 0^'+]). - -Lemma integralT_dirac_delta m : - \int[mu]_x dirac_delta m x = 1. -Proof. -rewrite [LHS](_:_= lim ((\int[mu]_x (normal_pdf m s x)%:E) @[s --> 0^'+])); last first. - rewrite /dirac_delta. - (* rewrite monotone_convergence *) - admit. -apply: lim_near_cst => //. -near=> s. -have [s0 s0s]: exists s' : {posnum R}, s'%:num = s. - have s0 : (0 < s)%R by []. - by exists (PosNum s0). -rewrite -s0s. -apply: integral_normal. -Abort. - -Lemma dirac_deltaE x A : - \int[mu]_(y in A) dirac_delta x y = \d_x A. -Proof. -rewrite diracE. -Abort. - -End dirac_delta. +Section move. + +Lemma cvg_comp_filter {R : realType} (f g : R -> R) (r l : R) : + continuous f -> + (f \o g) x @[x --> r] --> l -> + f x @[x --> g r] --> l. +Proof. +move=> cf fgrl. +apply/(@cvgrPdist_le _ R^o) => /= e e0. +have e20 : 0 < e / 2 by rewrite divr_gt0. +move/(@cvgrPdist_le _ R^o) : fgrl => /(_ _ e20) fgrl. +have := cf (g r). +move=> /(@cvgrPdist_le _ R^o) => /(_ _ e20)[x x0]H. +exists (minr x (e/2)). + by rewrite lt_min x0. +move=> z. +rewrite /ball_ /= => grze. +rewrite -[X in X - _](subrK (f (g r))). +rewrite -(addrA _ _ (- f z)). +apply: (le_trans (ler_normD _ _)). +rewrite (splitr e) lerD//. + case: fgrl => d /= d0 K. + apply: K. + by rewrite /ball_/= subrr normr0. +apply: H => /=. +by rewrite (lt_le_trans grze)// ge_min lexx. +Qed. + +End move. From 2622724fd19bda40265874a4598fe1076916833e Mon Sep 17 00:00:00 2001 From: IshYosh <103252572+IshiguroYoshihiro@users.noreply.github.com> Date: Tue, 25 Mar 2025 16:17:57 +0900 Subject: [PATCH 41/48] exponential_prob and noisyAB'_rearrange (#35) * noisyAB'_rearrange * add exponential_prob --- theories/lang_syntax_noisy.v | 130 +++++++++++++++++------------------ theories/prob_lang_wip.v | 62 ----------------- theories/probability.v | 27 +++++++- 3 files changed, 90 insertions(+), 129 deletions(-) diff --git a/theories/lang_syntax_noisy.v b/theories/lang_syntax_noisy.v index d44bbe81d4..1deeea96b3 100644 --- a/theories/lang_syntax_noisy.v +++ b/theories/lang_syntax_noisy.v @@ -207,6 +207,11 @@ Section conjugate_normal_property. Context {R : realType}. Local Notation mu := lebesgue_measure. +Let normal_pdf0 m s x : R := normal_peak s * normal_fun m s x. + +Let measurable_normal_pdf0 m s : measurable_fun setT (normal_pdf0 m s). +Proof. by apply: measurable_funM => //=; exact: measurable_normal_fun. Qed. + Lemma conjugate_normal1 (m1 m2 s1 s2 : R) V : measurable V -> s1 != 0%R -> s2 != 0%R -> \int[normal_prob m1 s1]_x normal_prob (m2 + x) s2 V = @@ -287,23 +292,14 @@ Proof. move=> s10 s20. rewrite -ge0_integralZl//=; last 3 first. - apply/measurable_EFinP => //=; apply: measurable_funM => //=. - + apply: measurableT_comp => //=; apply: measurable_funM => //=. - apply/measurableT_comp => //; apply: measurable_funX. - under eq_fun do rewrite opprD. - apply: measurable_funD => //=. - exact: measurable_funB. - + apply: measurableT_comp => //=. - apply: measurable_funM => //. - apply: measurableT_comp => //. - by apply: measurable_funX; exact: measurable_funD. + + rewrite /normal_fun. + under eq_fun do rewrite -(sqrrN (y - _)) opprB (addrC m1) -addrA -opprB. + exact: measurable_normal_fun. + + exact: measurable_normal_fun. - by move=> z _; rewrite lee_fin mulr_ge0// expR_ge0. - by rewrite lee_fin mulr_ge0// ?normal_peak_ge0. apply: eq_integral => /= z _. -rewrite /normal_pdf (negbTE s10). -rewrite (negbTE s20). -rewrite /normal_pdf0. -rewrite mulrACA. -rewrite /normal_fun. +rewrite 2?normal_pdfE// /normal_pdf0 mulrACA /normal_fun. by congr EFin. Qed. @@ -478,30 +474,27 @@ Local Notation mu := lebesgue_measure. Local Definition noisyA_semantics_normal (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := \int[normal_prob 0 1]_x (fun z => - `|expR (- ((y.1 - z) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * + (expR (- ((y.1 - z) ^+ 2%R / 2)) / Num.sqrt (2 * pi))%:E * normal_prob z 1 V) x. Lemma noisyA_semantics_normalE y V : measurable V -> noisyA_semantics_normal y V = \int[mu]_x - (`|expR (- ((y.1 - x) ^+ 2%R / 2)) / Num.sqrt (2 * pi)|%:E * + ((expR (- ((y.1 - x) ^+ 2%R / 2)) / Num.sqrt (2 * pi))%:E * normal_prob x 1 V * (normal_pdf 0 1 x)%:E). Proof. move=> mV; rewrite /noisyA_semantics_normal. rewrite integral_normal_prob//. apply: integrableMr => //. - apply: (@measurableT_comp _ _ _ _ _ _ (@normr _ R^o)) => //. apply: measurable_funM => //. - apply: measurableT_comp => //. - apply: measurableT_comp => //. - apply: measurable_funM => //. - apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => (t ^+ 2%R)%R)) => //. - exact: measurable_funB. + under eq_fun do rewrite -sqrrN opprB -mulNr. + rewrite [X in fun _ => expR (_ / X)%R](_:2 = 1 ^+ 2 *+ 2)%R; last first. + by rewrite expr1n. + exact: measurable_normal_fun. exists (Num.sqrt (2 * pi))^-1%R; split; first exact: num_real. - move=> x x_gt. - move=> p _. - rewrite /= normr_id ger0_norm; last by rewrite mulr_ge0// expR_ge0. + move=> x x_gt p _. + rewrite /= ger0_norm; last by rewrite mulr_ge0// expR_ge0. apply/ltW; apply: le_lt_trans x_gt. rewrite -[leRHS]mul1r ler_pM ?expR_ge0//. by rewrite -expR0 ler_expR oppr_le0 mulr_ge0// ?sqr_ge0// expR0 invr_ge0. @@ -518,14 +511,14 @@ Qed. Local Definition noisyB_semantics_normal (y : (@mctx R [:: ("y0", Real)])) (V : set (@mtyp R Real)) := \int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x - ((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / - Num.sqrt (4 * pi))))%:num%:E * + ((expR (- (y.1 ^+ 2%R / 4)) / + Num.sqrt (4 * pi))%:E * normal_prob x 1 V)%E. Lemma noisyB_semantics_normalE y V : measurable V -> noisyB_semantics_normal y V = \int[mu]_x - (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E * + ((expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))%:E * (normal_prob x 1 V * (normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 x)%:E)). Proof. move=> mV; rewrite /noisyB_semantics_normal integral_normal_prob//. @@ -544,12 +537,41 @@ apply: ge0_le_integral; [by []|by []| | | |]. - rewrite /= mule1. exact/measurable_EFinP. - move=> x _/=. - rewrite gee0_abs; last exact: mule_ge0. - by rewrite lee_pmul// probability_le1. + rewrite abseM/= lee_pmul//. + rewrite gee0_abs; last exact: measure_ge0. + by rewrite probability_le1. - rewrite integralZr//; last exact: finite_measure_integrable_cst. by rewrite integral_cst// mule1 probability_setT ltry. Qed. +Local Definition noisyA'_part + (y : (@mctx R [:: ("y0", Real)])) (x : R) := + ((expR (- ((y.1 - x) ^+ 2%R / 2)) / Num.sqrt (2 * pi)) * (normal_pdf 0 1 x))%R. + +Local Definition noisyB'_part + (y : (@mctx R [:: ("y0", Real)])) (x : R) := + ((expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)) * (normal_pdf (y.1 / 2) (Num.sqrt 2)^-1 x))%R. + +Lemma noisyAB'_rearrange (y : (@mctx R [:: ("y0", Real)])) x : + noisyA'_part y x = noisyB'_part y x. +Proof. +rewrite /noisyA'_part/noisyB'_part !normal_pdfE//. +rewrite mulrA mulrAC -(@sqrtrV _ 2)//. +rewrite /normal_peak sqr_sqrtr; last by rewrite invr_ge0. +rewrite /normal_fun subr0 sqr_sqrtr; last by rewrite invr_ge0. +rewrite -mulrA mulrAC mulrA. +rewrite [X in (X / _ / _)%R = _](_ : _ = expR (- x ^+ 2 / (1 ^+ 2 *+ 2) - (y.1 - x) ^+ 2%R / 2)); last first. + by rewrite mulrC -expRD -mulrA. +rewrite [RHS]mulrAC [X in _ = (_ * X / _)%R]mulrC mulrA -mulrA -[RHS]mulrA -expRD -2!invfM. +congr ((expR _) * _^-1)%R. + lra. +rewrite -2?sqrtrM; last 2 first. +- by rewrite -mulr_natr mulrAC mulVf// mul1r pi_ge0. +- by rewrite expr1n mul1r mulrn_wge0// pi_ge0. +congr Num.sqrt. +lra. +Qed. + Lemma noisyC_semanticsE (y : R) V : measurable V -> \int[normal_prob y (Num.sqrt 2)^-1]_x normal_prob x 1 V = normal_prob y (Num.sqrt (3 / 2)) V. @@ -564,6 +586,7 @@ rewrite (_ : ((Num.sqrt 2)^-1 ^+ 2 + 1 ^+ 2 = 3 / 2)%R)//; last first. exact. Qed. + End noisy_subproofs. (* this section contains the verification of the noisy program per se *) @@ -659,10 +682,9 @@ transitivity (noisyA_semantics_normal y V). move=> u _. rewrite letin'E/= integral_normal_prob_dirac//. over. - rewrite /=. - rewrite ge0_integral_mscale/=; [|by []..]. - rewrite integral_dirac; [|by []..]. - by rewrite diracT mul1e sub0r -expRM mul1r. + rewrite /= ge0_integral_mscale//. + rewrite integral_dirac// diracT mul1e sub0r -expRM mul1r/=. + by rewrite ger0_norm; last by rewrite mulr_ge0// expR_ge0. transitivity (noisyB_semantics_normal y V); last first. rewrite letin'E/=. under eq_integral. @@ -677,41 +699,19 @@ transitivity (noisyB_semantics_normal y V); last first. rewrite /= ge0_integral_mscale//; first last. by move=> ? _; exact: integral_ge0. rewrite integral_dirac// diracT mul1e sub0r -expRM mul1r/=. + rewrite ger0_norm; last by rewrite mulr_ge0// expR_ge0. rewrite -(@ge0_integralZl _ _ R (normal_prob _ _) _ measurableT _ _ - (`|expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi)|%:E))//. + (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))%:E)//. exact: emeasurable_normal_prob. rewrite noisyA_semantics_normalE//. rewrite noisyB_semantics_normalE//. -apply: eq_integral => /= x _. -rewrite [in LHS]muleAC. -rewrite [in RHS](muleC (normal_prob x 1 V)) muleA. -congr (fun t => t * normal_prob x 1 V)%E. -rewrite [in LHS]ger0_norm; last by rewrite mulr_ge0// expR_ge0. -rewrite [in RHS]ger0_norm; last by rewrite mulr_ge0// expR_ge0. -rewrite -!EFinM; congr EFin. -rewrite !normal_pdfE// /normal_pdf0. -rewrite mulrA. -rewrite mulrAC. -rewrite -(@sqrtrV _ 2)//. -rewrite /normal_peak. -rewrite sqr_sqrtr; last by rewrite invr_ge0. -rewrite /normal_fun. -rewrite subr0. -rewrite sqr_sqrtr; last by rewrite invr_ge0. -rewrite [X in (X / _)%R = _]mulrC mulrA -expRD -mulrA -invfM. -rewrite [RHS]mulrAC [X in _ = (_ * X / _)%R]mulrC mulrA -mulrA -expRD -invfM. -congr *%R. - congr expR. - lra. -congr GRing.inv. -rewrite expr1n mul1r -mulrnAr -(mulr_natl pi) mulrA mulVf// mul1r -expr2. -rewrite sqr_sqrtr; last by rewrite mulr_ge0// pi_ge0. -rewrite !sqrtrM// mulrCA -expr2 sqr_sqrtr; last exact: pi_ge0. -congr *%R. -rewrite -[LHS]gtr0_norm//. -rewrite -(sqrtr_sqr 2%R). -congr Num.sqrt. -lra. +apply: eq_integral => x _. +transitivity ((noisyA'_part y x)%:E * normal_prob x 1 V). + by rewrite muleAC. +transitivity ((noisyB'_part y x)%:E * normal_prob x 1 V); last first. + by rewrite (muleC (normal_prob x 1 V)) muleA. +congr (fun t => t%:E * normal_prob x 1 V)%E. +exact: noisyAB'_rearrange. Qed. (* from (7) to (9) in [Shan, POPL 2018] *) diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 4dac4d1814..0bb9b0a6a0 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -114,65 +114,3 @@ Definition staton_counting : R.-sfker X ~> _ := (ret macc1of3)). End staton_counting. - -Section exponential_pdf. -Context {R : realType}. -Notation mu := lebesgue_measure. -Variable (mean : R). -Hypothesis mean0 : (0 < mean)%R. - -Definition exponential_pdf' (x : R) := (mean^-1 * expR (- x / mean))%R. -Definition exponential_pdf := exponential_pdf' \_ `[0%R, +oo[. - -Lemma exponential_pdf_ge0 (x : R) : (0 <= exponential_pdf x)%R. -Proof. -apply: restrict_ge0 => {}x _. -apply: mulr_ge0; last exact: expR_ge0. -by rewrite invr_ge0 ltW. -Qed. - -End exponential_pdf. - -Definition exponential {R : realType} (m : R) - of \int[@lebesgue_measure R]_x (exponential_pdf m x)%:E = 1%E : set _ -> \bar R := - fun V => (\int[lebesgue_measure]_(x in V) (exponential_pdf m x)%:E)%E. - -Section exponential. -Context {R : realType}. -Local Open Scope ring_scope. - -Notation mu := lebesgue_measure. -Variable (mean : R). -Hypothesis mean0 : (0 < mean)%R. - -Hypothesis integrable_exponential : forall (m : R), (0 < m)%R -> - mu.-integrable [set: _] (EFin \o (exponential_pdf m)). - -Hypothesis integral_exponential_pdf : forall (m : R), (0 < m)%R -> - (\int[mu]_x (exponential_pdf m x)%:E = 1)%E. - -Local Notation exponential := (exponential (integral_exponential_pdf mean0)). - -Let exponential0 : exponential set0 = 0%E. -Proof. by rewrite /exponential integral_set0. Qed. - -Let exponential_ge0 A : (0 <= exponential A)%E. -Proof. -rewrite /exponential integral_ge0//= => x _. -by rewrite lee_fin exponential_pdf_ge0. -Qed. - -Let exponential_sigma_additive : semi_sigma_additive exponential. -Proof. -Admitted. - -HB.instance Definition _ := isMeasure.Build _ _ _ - exponential exponential0 exponential_ge0 exponential_sigma_additive. - -Let exponential_setT : exponential [set: _] = 1%E. -Proof. by rewrite /exponential integral_exponential_pdf. Qed. - -HB.instance Definition _ := - @Measure_isProbability.Build _ _ R exponential exponential_setT. - -End exponential. diff --git a/theories/probability.v b/theories/probability.v index c40cedc975..47b4eb3c8a 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -3743,6 +3743,29 @@ Admitted. added by the integral of `normal_pdf a s x` on ]a - e, a + e[ *) +Let normal_pdf0 m s x : R := normal_peak s * normal_fun m s x. + +Let normal_pdf0_ge0 m x : 0 <= normal_pdf0 m s x. +Proof. by rewrite mulr_ge0 ?normal_peak_ge0 ?expR_ge0. Qed. + +Let continuous_normal_pdf0 m : continuous (normal_pdf0 m s). +Proof. +move=> x; apply: cvgM; first exact: cvg_cst. +apply: (@cvg_comp _ R^o _ _ _ _ + (nbhs (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)))); last exact: continuous_expR. +apply: cvgM; last exact: cvg_cst; apply: (@cvgN _ R^o). +apply: (@cvg_comp _ _ _ _ (@GRing.exp R^~ 2) _ (nbhs (x - m))). + apply: (@cvgB _ R^o) => //; exact: cvg_cst. +exact: sqr_continuous. +Qed. + +Let normal_pdf0_ub m x : normal_pdf0 m s x <= normal_peak s. +Proof. +rewrite /normal_pdf0 ler_piMr ?normal_peak_ge0//. +rewrite -[leRHS]expR0 ler_expR mulNr oppr_le0 mulr_ge0// ?sqr_ge0//. +by rewrite invr_ge0 mulrn_wge0// sqr_ge0. +Qed. + Let g' a e : R -> R := fun x => if x \in (ball a e : set R^o) then normal_peak s else normal_pdf0 e s `|x - a|. @@ -4006,7 +4029,8 @@ apply: (@continuity_under_integral _ _ _ mu _ _ _ _ (a - e) (a + e) _ _ _ g) => apply/abse_integralP => //=. by apply/measurable_EFinP; exact: measurable_normal_pdf. by rewrite integral_normal_pdf ltry. -move=> x ax. +move=> x. +rewrite !in_itv/= => /andP[aex xae]. apply/aeW => y Vy. rewrite ger0_norm; last exact: normal_pdf_ge0. rewrite normal_pdfE// /g /g'. @@ -4016,7 +4040,6 @@ rewrite /normal_pdf0 ler_pM//. rewrite ler_expR !mulNr lerN2 ler_pM //. exact: sqr_ge0. by rewrite invr_ge0 mulrn_wge0// sqr_ge0. -move: ax; rewrite in_itv/= => /andP[aex xae]. move: aey; move/negP/nandP; rewrite -!leNgt => -[yae|aey]. rewrite -normrN opprB ger0_norm; last first. by rewrite subr_ge0 (le_trans yae)// gerBl. From ed9456694ada2ad36aec7bfee42d085ed2cdbf8e Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Tue, 25 Mar 2025 19:19:14 +0900 Subject: [PATCH 42/48] minor cleaning --- theories/lang_syntax.v | 19 ++-- theories/lang_syntax_noisy.v | 27 +++-- theories/lang_syntax_table_game.v | 3 +- theories/prob_lang_wip.v | 69 +------------ theories/probability.v | 163 ++---------------------------- 5 files changed, 39 insertions(+), 242 deletions(-) diff --git a/theories/lang_syntax.v b/theories/lang_syntax.v index 933b5f89ca..7e2a0045b9 100644 --- a/theories/lang_syntax.v +++ b/theories/lang_syntax.v @@ -71,15 +71,13 @@ Reserved Notation "e -P> k" (at level 40). Local Open Scope classical_set_scope. Local Open Scope ring_scope. -(* In this module, we use our lemma continuous_FTC2 to compute the value of - * integration of the indicator function over the interval [0, 1]. - * we can use our lemma continuous_FTC2 because it requires continuous - * within [0, 1], which the indicator function satisfies. - * we also shows that the indicator function is not continuous in [0, 1], - * required by previous version of lemma continuous_FTC2. This shows that - * our lemma continuous_FTC2 is - * enough weak to be usable in practice. - *) +(* In this module, we use the lemma continuous_FTC2 to compute the value of + integration of the indicator function over the interval [0, 1]. + We can use the lemma continuous_FTC2 because it requires continuity within + [0, 1], which the indicator function satisfies. + We also show that the indicator function is not continuous in [0, 1]. + This shows that the lemma continuous_FTC2 is * enough weak to be usable + in practice. *) Module integral_indicator_function. Section integral_indicator_function. @@ -450,8 +448,7 @@ Defined. End binop. -(* TODO: rename *) -(* TODO: generalize? *) +(* TODO: rename, generalize? *) Section relop. Inductive relop := | relop_le | relop_lt | relop_eq . diff --git a/theories/lang_syntax_noisy.v b/theories/lang_syntax_noisy.v index 1deeea96b3..735eb40497 100644 --- a/theories/lang_syntax_noisy.v +++ b/theories/lang_syntax_noisy.v @@ -42,11 +42,11 @@ Local Open Scope ereal_scope. Local Open Scope string_scope. (* TODO: PR *) -Section ge0_fin_measurable_probability_integrable. +Section ge0_bounded_measurable_probability_integrable. Context d {T : measurableType d} {R : realType} {p : probability T R} {f : T -> \bar R}. -Lemma ge0_fin_measurable_probability_integrable : +Lemma ge0_bounded_measurable_probability_integrable : (forall x, 0 <= f x) -> (exists M : R, forall x, f x <= M%:E) -> measurable_fun setT f -> @@ -64,7 +64,7 @@ apply (@le_lt_trans _ _ (\int[p]_x M%:E)). by rewrite integral_cst// probability_setT mule1 ltry. Qed. -End ge0_fin_measurable_probability_integrable. +End ge0_bounded_measurable_probability_integrable. (* TODO: PR *) Section pkernel_probability_integrable. @@ -75,9 +75,8 @@ Lemma pkernel_probability_integrable V : measurable V -> p.-integrable setT (fun x => f x V). Proof. move=> mV. -apply: ge0_fin_measurable_probability_integrable => //. - exists 1%R. - move=> x. +apply: ge0_bounded_measurable_probability_integrable => //. + exists 1%R => x. apply: (@le_trans _ _ (f x setT)). by apply: le_measure; rewrite ?inE. by rewrite prob_kernel. @@ -203,7 +202,8 @@ Qed. End normal_prob_lemmas. -Section conjugate_normal_property. +(* TODO: move to probability.v *) +Section normal_probD. Context {R : realType}. Local Notation mu := lebesgue_measure. @@ -212,13 +212,13 @@ Let normal_pdf0 m s x : R := normal_peak s * normal_fun m s x. Let measurable_normal_pdf0 m s : measurable_fun setT (normal_pdf0 m s). Proof. by apply: measurable_funM => //=; exact: measurable_normal_fun. Qed. -Lemma conjugate_normal1 (m1 m2 s1 s2 : R) V : measurable V -> +Lemma normal_probD1 (m1 m2 s1 s2 : R) V : measurable V -> s1 != 0%R -> s2 != 0%R -> \int[normal_prob m1 s1]_x normal_prob (m2 + x) s2 V = \int[mu]_(y in V) \int[mu]_x (normal_pdf (m2 + x) s2 y * normal_pdf m1 s1 x)%:E. Proof. move=> mV s10 s20; rewrite integral_normal_prob//; last first. - apply: ge0_fin_measurable_probability_integrable => //=. + apply: ge0_bounded_measurable_probability_integrable => //=. by exists 1%R => ?; exact: probability_le1. apply: (@measurableT_comp _ _ _ _ _ _ (fun x => normal_prob x s2 V) _ (fun x => m2 + x)). @@ -284,7 +284,7 @@ move/negP in xV. by move/mem_set. Qed. -Lemma conjugate_normal2 (y m1 m2 s1 s2 : R) : s1 != 0%R -> s2 != 0%R -> +Lemma normal_probD2 (y m1 m2 s1 s2 : R) : s1 != 0%R -> s2 != 0%R -> \int[mu]_x (normal_pdf (m1 + x)%E s1 y * normal_pdf m2 s2 x)%:E = (normal_peak s1 * normal_peak s2)%:E * \int[mu]_z (normal_fun (m1 + z) s1 y * normal_fun m2 s2 z)%:E. @@ -315,9 +315,9 @@ Lemma normal_probD (m1 s1 m2 s2 : R) V : s1 != 0%R -> s2 != 0%R -> normal_prob (m1 + m2) (Num.sqrt (s1 ^+ 2 + s2 ^+ 2)) V. Proof. move=> s10 s20 mV. -rewrite conjugate_normal1//; apply: eq_integral => y _. +rewrite normal_probD1//; apply: eq_integral => y _. clear V mV. -rewrite conjugate_normal2//. +rewrite normal_probD2//. have s1s20 : (s1 ^+ 2 + s2 ^+ 2 != 0)%R. by rewrite lt0r_neq0// addr_gt0// exprn_even_gt0. have sqs1s20 : Num.sqrt (s1 ^+ 2 + s2 ^+ 2) != 0%R. @@ -418,7 +418,7 @@ rewrite sqrtrM; last by rewrite addr_ge0 ?sqr_ge0. by rewrite mulrC. Qed. -End conjugate_normal_property. +End normal_probD. Section noisy_programs. Local Open Scope lang_scope. @@ -586,7 +586,6 @@ rewrite (_ : ((Num.sqrt 2)^-1 ^+ 2 + 1 ^+ 2 = 3 / 2)%R)//; last first. exact. Qed. - End noisy_subproofs. (* this section contains the verification of the noisy program per se *) diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v index 0ae45e9d59..dd8064644e 100644 --- a/theories/lang_syntax_table_game.v +++ b/theories/lang_syntax_table_game.v @@ -67,6 +67,7 @@ by rewrite {}/A {}/B !setTI ysectionE. Qed. Local Open Scope lang_scope. +(* NB: consider moving to lang_syntax.v *) Lemma execP_letin_uniform {R : realType} g t str (s0 s1 : exp P ((str, Real) :: g) t) : (forall (p : R) x U, (0 <= p <= 1)%R -> @@ -86,7 +87,7 @@ Qed. End execP_letin_uniform. -(* NB: generic lemmas about bouned, to be moved *) +(* NB: generic lemmas about bounded, to be moved *) Section bounded. Local Open Scope ring_scope. Local Open Scope lang_scope. diff --git a/theories/prob_lang_wip.v b/theories/prob_lang_wip.v index 0bb9b0a6a0..d080a274cc 100644 --- a/theories/prob_lang_wip.v +++ b/theories/prob_lang_wip.v @@ -25,65 +25,6 @@ Local Open Scope classical_set_scope. Local Open Scope ring_scope. Local Open Scope ereal_scope. -Section poisson. -Variable R : realType. -Local Open Scope ring_scope. -Notation mu := (@lebesgue_measure R). -Hypothesis integral_poisson_density : forall k, - (\int[mu]_x (@poisson_pmf R x k)%:E = 1%E)%E. - -(* density function for poisson *) -Definition poisson1 r := @poisson_pmf R r 1%N. - -Lemma poisson1_ge0 (x : R) : 0 <= poisson1 x. -Proof. exact: poisson_pmf_ge0. Qed. - -Definition mpoisson1 (V : set R) : \bar R := - (\int[lebesgue_measure]_(x in V) (poisson1 x)%:E)%E. - -Lemma measurable_fun_poisson1 : measurable_fun setT poisson1. -Proof. exact: measurable_poisson_pmf. Qed. - -Let mpoisson10 : mpoisson1 set0 = 0%E. -Proof. by rewrite /mpoisson1 integral_set0. Qed. - -Lemma mpoisson1_ge0 A : (0 <= mpoisson1 A)%E. -Proof. -apply: integral_ge0 => x Ax. -rewrite lee_fin poisson1_ge0//. -Qed. - -Let mpoisson1_sigma_additive : semi_sigma_additive mpoisson1. -Proof. -move=> /= F mF tF mUF. -rewrite /mpoisson1/= integral_bigcup//=; last first. - apply/integrableP; split. - apply/measurable_EFinP. - exact: measurable_funS (measurable_poisson_pmf _ measurableT). - rewrite (_ : (fun x => _) = (EFin \o poisson1)); last first. - by apply/funext => x; rewrite gee0_abs// lee_fin poisson1_ge0//. - apply: le_lt_trans. - apply: (@ge0_subset_integral _ _ _ _ _ setT) => //=. - by apply/measurable_EFinP; exact: measurable_poisson_pmf. - by move=> ? _; rewrite lee_fin poisson1_ge0//. - by rewrite /= integral_poisson_density// ltry. -apply: is_cvg_ereal_nneg_natsum_cond => n _ _. -by apply: integral_ge0 => /= x ?; rewrite lee_fin poisson1_ge0. -Qed. - -HB.instance Definition _ := isMeasure.Build _ _ _ - mpoisson1 mpoisson10 mpoisson1_ge0 mpoisson1_sigma_additive. - -Let mpoisson1_setT : mpoisson1 [set: _] = 1%E. -Proof. exact: integral_poisson_density. Qed. - -HB.instance Definition _ := @Measure_isProbability.Build _ _ R - mpoisson1 mpoisson1_setT. - -Definition poisson' := [the probability _ _ of mpoisson1]. - -End poisson. - (* Staton's definition of the counting measure [equation (13), Sect. 4.2, Staton ESOP 2017] *) Section staton_counting. @@ -94,21 +35,21 @@ Import Notations. Hypothesis integral_poisson_density : forall k, (\int[mu]_x (@poisson_pmf R x k)%:E = 1%E)%E. -Let f1 x := ((poisson1 (x : R))^-1)%R. +Let f1 n := (@poisson_pmf R 1%R n)^-1%R. Let mf1 : measurable_fun setT f1. -rewrite /f1 /poisson1 /poisson_pmf. +Proof. +rewrite /f1 /poisson_pmf. apply: (measurable_comp (F := [set r : R | r != 0%R])) => //. - exact: open_measurable. - move=> /= r [t ? <-]. by case: ifPn => // t0; rewrite gt_eqF ?mulr_gt0 ?expR_gt0//= invrK ltr0n. - apply: open_continuous_measurable_fun => //. by apply/in_setP => x /= x0; exact: inv_continuous. -- exact: measurable_poisson_pmf. Qed. -Definition staton_counting : R.-sfker X ~> _ := - letin (sample_cst (@poisson' R integral_poisson_density : pprobability _ _)) +Definition staton_counting (r : R) : R.-sfker X ~> _ := + letin (sample_cst (@poisson_prob R r 1%N)) (letin (score (measurableT_comp mf1 macc1of2)) (ret macc1of3)). diff --git a/theories/probability.v b/theories/probability.v index 47b4eb3c8a..82f79d2193 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -3585,145 +3585,6 @@ under [X in measurable_fun _ X]eq_fun. apply: continuous_measurable_fun. *) -Local Open Scope ereal_scope. - -Local Close Scope ereal_scope. - -(* -Lemma continuousT_integralT (f : R -> R -> R) : -(forall l, mu.-integrable setT (fun x => (f l x)%:E)) -> -{ae mu, forall x, {for x, continuous (fun l => f l x)}} -> -(exists g : R -> R, forall l x, `|f l x| <= g x) -> -continuous (fun l => \int[mu]_x f l x). -Proof. -Abort. -*) - - -Lemma normal_prob_fin_num (m : R) U : normal_prob m s U \is a fin_num. -Proof. -Admitted. - -Lemma integrable_normal_pdf0 U z : mu.-integrable U - (fun x : R => (normal_pdf 0 s (x - z))%:E). -Proof. -Abort. - -(* -Lemma shift_continuous (x : R) : continuous (shift x). -Proof. -Admitted. - -Lemma add_continuous (x : R) : continuous (+%R x). -Proof. -Admitted. -*) - -Local Import Num.ExtraDef. - -Local Definition f (m x : R) := - (fun n => let sigma := s ^+ 2 in - ((sqrtr (sigma * pi *+ 2))^-1 * - series (exp_coeff (- ((x - m) ) ^+ 2 / (sigma *+ 2))) (2 * n)%N)%:E). - -Lemma near_f_ge0 : \forall n \near \oo, forall m x, (0 <= f m x n)%E. -Proof. -Admitted. -(* -apply: lt_lim. -- move => n0 n1 n01. - rewrite ler_pM2l; last first. - rewrite invr_gt0. - rewrite sqrtr_gt0. - rewrite pmulrn_rgt0//. - rewrite mulr_gt0//. - by rewrite exprn_even_gt0. - exact: pi_gt0. - set X := (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)). - have [N] := (exp_coeff2_near_in_increasing X). -(* -near=> n. -apply: mulr_ge0. - rewrite invr_ge0. - exact: sqrtr_ge0. -have : series (exp_coeff (- (x - m) ^+ 2 / (s ^+ 2 *+ 2))) (2 * n)%N @[n --> \oo] --> expR (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)). -*) - admit. - admit. -rewrite /expR. -Admitted. -*) - -Lemma near_nd_f : \forall N \near \oo, forall m x, - {in [set n | (N <= n)%N]&, nondecreasing (f m x)}. -Proof. -near=> N. -move=> m x n0 n1; rewrite !inE/= => Nn0 Nn1 n01. -rewrite /f lee_fin ler_pM2l; last first. - rewrite invr_gt0 sqrtr_gt0. - rewrite pmulrn_lgt0//. - rewrite mulr_gt0//. - by rewrite exprn_even_gt0. - exact: pi_gt0. -Admitted. - -Lemma f_first_continuous (n : nat) (x : R) : - continuous (fun m => fine (f m x n)). -Proof. -Admitted. - -Lemma f_second_continuous (n : nat) (x : R) : - continuous (fine \o (f x ^~ n)). -Proof. -Admitted. - -Lemma measurable_f_second (m : R) n : measurable_fun setT (f m ^~ n). -Proof. -Abort. - -Lemma cvg_f m x : (f m x n) @[n --> \oo] --> (normal_pdf m s x)%:E. -Proof. -rewrite /f/normal_pdf. -rewrite ifF; last exact/negP/negP. -apply/cvg_EFin. - exact/nearW. -apply: cvgMr. -have : (2 * n)%N @[n --> \oo] --> \oo. - under eq_cvg do rewrite mulnC. - exact: cvg_mulnr. -move/cvg_comp; apply. -exact: is_cvg_series_exp_coeff. -Qed. - - -Lemma integral_f_fin_num x Ys n : (\int[mu]_(x0 in Ys) f x x0 n)%E \is a fin_num. -Admitted. - -Lemma cvg_integral_f m Ys : measurable Ys -> - (\int[mu]_(x in Ys) f m x n)%E @[n --> \oo] --> normal_prob m s Ys. -Proof. -move=> mYs. -rewrite /normal_prob. -rewrite [X in _ --> X](_ : _ = (\int[mu]_(x in Ys) (limn (f m x)))%E); last first. - apply: eq_integral => y _. - apply/esym/cvg_lim => //. - exact: cvg_f. -apply: cvg_near_monotone_convergence => //. -move=> n. -rewrite /f. -(* ge0_emeasurable_sum *) -- admit. -- near=> n. - move=> x Ysx. - rewrite /f. -admit. -near=> n. -move=> x Ysx. -Admitted. - -Lemma integrable_f x Ys n : mu.-integrable Ys ((f x)^~ n). -Admitted. - (* outline of proof: 1. It is enough to prove that `(fun x => normal_prob x s Ys)` is continuous for all measurable set `Ys`. @@ -3787,7 +3648,7 @@ rewrite ball_itv/= in_itv/=; apply/negP/andP/not_andP; right. by apply/negP; rewrite -leNgt. Qed. -Lemma g'a0 (a : R) : g' a 0 = normal_pdf0 a s. +Let g'a0 (a : R) : g' a 0 = normal_pdf0 a s. Proof. apply/funext => x; rewrite /g'. have /orP [x0|x0] := le_total x a. @@ -3797,7 +3658,7 @@ rewrite ballFE_ge; last by rewrite addr0. by rewrite /normal_pdf0 /normal_fun subr0 real_normK// num_real. Qed. -Lemma mg' a e : measurable_fun setT (g' a e). +Let mg' a e : measurable_fun setT (g' a e). Proof. apply: measurable_fun_if => //. apply: (measurable_fun_bool true) => /=. @@ -3808,13 +3669,13 @@ apply: measurableT_comp => //; first exact: measurable_normal_fun. by apply: measurableT_comp => //; exact: measurable_funD. Qed. -Lemma g'_ge0 a e x : 0 <= g' a e x. +Let g'_ge0 a e x : 0 <= g' a e x. Proof. rewrite /g'; case: ifP => _; first by rewrite normal_peak_ge0. exact: normal_pdf0_ge0. Qed. -Lemma continuous_g' (a e : R) : 0 <= e -> continuous (g' a e). +Let continuous_g' (a e : R) : 0 <= e -> continuous (g' a e). Proof. move=> e0. have aNe k : k < a - e -> (`|k - a| - e) ^+ 2 = (k - (a - e)) ^+ 2. @@ -3926,7 +3787,7 @@ apply: withinU_continuous. exact: continuous_normal_pdf0. Unshelve. all: end_near. Qed. -Lemma gE_Ny a e : 0 <= e -> +Let gE_Ny a e : 0 <= e -> (\int[mu]_(x in `]-oo, (a - e)%R]) `|g' a e x|%:E = \int[mu]_(x in `]-oo, a]) `|normal_pdf a s x|%:E)%E. Proof. @@ -3949,7 +3810,7 @@ under eq_integral. by apply: eq_integral => /= x xay; rewrite /normal_pdf (negbTE s0). Qed. -Lemma gE_y a e : 0 <= e -> +Let gE_y a e : 0 <= e -> (\int[mu]_(x in `[a + e, +oo[) `|g' a e x|%:E = \int[mu]_(x in `[a, +oo[) `|normal_pdf a s x|%:E)%E. Proof. @@ -3981,10 +3842,8 @@ apply: (@continuity_under_integral _ _ _ mu _ _ _ _ (a - e) (a + e) _ _ _ g) => - rewrite in_itv/=. by rewrite ltrDl gtrBl andbb. - move=> x _. - apply: (integrableS measurableT) => //. - exact: integrable_normal_pdf. -- apply/aeW => y _. - move=> x. + by apply: (integrableS measurableT) => //=; exact: integrable_normal_pdf. +- apply/aeW => y _ x. under [X in _ _ X]eq_fun. move=> x0. rewrite normal_pdfE// /normal_pdf0 /normal_fun -(sqrrN (y - _)) opprB. @@ -4029,9 +3888,8 @@ apply: (@continuity_under_integral _ _ _ mu _ _ _ _ (a - e) (a + e) _ _ _ g) => apply/abse_integralP => //=. by apply/measurable_EFinP; exact: measurable_normal_pdf. by rewrite integral_normal_pdf ltry. -move=> x. -rewrite !in_itv/= => /andP[aex xae]. -apply/aeW => y Vy. +move=> x; rewrite in_itv/= => /andP[aex xae]. +apply: aeW => /= y Vy. rewrite ger0_norm; last exact: normal_pdf_ge0. rewrite normal_pdfE// /g /g'. case: ifPn => [_|]; first exact: normal_pdf0_ub. @@ -4076,6 +3934,7 @@ Qed. End normal_kernel. +(* TODO: move to derive.v, etc. *) Section move. Lemma cvg_comp_filter {R : realType} (f g : R -> R) (r l : R) : From a57636fcd10a869d6895b650b8995e7384d486eb Mon Sep 17 00:00:00 2001 From: IshYosh <103252572+IshiguroYoshihiro@users.noreply.github.com> Date: Fri, 9 May 2025 14:10:00 +0900 Subject: [PATCH 43/48] to avoid overflow with small RAM (#39) --- theories/lang_syntax_examples.v | 19 ++++++++++++++++--- theories/lang_syntax_noisy.v | 2 +- theories/lang_syntax_table_game.v | 8 +++++++- theories/lang_syntax_toy.v | 3 +-- 4 files changed, 25 insertions(+), 7 deletions(-) diff --git a/theories/lang_syntax_examples.v b/theories/lang_syntax_examples.v index c13f8ca378..4434c118c1 100644 --- a/theories/lang_syntax_examples.v +++ b/theories/lang_syntax_examples.v @@ -1,5 +1,5 @@ (* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *) -From Coq Require Import String. +Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. From mathcomp Require Import interval_inference. @@ -65,6 +65,8 @@ Local Open Scope ring_scope. Local Open Scope ereal_scope. Local Open Scope string_scope. +Local Open Scope string_scope. + (* simple tests to check bidirectional hints *) Module bidi_tests. Section bidi_tests. @@ -696,9 +698,20 @@ transitivity (beta_prob_bernoulli_prob 6 4 1 0 U : \bar R). rewrite beta_prob_bernoulli_probE// !bernoulli_probE//=; last 2 first. lra. by rewrite div_beta_fun_ge0 div_beta_fun_le1. -by congr (_ * _ + _ * _)%:E; - rewrite /div_beta_fun/= /onem !beta_funE/=; repeat rewrite !factE/=; field. +(*by congr (_ * _ + _ * _)%:E; + rewrite /div_beta_fun/= /onem !beta_funE/=; repeat rewrite !factE/=; field.*) + (* temporary measure to avoid stack overflow *) +suff : div_beta_fun 6 4 1 0 = 3 / 5 :> R by move->. +rewrite /div_beta_fun/= /onem !beta_funE. +rewrite addn0 invfM mulrCA invrK. +rewrite addn1 8!addnS 2!addn0. +by rewrite (factS 9) !factS fact0; field. +Qed. +(* +congr (_ * _ + _ * _)%:E. +rewrite !factE/= !factE; field. Qed. +*) End beta_bernoulli_bernoulli. diff --git a/theories/lang_syntax_noisy.v b/theories/lang_syntax_noisy.v index 735eb40497..d4cc276c9d 100644 --- a/theories/lang_syntax_noisy.v +++ b/theories/lang_syntax_noisy.v @@ -1,4 +1,4 @@ -From Coq Require Import String. +Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg ssrnum ssrint interval. From mathcomp.classical Require Import mathcomp_extra boolp. diff --git a/theories/lang_syntax_table_game.v b/theories/lang_syntax_table_game.v index dd8064644e..95d7d36490 100644 --- a/theories/lang_syntax_table_game.v +++ b/theories/lang_syntax_table_game.v @@ -564,7 +564,13 @@ under eq_integral. rewrite patchE x0 XMonemX0n. over. rewrite /= => ->; congr bernoulli_prob. -by rewrite /div_beta_fun addn0 !beta_funE/=; repeat rewrite !factE/=; field. +(*by rewrite /div_beta_fun addn0 !beta_funE/=; repeat rewrite !factE/=; field.*) +rewrite /div_beta_fun addn0 !beta_funE/=. +(* temporary measure to avoid stack overflow *) +rewrite mulrAC -mulrA mulrAC 2!invfM 3!mulrA mulfV ?gt_eqF// 2!div1r. +rewrite !addnS !addn0. +rewrite (factS 11) (factS 10) (factS 9). +by rewrite !factE; field. Qed. Lemma dirac_bool {R : realType} (U : set bool) : diff --git a/theories/lang_syntax_toy.v b/theories/lang_syntax_toy.v index 7a4ca222f0..dd30e9e854 100644 --- a/theories/lang_syntax_toy.v +++ b/theories/lang_syntax_toy.v @@ -1,5 +1,5 @@ (* mathcomp analysis (c) 2025 Inria and AIST. License: CeCILL-C. *) -From Coq Require Import String Classical. +Require Import String. From HB Require Import structures. From mathcomp Require Import all_ssreflect ssralg interval_inference. From mathcomp Require Import mathcomp_extra boolp. @@ -32,7 +32,6 @@ Import numFieldTopology.Exports. Local Open Scope classical_set_scope. Local Open Scope ring_scope. -Local Open Scope ereal_scope. Local Open Scope string_scope. Section type. From c55de4d107eeeacd1de885e8abbd5471c8db56c4 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 5 Jun 2025 19:05:02 +0900 Subject: [PATCH 44/48] cleaning - adjust XMoneMX01 - restrict opam --- coq-mathcomp-analysis.opam | 2 +- coq-mathcomp-classical.opam | 4 +- theories/lang_syntax_noisy.v | 226 ++++++++++++++--------------------- theories/probability.v | 39 ++---- 4 files changed, 105 insertions(+), 166 deletions(-) diff --git a/coq-mathcomp-analysis.opam b/coq-mathcomp-analysis.opam index f074956f7a..0d54ca9fd3 100644 --- a/coq-mathcomp-analysis.opam +++ b/coq-mathcomp-analysis.opam @@ -19,7 +19,7 @@ depends: [ "coq-mathcomp-solvable" "coq-mathcomp-field" "coq-mathcomp-bigenough" { (>= "1.0.0") } - "coq-mathcomp-algebra-tactics" { (>= "1.2.3") } + "coq-mathcomp-algebra-tactics" { (>= "1.2.4") } ] tags: [ diff --git a/coq-mathcomp-classical.opam b/coq-mathcomp-classical.opam index e32f78601a..5cc88890ee 100644 --- a/coq-mathcomp-classical.opam +++ b/coq-mathcomp-classical.opam @@ -20,8 +20,8 @@ depends: [ "coq-mathcomp-ssreflect" { (>= "2.4.0" & < "2.6~") | (= "dev") } "coq-mathcomp-fingroup" "coq-mathcomp-algebra" - "coq-mathcomp-finmap" { (>= "2.1.0") } - "coq-hierarchy-builder" { (>= "1.8.0") } + "coq-mathcomp-finmap" { (>= "2.2.0") } + "coq-hierarchy-builder" { (>= "1.8.1") } ] tags: [ diff --git a/theories/lang_syntax_noisy.v b/theories/lang_syntax_noisy.v index d4cc276c9d..d2a85d381c 100644 --- a/theories/lang_syntax_noisy.v +++ b/theories/lang_syntax_noisy.v @@ -13,7 +13,7 @@ From mathcomp Require Import lang_syntax_util lang_syntax lang_syntax_examples. (**md**************************************************************************) (* # Observing a noisy draw from a normal distribution *) (* *) -(* Formalization of the HelloRight example (Sect. 2.3 of [Shan, POPL 2018]). *) +(* Formalization of Shan's HelloRight example (Sec. 2.3 of [Shan, POPL 2018]).*) (* *) (* ref: *) (* - Chung-chieh Shan, Equational reasoning for probabilistic programming, *) @@ -41,26 +41,22 @@ Local Open Scope ring_scope. Local Open Scope ereal_scope. Local Open Scope string_scope. -(* TODO: PR *) +(* TODO: PR? *) Section ge0_bounded_measurable_probability_integrable. Context d {T : measurableType d} {R : realType} {p : probability T R} {f : T -> \bar R}. Lemma ge0_bounded_measurable_probability_integrable : - (forall x, 0 <= f x) -> - (exists M : R, forall x, f x <= M%:E) -> - measurable_fun setT f -> - p.-integrable setT f. + (forall x, 0 <= f x) -> (exists M : R, forall x, f x <= M%:E) -> + measurable_fun setT f -> p.-integrable setT f. Proof. move=> f0 [M fleM] mf. apply/integrableP; split => //. -apply (@le_lt_trans _ _ (\int[p]_x M%:E)). +rewrite (@le_lt_trans _ _ (\int[p]_x M%:E))//. apply: ge0_le_integral => //=. - exact: measurableT_comp. - move=> t _. - exact: (@le_trans _ _ (f t)). - move=> t _. - by rewrite gee0_abs. + - exact: measurableT_comp. + - by move=> t _; exact: (@le_trans _ _ (f t)). + - by move=> t _; rewrite gee0_abs. by rewrite integral_cst// probability_setT mule1 ltry. Qed. @@ -77,17 +73,17 @@ Proof. move=> mV. apply: ge0_bounded_measurable_probability_integrable => //. exists 1%R => x. - apply: (@le_trans _ _ (f x setT)). - by apply: le_measure; rewrite ?inE. + rewrite (@le_trans _ _ (f x setT))//. + by rewrite le_measure ?inE. by rewrite prob_kernel. exact: measurable_kernel. Qed. End pkernel_probability_integrable. -(* TODO: move to probability.v *) -Section normal_prob_properties. -Context {R : realType}. +(* TODO: move to probability.v? *) +Section normal_prob_lemmas. +Context {R: realType}. Local Notation mu := lebesgue_measure. Local Open Scope charge_scope. @@ -98,31 +94,24 @@ Lemma normal_pdf_uniq_ae (m s : R) (s0 : (s != 0)%R) : (EFin \o (@normal_pdf R m s)). Proof. apply: integral_ae_eq => //. - apply: Radon_Nikodym_integrable => /=. - exact: normal_prob_dominates. - apply/measurable_EFinP. - exact: measurable_normal_pdf. -move=> /= E _ mE. -by rewrite -Radon_Nikodym_integral//=; apply: normal_prob_dominates. +- by apply: Radon_Nikodym_integrable => /=; exact: normal_prob_dominates. +- by apply/measurable_EFinP; exact: measurable_normal_pdf. +- move=> /= E _ mE. + by rewrite -Radon_Nikodym_integral//=; exact: normal_prob_dominates. Qed. -End normal_prob_properties. - -(* TODO: move *) -Section normal_prob_lemmas. -Context {R: realType}. -Local Notation mu := lebesgue_measure. +Local Close Scope charge_scope. -Lemma emeasurable_normal_prob (s : R) U : s != 0%R -> measurable U -> - measurable_fun setT (fun x => normal_prob x s U). +Lemma measurable_normal_prob (s : R) U : s != 0%R -> measurable U -> + measurable_fun setT (fun x => normal_prob x s U). Proof. move=> s0 mU. under [X in _ _ X]eq_fun. - move=> x. + move=> /= x. rewrite -(@fineK _ (_ x _ _)); last first. rewrite ge0_fin_numE//. - apply: (@le_lt_trans _ _ (normal_prob x s setT)). - by apply: le_measure; rewrite ?inE. + rewrite (@le_lt_trans _ _ (normal_prob x s setT))//. + by rewrite le_measure ?inE. by rewrite probability_setT ltry. over. apply/measurable_EFinP. @@ -140,19 +129,18 @@ move=> mU intf. move/integrableP : (intf) => [mf intf_lty]. rewrite -(Radon_Nikodym_change_of_variables (normal_prob_dominates m s))//=. apply: ae_eq_integral => //=. - apply: emeasurable_funM => //. - apply: measurable_funTS. - have : charge_of_finite_measure (normal_prob m s) `<< mu. - exact: normal_prob_dominates m s. - move/Radon_Nikodym_integrable. - by move/integrableP => []. - apply: emeasurable_funM => //. +- apply: emeasurable_funM => //. + apply: measurable_funTS. + have : charge_of_finite_measure (normal_prob m s) `<< mu. + exact: normal_prob_dominates m s. + by move=> /Radon_Nikodym_integrable /integrableP[]. +- apply: emeasurable_funM => //. apply/measurable_EFinP. apply: measurable_funTS. exact: measurable_normal_pdf. -apply: ae_eqe_mul2l. -apply: (ae_eq_subset (@subsetT _ U)). -exact: (normal_pdf_uniq_ae m s0). +- apply: ae_eqe_mul2l. + apply: (ae_eq_subset (@subsetT _ U)). + exact: (normal_pdf_uniq_ae m s0). Qed. Lemma normal_prob_integrable_dirac (m s : R) (V : set R): measurable V -> @@ -160,21 +148,20 @@ Lemma normal_prob_integrable_dirac (m s : R) (V : set R): measurable V -> Proof. move=> mV. apply/integrableP; split; first exact: measurable_fun_dirac. -rewrite -(setUv V). -rewrite ge0_integral_setU => //; first last. - exact/disj_setPCl. - rewrite setUv. - apply: measurableT_comp => //. - exact: measurable_fun_dirac. +rewrite -(setUv V) ge0_integral_setU//; last 3 first. exact: measurableC. + rewrite setUv. + apply: measurableT_comp => //. + exact: measurable_fun_dirac. + exact/disj_setPCl. under eq_integral. move=> x Vx. rewrite diracE Vx/= normr1. over. under [X in _ + X < _]eq_integral. - move=> x. + move=> /= x. rewrite inE/= => nVx. - have {}nVx := (memNset nVx). + have {}nVx := memNset nVx. rewrite indicE nVx/= normr0. over. rewrite !integral_cst//=; last exact: measurableC. @@ -184,16 +171,13 @@ exact: ltey. Qed. Lemma integral_normal_prob_dirac (s : R) (m : R) V : - (s != 0)%R -> - measurable V -> - \int[normal_prob m s]_x0 (\d_x0 V) = normal_prob m s V. + (s != 0)%R -> measurable V -> + \int[normal_prob m s]_x (\d_x V) = normal_prob m s V. Proof. move=> s0 mV. -rewrite integral_normal_prob => //; last first. - exact: (normal_prob_integrable_dirac m s). +rewrite integral_normal_prob//; last exact: normal_prob_integrable_dirac. under eq_integral do rewrite diracE. -rewrite /normal_prob. -rewrite [in RHS]integral_mkcond. +rewrite /= /normal_prob [in RHS]integral_mkcond. under [in RHS]eq_integral do rewrite patchE. rewrite /=. apply: eq_integral => x _. @@ -222,66 +206,44 @@ move=> mV s10 s20; rewrite integral_normal_prob//; last first. by exists 1%R => ?; exact: probability_le1. apply: (@measurableT_comp _ _ _ _ _ _ (fun x => normal_prob x s2 V) _ (fun x => m2 + x)). - exact: emeasurable_normal_prob. + exact: measurable_normal_prob. exact: measurable_funD. -transitivity (\int[mu]_x - \int[mu]_y ((normal_pdf (m2 + x) s2 y * - normal_pdf m1 s1 x)%:E * (\1_V y)%:E)). +transitivity (\int[mu]_x \int[mu]_y + ((normal_pdf (m2 + x) s2 y * normal_pdf m1 s1 x)%:E * (\1_V y)%:E)). apply: eq_integral => y _. - rewrite /normal_prob. - rewrite -integralZr//; last first. - apply: (integrableS measurableT) => //. - exact: integrable_normal_pdf. - transitivity (\int[mu]_(x in V) (normal_pdf (m2 + y) s2 x * - normal_pdf m1 s1 y)%:E). - apply: eq_integral => z _. - by rewrite -EFinM. - rewrite integral_mkcond. - by rewrite epatch_indic. + rewrite /normal_prob -integralZr//; last first. + by apply: (integrableS measurableT) => //; exact: integrable_normal_pdf. + transitivity (\int[mu]_(x in V) + (normal_pdf (m2 + y) s2 x * normal_pdf m1 s1 y)%:E). + by apply: eq_integral => z _; rewrite -EFinM. + by rewrite integral_mkcond epatch_indic. rewrite (@fubini_tonelli _ _ _ _ _ mu mu (EFin \o ((fun xz : R * R => (normal_pdf (m2 + xz.1) s2 xz.2 * - normal_pdf m1 s1 xz.1)%R - * \1_V xz.2)%R)))/=; last 2 first. + normal_pdf m1 s1 xz.1)%R * \1_V xz.2)%R)))/=; last 2 first. apply/measurable_EFinP; apply: measurable_funM => /=; last first. - apply: measurable_indic. - rewrite -[X in measurable X]setTX. + apply: measurable_indic; rewrite -[X in measurable X]setTX. exact: measurableX. - apply: measurable_funM. - rewrite [X in measurable_fun _ X](_ : _ = (fun x0 => - normal_pdf0 0 s2 (x0.2 - (m2 + x0.1)%E))) /=; last first. - apply/funext=> x0. - rewrite /normal_pdf0. - rewrite normal_pdfE//. - by rewrite normal_fun_center. - apply: measurableT_comp. - exact: measurable_normal_pdf0. - rewrite /=. + apply: measurable_funM => /=. + rewrite [X in measurable_fun _ X](_ : _ = (fun x => + normal_pdf0 0 s2 (x.2 - (m2 + x.1)%E)))/=; last first. + by apply/funext=> x0; rewrite /normal_pdf0 normal_pdfE// normal_fun_center. + apply: measurableT_comp => /=; first exact: measurable_normal_pdf0. under eq_fun do rewrite opprD. - apply: measurable_funD => //=. - exact: measurable_funB. + by apply: measurable_funD => //=; exact: measurable_funB. by apply: measurableT_comp => //; exact: measurable_normal_pdf. - move=> x/=. - by rewrite lee_fin !mulr_ge0 ?normal_pdf_ge0. -transitivity (\int[mu]_x - \int[mu]_x0 - ((fun x0 => - (normal_pdf (m2 + x0) s2 x * normal_pdf m1 s1 x0)%:E) \_ (fun=> V x)) x0). - apply: eq_integral => x0 _. - rewrite /=. + by move=> x/=; rewrite lee_fin !mulr_ge0 ?normal_pdf_ge0. +transitivity (\int[mu]_x \int[mu]_y + ((fun y => (normal_pdf (m2 + y) s2 x * normal_pdf m1 s1 y)%:E) \_ (fun=> V x)) y). + apply: eq_integral => x0 _ /=. under eq_integral do rewrite EFinM. by rewrite -epatch_indic. rewrite [RHS]integral_mkcond/=. apply: eq_integral => /= x0 _. -rewrite patchE. -case: ifPn => xV. - apply: eq_integral => z _/=. - rewrite patchE. - by rewrite ifT. +rewrite patchE; case: ifPn => xV. + by apply: eq_integral => z _/=; rewrite patchE ifT. apply: integral0_eq => /= z _. -rewrite patchE ifF//; apply/negbTE. -rewrite notin_setE/=. -move/negP in xV. -by move/mem_set. +rewrite patchE ifF//; apply/negbTE; rewrite notin_setE/=. +by move/negP : xV; rewrite inE. Qed. Lemma normal_probD2 (y m1 m2 s1 s2 : R) : s1 != 0%R -> s2 != 0%R -> @@ -291,16 +253,15 @@ Lemma normal_probD2 (y m1 m2 s1 s2 : R) : s1 != 0%R -> s2 != 0%R -> Proof. move=> s10 s20. rewrite -ge0_integralZl//=; last 3 first. -- apply/measurable_EFinP => //=; apply: measurable_funM => //=. - + rewrite /normal_fun. + apply/measurable_EFinP => //=; apply: measurable_funM => //=. + - rewrite /normal_fun. under eq_fun do rewrite -(sqrrN (y - _)) opprB (addrC m1) -addrA -opprB. exact: measurable_normal_fun. - + exact: measurable_normal_fun. -- by move=> z _; rewrite lee_fin mulr_ge0// expR_ge0. -- by rewrite lee_fin mulr_ge0// ?normal_peak_ge0. + - exact: measurable_normal_fun. + by move=> z _; rewrite lee_fin mulr_ge0// expR_ge0. + by rewrite lee_fin mulr_ge0// ?normal_peak_ge0. apply: eq_integral => /= z _. -rewrite 2?normal_pdfE// /normal_pdf0 mulrACA /normal_fun. -by congr EFin. +by rewrite 2?normal_pdfE// /normal_pdf0 mulrACA /normal_fun. Qed. Lemma normal_peak1 : normal_peak 1 = (Num.sqrt (pi *+ 2))^-1%R :> R. @@ -347,17 +308,14 @@ transitivity (((Num.sqrt S1 * Num.sqrt S2 * pi *+ 2)^-1)%:E * apply: eq_integral. move=> x _. rewrite -expRD. - congr EFin. - congr expR. + congr ((expR _)%:E). rewrite sqr_sqrtr; last first. rewrite mulr_ge0 ?invr_ge0// ?addr_ge0 ?(@mulr_ge0 _ (_ ^+ 2))// ?sqr_ge0//. field. by apply/and3P; split. - set DS12 := S1 + S2. set MS12 := (S1 * S2)%R. set C := ((((y * s1 ^+ 2)%R + (m1 * s2 ^+ 2)%R)%E - m2 * s1 ^+ 2) / DS12)%R. - under eq_integral do rewrite expRD EFinM. rewrite ge0_integralZr//=; last first. apply/measurable_EFinP. @@ -371,7 +329,6 @@ rewrite [in RHS]EFinM. rewrite [in RHS]sqr_sqrtr//; last first. by rewrite addr_ge0// sqr_ge0. rewrite muleA; congr *%E; last by rewrite -mulNr. - (* gauss integral *) have MS12DS12_gt0 : (0 < MS12 / DS12)%R. rewrite divr_gt0//. @@ -391,12 +348,12 @@ transitivity (((Num.sqrt S1 * Num.sqrt S2 * pi *+ 2)^-1)%:E rewrite exprn_even_gt0//=. by rewrite lt0r_neq0// sqrtr_gt0. rewrite ge0_integralZl//; last 3 first. -- apply/measurable_EFinP. + apply/measurable_EFinP. exact: measurable_normal_pdf. -- move=> x _. + move=> x _. rewrite lee_fin. exact: normal_pdf_ge0. -- rewrite lee_fin invr_ge0. + rewrite lee_fin invr_ge0. exact: normal_peak_ge0. rewrite integral_normal_pdf. rewrite mule1 -EFinM; congr EFin. @@ -498,11 +455,11 @@ apply: integrableMr => //. apply/ltW; apply: le_lt_trans x_gt. rewrite -[leRHS]mul1r ler_pM ?expR_ge0//. by rewrite -expR0 ler_expR oppr_le0 mulr_ge0// ?sqr_ge0// expR0 invr_ge0. -apply/integrableP; split; first exact: emeasurable_normal_prob. -apply/abse_integralP => //; first exact: emeasurable_normal_prob. +apply/integrableP; split; first exact: measurable_normal_prob. +apply/abse_integralP => //; first exact: measurable_normal_prob. rewrite gee0_abs//; last exact: integral_ge0. apply: (@le_lt_trans _ _ (\int[normal_prob 0 1]_x (cst 1%R x)%:E)). - apply: ge0_le_integral => //; first exact: emeasurable_normal_prob. + apply: ge0_le_integral => //; first exact: measurable_normal_prob. by apply/measurable_EFinP; exact: measurable_cst. by move=> x _; exact: probability_le1. by rewrite /= integral_cst// mul1e probability_setT ltry. @@ -524,18 +481,15 @@ Proof. move=> mV; rewrite /noisyB_semantics_normal integral_normal_prob//. by under eq_integral do rewrite -muleA. apply/integrableP; split. - apply: measurable_funeM. - exact: emeasurable_normal_prob. + by apply: measurable_funeM; exact: measurable_normal_prob. apply: (@le_lt_trans _ _ (\int[normal_prob (y.1 / 2) (Num.sqrt 2)^-1]_x ((NngNum (normr_ge0 (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))))%:num%:E * (cst 1%R x)%:E))). apply: ge0_le_integral; [by []|by []| | | |]. - apply: measurableT_comp => //. - apply: measurable_funeM. - exact: emeasurable_normal_prob. + by apply: measurable_funeM; exact: measurable_normal_prob. - by move=> x _; rewrite /= mule1. -- rewrite /= mule1. - exact/measurable_EFinP. +- by rewrite /= mule1; exact/measurable_EFinP. - move=> x _/=. rewrite abseM/= lee_pmul//. rewrite gee0_abs; last exact: measure_ge0. @@ -560,14 +514,16 @@ rewrite mulrA mulrAC -(@sqrtrV _ 2)//. rewrite /normal_peak sqr_sqrtr; last by rewrite invr_ge0. rewrite /normal_fun subr0 sqr_sqrtr; last by rewrite invr_ge0. rewrite -mulrA mulrAC mulrA. -rewrite [X in (X / _ / _)%R = _](_ : _ = expR (- x ^+ 2 / (1 ^+ 2 *+ 2) - (y.1 - x) ^+ 2%R / 2)); last first. +rewrite [X in (X / _ / _)%R = _](_ : _ = + expR (- x ^+ 2 / (1 ^+ 2 *+ 2) - (y.1 - x) ^+ 2%R / 2)); last first. by rewrite mulrC -expRD -mulrA. -rewrite [RHS]mulrAC [X in _ = (_ * X / _)%R]mulrC mulrA -mulrA -[RHS]mulrA -expRD -2!invfM. +rewrite [RHS]mulrAC [X in _ = (_ * X / _)%R]mulrC mulrA -mulrA -[RHS]mulrA. +rewrite -expRD -2!invfM. congr ((expR _) * _^-1)%R. lra. rewrite -2?sqrtrM; last 2 first. -- by rewrite -mulr_natr mulrAC mulVf// mul1r pi_ge0. -- by rewrite expr1n mul1r mulrn_wge0// pi_ge0. + by rewrite -mulr_natr mulrAC mulVf// mul1r pi_ge0. + by rewrite expr1n mul1r mulrn_wge0// pi_ge0. congr Num.sqrt. lra. Qed. @@ -701,7 +657,7 @@ transitivity (noisyB_semantics_normal y V); last first. rewrite ger0_norm; last by rewrite mulr_ge0// expR_ge0. rewrite -(@ge0_integralZl _ _ R (normal_prob _ _) _ measurableT _ _ (expR (- (y.1 ^+ 2%R / 4)) / Num.sqrt (4 * pi))%:E)//. - exact: emeasurable_normal_prob. + exact: measurable_normal_prob. rewrite noisyA_semantics_normalE//. rewrite noisyB_semantics_normalE//. apply: eq_integral => x _. diff --git a/theories/probability.v b/theories/probability.v index 82f79d2193..4c80687b24 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -3934,33 +3934,16 @@ Qed. End normal_kernel. -(* TODO: move to derive.v, etc. *) -Section move. - -Lemma cvg_comp_filter {R : realType} (f g : R -> R) (r l : R) : - continuous f -> - (f \o g) x @[x --> r] --> l -> - f x @[x --> g r] --> l. +Lemma continuous_comp_cvg {R : numFieldType} (U V : pseudoMetricNormedZmodType R) + (f : U -> V) (g : R -> U) (r : R) (l : V) : continuous f -> + (f \o g) x @[x --> r] --> l -> f x @[x --> g r] --> l. Proof. -move=> cf fgrl. -apply/(@cvgrPdist_le _ R^o) => /= e e0. +move=> cf fgl; apply/(@cvgrPdist_le _ V) => /= e e0. have e20 : 0 < e / 2 by rewrite divr_gt0. -move/(@cvgrPdist_le _ R^o) : fgrl => /(_ _ e20) fgrl. -have := cf (g r). -move=> /(@cvgrPdist_le _ R^o) => /(_ _ e20)[x x0]H. -exists (minr x (e/2)). - by rewrite lt_min x0. -move=> z. -rewrite /ball_ /= => grze. -rewrite -[X in X - _](subrK (f (g r))). -rewrite -(addrA _ _ (- f z)). -apply: (le_trans (ler_normD _ _)). -rewrite (splitr e) lerD//. - case: fgrl => d /= d0 K. - apply: K. - by rewrite /ball_/= subrr normr0. -apply: H => /=. -by rewrite (lt_le_trans grze)// ge_min lexx. -Qed. - -End move. +move/(@cvgrPdist_le _ V) : fgl => /(_ _ e20) fgl. +have /(@cvgrPdist_le _ V) /(_ _ e20) fgf := cf (g r). +rewrite !near_simpl/=; near=> t. +rewrite -(@subrK _ (f (g r)) l) -(addrA (_ + _)) (le_trans (ler_normD _ _))//. +rewrite (splitr e) lerD//; last by near: t. +by case: fgl => d /= d0; apply; rewrite /ball_/= subrr normr0. +Unshelve. all: by end_near. Qed. From 8f5d6ee0508c1805902bad068bab6af5b013be10 Mon Sep 17 00:00:00 2001 From: IshYosh <103252572+IshiguroYoshihiro@users.noreply.github.com> Date: Sat, 7 Jun 2025 21:55:01 +0900 Subject: [PATCH 45/48] poisson probability (#41) * add measurable_poisson_prob --- theories/probability.v | 3 +++ 1 file changed, 3 insertions(+) diff --git a/theories/probability.v b/theories/probability.v index 4c80687b24..4155f49c8f 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -75,6 +75,9 @@ From mathcomp Require Import charge ftc gauss_integral hoelder. (* beta_pdf == probability density function for beta *) (* beta_prob == beta probability measure *) (* div_beta_fun a b c d := beta_fun (a + c) (b + d) / beta_fun a b *) +(* poisson_pmf r == a pmf of Poisson distribution with parameter r : R *) +(* poisson_prob r == probability measure of Poisson distribution when *) +(* 0 <= r and \d_0%N otherwise *) (* ``` *) (* *) (******************************************************************************) From 51f89985da63f3b87cb1a449e8a1c45c0e062987 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Sun, 8 Jun 2025 15:56:09 +0900 Subject: [PATCH 46/48] rm dup sections normal_density and normal_probability --- theories/probability.v | 156 ----------------------------------------- 1 file changed, 156 deletions(-) diff --git a/theories/probability.v b/theories/probability.v index 4155f49c8f..c65cbb9c8c 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -75,9 +75,6 @@ From mathcomp Require Import charge ftc gauss_integral hoelder. (* beta_pdf == probability density function for beta *) (* beta_prob == beta probability measure *) (* div_beta_fun a b c d := beta_fun (a + c) (b + d) / beta_fun a b *) -(* poisson_pmf r == a pmf of Poisson distribution with parameter r : R *) -(* poisson_prob r == probability measure of Poisson distribution when *) -(* 0 <= r and \d_0%N otherwise *) (* ``` *) (* *) (******************************************************************************) @@ -3204,159 +3201,6 @@ Qed. End exp_coeff_properties. -Section normal_density. -Context {R : realType}. -Local Open Scope ring_scope. -Local Import Num.ExtraDef. -Implicit Types m s x : R. - -Definition normal_pdf0 m s x : R := normal_peak s * normal_fun m s x. - -Lemma normal_pdf0_center m s : normal_pdf0 m s = normal_pdf0 0 s \o center m. -Proof. by rewrite /normal_pdf0 normal_fun_center. Qed. - -Lemma normal_pdf0_ge0 m s x : 0 <= normal_pdf0 m s x. -Proof. by rewrite mulr_ge0 ?normal_peak_ge0 ?expR_ge0. Qed. - -Lemma normal_pdf0_gt0 m s x : s != 0 -> 0 < normal_pdf0 m s x. -Proof. by move=> s0; rewrite mulr_gt0 ?expR_gt0// normal_peak_gt0. Qed. - -Lemma measurable_normal_pdf0 m s : measurable_fun setT (normal_pdf0 m s). -Proof. by apply: measurable_funM => //=; exact: measurable_normal_fun. Qed. - -Lemma continuous_normal_pdf0 m s : continuous (normal_pdf0 m s). -Proof. -move=> x; apply: cvgM; first exact: cvg_cst. -apply: (@cvg_comp _ R^o _ _ _ _ - (nbhs (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)))); last exact: continuous_expR. -apply: cvgM; last exact: cvg_cst; apply: (@cvgN _ R^o). -apply: (@cvg_comp _ _ _ _ (@GRing.exp R^~ 2) _ (nbhs (x - m))). - apply: (@cvgB _ R^o) => //; exact: cvg_cst. -exact: sqr_continuous. -Qed. - -Lemma normal_pdf0_ub m s x : normal_pdf0 m s x <= normal_peak s. -Proof. -rewrite /normal_pdf0 ler_piMr ?normal_peak_ge0//. -rewrite -[leRHS]expR0 ler_expR mulNr oppr_le0 mulr_ge0// ?sqr_ge0//. -by rewrite invr_ge0 mulrn_wge0// sqr_ge0. -Qed. - -End normal_density. - -Section normal_probability. -Variables (R : realType) (m sigma : R). -Local Open Scope ring_scope. -Notation mu := lebesgue_measure. - -Local Notation normal_peak := (normal_peak sigma). -Local Notation normal_fun := (normal_fun m sigma). - -Let measurable_normal_fun : measurable_fun setT normal_fun. -Proof. -apply: measurableT_comp => //=; apply: measurable_funM => //=. -apply: measurableT_comp => //=; apply: measurable_funX => //=. -exact: measurable_funD. -Qed. - -Let F (x : R^o) := (x - m) / (Num.sqrt (sigma ^+ 2 *+ 2)). -Lemma F'E : F^`()%classic = cst (Num.sqrt (sigma ^+ 2 *+ 2))^-1. -Proof. -apply/funext => x; rewrite /F derive1E deriveM// deriveD// derive_cst scaler0. -by rewrite add0r derive_id derive_cst addr0 scaler1. -Qed. - -Let int_gaussFE : sigma != 0 -> - (\int[mu]_x ((((gauss_fun \o F) * - (F^`())%classic) x)%:E * (Num.sqrt (sigma ^+ 2 *+ 2))%:E))%E = - (Num.sqrt (sigma ^+ 2 * pi *+ 2))%:E. -Proof. -move=> s0. -rewrite -mulrnAr -[in RHS]mulr_natr sqrtrM ?(sqrtrM 2) ?sqr_ge0 ?pi_ge0// !EFinM. -rewrite muleCA ge0_integralZr//=; last first. - by move=> x _; rewrite lee_fin mulr_ge0//= ?gauss_fun_ge0// F'E/= invr_ge0. - rewrite F'E; apply/measurable_EFinP/measurable_funM => //. - apply/measurableT_comp => //; first exact: measurable_gauss_fun. - by apply: measurable_funM => //; exact: measurable_funD. -congr *%E. -rewrite -increasing_ge0_integration_by_substitutionT//; first last. -- by move=> x; rewrite gauss_fun_ge0. -- exact: continuous_gauss_fun. -- apply/gt0_cvgMly; last exact: cvg_addrr. - by rewrite invr_gt0// sqrtr_gt0 -mulr_natr mulr_gt0// exprn_even_gt0. -- apply/gt0_cvgMlNy; last exact: cvg_addrr_Ny. - by rewrite invr_gt0// sqrtr_gt0 -mulr_natr mulr_gt0// exprn_even_gt0. -- by rewrite F'E; exact: is_cvg_cst. -- by rewrite F'E; exact: is_cvg_cst. -- by rewrite F'E => ?; exact: cvg_cst. -- move=> x y xy; rewrite /F ltr_pM2r ?ltr_leB//. - rewrite invr_gt0 ?sqrtr_gt0 ?pmulrn_lgt0 ?exprn_even_gt0//. -apply: integralT_gauss. -by rewrite -(mulr_natr (_ ^+ 2)) sqrtrM ?sqr_ge0. -Qed. - -Let integral_normal_pdf_part : sigma != 0 -> - (\int[mu]_x (normal_fun x)%:E = - (Num.sqrt (sigma ^+ 2 * pi *+ 2))%:E)%E. -Proof. -move=> s0; rewrite -int_gaussFE//; apply: eq_integral => /= x _. -rewrite F'E !fctE/= EFinM -muleA -EFinM mulVf ?mulr1; last first. - by rewrite gt_eqF// sqrtr_gt0 pmulrn_lgt0// exprn_even_gt0. -congr (expR _)%:E; rewrite mulNr exprMn exprVn; congr (- (_ * _^-1)%R). -by rewrite sqr_sqrtr// pmulrn_lge0// sqr_ge0. -Qed. - -Let integrable_normal_pdf_part : sigma != 0 -> - mu.-integrable setT (EFin \o normal_fun). -Proof. -move=> s0; apply/integrableP; split. - by apply/measurable_EFinP; apply: measurable_normal_fun. -under eq_integral do rewrite /= ger0_norm ?expR_ge0//. -by rewrite /= integral_normal_pdf_part// ltry. -Qed. - -Local Notation normal_pdf := (normal_pdf m sigma). -Local Notation normal_prob := (normal_prob m sigma). - -Let normal0 : normal_prob set0 = 0%E. -Proof. by rewrite /normal_prob integral_set0. Qed. - -Let normal_ge0 A : (0 <= normal_prob A)%E. -Proof. -rewrite /normal_prob integral_ge0//= => x _. -by rewrite lee_fin normal_pdf_ge0 ?ltW. -Qed. - -Let normal_sigma_additive : semi_sigma_additive normal_prob. -Proof. -move=> /= A mA tA mUA. -rewrite /normal_prob/= integral_bigcup//=; last first. - apply: (integrableS _ _ (@subsetT _ _)) => //; exact: integrable_normal_pdf. -apply: is_cvg_ereal_nneg_natsum_cond => n _ _. -by apply: integral_ge0 => /= x ?; rewrite lee_fin normal_pdf_ge0 ?ltW. -Qed. - -HB.instance Definition _ := isMeasure.Build _ _ _ - normal_prob normal0 normal_ge0 normal_sigma_additive. - -Let normal_setT : normal_prob [set: _] = 1%E. -Proof. by rewrite /normal_prob integral_normal_pdf. Qed. - -HB.instance Definition _ := - @Measure_isProbability.Build _ _ R normal_prob normal_setT. - -Lemma integrable_indic_itv (a b : R) (b0 b1 : bool) : a < b -> - mu.-integrable setT (EFin \o \1_[set` (Interval (BSide b0 a) (BSide b1 b))]). -Proof. -move=> ab. -apply/integrableP; split; first by apply/measurable_EFinP/measurable_indic. -apply/abse_integralP => //; first by apply/measurable_EFinP/measurable_indic. -rewrite integral_indic// setIT/= lebesgue_measure_itv/=. -by rewrite lte_fin ab -EFinD/= ltry. -Qed. - -End normal_probability. - (* TODO: move *) Section shift_properties. Variable R : realType. From e69653cde4b069caad65a59c1babc53760a1f29c Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Thu, 18 Sep 2025 21:40:00 +0900 Subject: [PATCH 47/48] to accommodate for the integration of lra --- theories/lang_syntax_noisy.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/theories/lang_syntax_noisy.v b/theories/lang_syntax_noisy.v index d2a85d381c..cad6006590 100644 --- a/theories/lang_syntax_noisy.v +++ b/theories/lang_syntax_noisy.v @@ -311,8 +311,7 @@ transitivity (((Num.sqrt S1 * Num.sqrt S2 * pi *+ 2)^-1)%:E * congr ((expR _)%:E). rewrite sqr_sqrtr; last first. rewrite mulr_ge0 ?invr_ge0// ?addr_ge0 ?(@mulr_ge0 _ (_ ^+ 2))// ?sqr_ge0//. - field. - by apply/and3P; split. + by field; do ?[apply/and3P; split]. set DS12 := S1 + S2. set MS12 := (S1 * S2)%R. set C := ((((y * s1 ^+ 2)%R + (m1 * s2 ^+ 2)%R)%E - m2 * s1 ^+ 2) / DS12)%R. From dc8008ef808e62ee65c4fae075050cfdb0c25ad9 Mon Sep 17 00:00:00 2001 From: Reynald Affeldt Date: Wed, 29 Oct 2025 13:20:32 +0900 Subject: [PATCH 48/48] gen cont under int --- theories/lang_syntax_noisy.v | 24 +- .../lebesgue_integral_under.v | 18 +- theories/probability.v | 318 ++++++------------ 3 files changed, 128 insertions(+), 232 deletions(-) diff --git a/theories/lang_syntax_noisy.v b/theories/lang_syntax_noisy.v index cad6006590..950d5cf2b0 100644 --- a/theories/lang_syntax_noisy.v +++ b/theories/lang_syntax_noisy.v @@ -317,16 +317,15 @@ set MS12 := (S1 * S2)%R. set C := ((((y * s1 ^+ 2)%R + (m1 * s2 ^+ 2)%R)%E - m2 * s1 ^+ 2) / DS12)%R. under eq_integral do rewrite expRD EFinM. rewrite ge0_integralZr//=; last first. - apply/measurable_EFinP. - apply: measurableT_comp => //. - apply: measurable_funM => //. - apply: measurableT_comp => //. - apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)%R) => //. - exact: measurable_funD. + apply/measurable_EFinP. + apply: measurableT_comp => //. + apply: measurable_funM => //. + apply: measurableT_comp => //. + apply: (@measurableT_comp _ _ _ _ _ _ (fun t : R => t ^+ 2)%R) => //. + exact: measurable_funD. rewrite /normal_peak /normal_fun. rewrite [in RHS]EFinM. -rewrite [in RHS]sqr_sqrtr//; last first. - by rewrite addr_ge0// sqr_ge0. +rewrite [in RHS]sqr_sqrtr//; last by rewrite addr_ge0// sqr_ge0. rewrite muleA; congr *%E; last by rewrite -mulNr. (* gauss integral *) have MS12DS12_gt0 : (0 < MS12 / DS12)%R. @@ -339,14 +338,11 @@ transitivity (((Num.sqrt S1 * Num.sqrt S2 * pi *+ 2)^-1)%:E congr *%E. apply: eq_integral => x _. rewrite -EFinM; congr EFin. - rewrite normal_pdfE; last first. - apply: lt0r_neq0. - by rewrite sqrtr_gt0. + rewrite normal_pdfE; last by rewrite lt0r_neq0// sqrtr_gt0. rewrite mulrA mulVf// ?mul1r//. rewrite lt0r_neq0// invr_gt0 sqrtr_gt0 pmulrn_lgt0// mulr_gt0// ?pi_gt0//. - rewrite exprn_even_gt0//=. - by rewrite lt0r_neq0// sqrtr_gt0. -rewrite ge0_integralZl//; last 3 first. + by rewrite exprn_even_gt0//= lt0r_neq0// sqrtr_gt0. +rewrite ge0_integralZl//=; last 3 first. apply/measurable_EFinP. exact: measurable_normal_pdf. move=> x _. diff --git a/theories/lebesgue_integral_theory/lebesgue_integral_under.v b/theories/lebesgue_integral_theory/lebesgue_integral_under.v index 2ceb1e5dfd..c0f9d5f063 100644 --- a/theories/lebesgue_integral_theory/lebesgue_integral_under.v +++ b/theories/lebesgue_integral_theory/lebesgue_integral_under.v @@ -41,12 +41,11 @@ Variable f : R -> Y -> R. Variable B : set Y. Hypothesis mB : measurable B. -Variable a u v : R. +Variable u v : R. Let I : set R := `]u, v[. -Hypothesis Ia : I a. Hypothesis int_f : forall x, I x -> mu.-integrable B (EFin \o (f x)). -Hypothesis cf : {ae mu, forall y, B y -> continuous (f ^~ y)}. +Hypothesis cf : {ae mu, forall y, B y -> {in I, continuous (f ^~ y)}}. Variable g : Y -> R. @@ -56,13 +55,14 @@ Hypothesis g_ub : forall x, I x -> {ae mu, forall y, B y -> `|f x y| <= g y}. Let F x := (\int[mu]_(y in B) f x y)%R. Lemma continuity_under_integral : - continuous_at a (fun l => \int[mu]_(x in B) f l x). + {in I, continuous (fun l => \int[mu]_(x in B) f l x)}. Proof. +move=> a /set_mem Ia. have [Z [mZ Z0 /subsetCPl ncfZ]] := cf. -have BZ_cf x : x \in B `\` Z -> continuous (f ^~ x). +have BZ_cf x : x \in B `\` Z -> {in I, continuous (f ^~ x)}. by rewrite inE/= => -[Bx nZx]; exact: ncfZ. have [vu|uv] := lerP v u. - by move: Ia; rewrite /I set_itv_ge// -leNgt bnd_simp. + by move: (Ia); rewrite /I set_itv_ge// -leNgt bnd_simp. apply/cvg_nbhsP => w wa. have /near_in_itvoo[e /= e0 aeuv] : a \in `]u, v[ by rewrite inE. move/cvgrPdist_lt : (wa) => /(_ _ e0)[N _ aue]. @@ -114,7 +114,8 @@ apply: (@dominated_cvg _ _ _ mu _ _ have : x \in B `\` Z. move: BZUUx; rewrite inE/= => -[Bx nZUUx]; rewrite inE/=; split => //. by apply: contra_not nZUUx; left. - by move/(BZ_cf x)/(_ a)/cvg_nbhsP; apply; rewrite (cvg_shiftn N). + move/(BZ_cf x)/(_ a); move/mem_set : Ia => /[swap] /[apply]. + by move/cvg_nbhsP; apply; rewrite (cvg_shiftn N). - by apply: (integrableS mB) => //; exact: measurableD. - move=> n x [Bx ZUUx]; rewrite lee_fin. move/subsetCPl : (Ug_ub n); apply => //=. @@ -125,7 +126,8 @@ End continuity_under_integral. Section differentiation_under_integral. -Definition partial1of2 {R : realType} {T : Type} (f : R -> T -> R) : R -> T -> R := fun x y => (f ^~ y)^`() x. +Definition partial1of2 {R : realType} {T : Type} (f : R -> T -> R) : R -> T -> R + := fun x y => (f ^~ y)^`() x. Local Notation "'d1 f" := (partial1of2 f). diff --git a/theories/probability.v b/theories/probability.v index c65cbb9c8c..34d7428939 100644 --- a/theories/probability.v +++ b/theories/probability.v @@ -3281,7 +3281,8 @@ rewrite /normal_pdf ifF//. exact/negP/negP. Qed. -Local Definition normal_prob2 := (fun m => normal_prob m s) : _ -> pprobability _ _. +Local Definition normal_prob2 := + (fun m => normal_prob m s) : R -> pprobability (measurableTypeR R) R. Lemma bij_shift x : bijective (id \+ @cst R R x). Proof. @@ -3334,104 +3335,6 @@ Qed. From mathcomp Require Import charge. Open Scope charge_scope. -(* -Lemma radon_nikodym_crestr_fin U (mU : measurable U) -(Uoo : (@lebesgue_measure R U < +oo)%E) : - ae_eq lebesgue_measure setT ('d charge_of_finite_measure (mfrestr mU Uoo) '/d - [the sigma_finite_measure _ _ of @lebesgue_measure R]) - (EFin \o \1_U). -Proof. -apply: integral_ae_eq => //=. -- admit. -- admit. -move=> E _ mE. -rewrite -Radon_Nikodym_integral. -rewrite integral_indic/=. -by rewrite /mfrestr/mrestr setIC. -Admitted. -*) - -(* -Lemma radon_nikodym_crestr U (mU : measurable U) : - ae_eq lebesgue_measure setT ('d charge_of_finite_measure (mfrestr mU Uoo) '/d - [the sigma_finite_measure _ _ of @lebesgue_measure R]) - (EFin \o \1_U). -Proof. -*) - -(* -rewrite [RHS](_:_= ('d charge_of_finite_measure (mfrestr mU Uoo) '/d - [the sigma_finite_measure _ _ of @lebesgue_measure R]) - (EFin \o \1_U) - move=> x _. - rewrite epatch_indic. - rewrite -radon_nikodym_crestr. -rewrite [RHS]integral_mkcond. -under [RHS]eq_integral do rewrite epatch_indic. - -rewrite -integral_pushforward. -apply: eq_integral. -move=> x _. -Admitted. -*) -(*Local Definition normal_prob2 := - (fun m => normal_prob m s) : _ -> pprobability _ _. -*) -(* -Lemma normal_shift0 x : -normal_prob2 x = - @pushforward _ _ _ - (measurableTypeR R) _ (normal_prob2 0%R) (fun z => z + x) - :> (set R -> \bar R). -Proof. -apply: funext. -move=> U. -rewrite /normal_prob2/=. -rewrite /pushforward/=. -rewrite /normal_prob. -rewrite shift_preimage. -rewrite integration_by_substitution_shift/=. -apply: eq_integral. -move=> z Uz. -congr EFin. -rewrite /normal_pdf/=. -rewrite ifF; last exact/negP/negP. -rewrite ifF; last exact/negP/negP. -rewrite {2}/normal_fun. -by rewrite subr0. -Qed. -*) - -(* -Lemma measurable_normal_prob2_ocitv a b: - measurable_fun [set: R] (normal_prob2 ^~ `]a, b]%classic). -Proof. -apply: (@measurability _ _ _ _ _ _ - (@pset _ _ _ : set (set (pprobability _ R)))) => //. -move=> _ -[_ [r r01] [Ys mYs <-]] <-; apply: emeasurable_fun_infty_o => //=. - -rewrite /normal_prob2/=. -rewrite /normal_prob. - -under [X in measurable_fun _ X]eq_fun. - move=> x. - rewrite (_: normal_kernel _ _ = (fine (normal_kernel x `]a, b]%classic))%:E); last first. - rewrite fineK//. - rewrite ge0_fin_numE//. - apply: (@le_lt_trans _ _ 1%E); last exact: ltey. - exact: probability_le1. - rewrite normal_shift0/=. - over. -apply: measurableT_comp; last by []. -apply: measurableT_comp; first exact: EFin_measurable. -rewrite /=. -under [X in measurable_fun _ X]eq_fun. - move=> x. - rewrite /normal_prob. -(pushforward_shift_itv (normal_kernel 0) a b x). -apply: continuous_measurable_fun. -*) - (* outline of proof: 1. It is enough to prove that `(fun x => normal_prob x s Ys)` is continuous for all measurable set `Ys`. @@ -3451,58 +3354,48 @@ apply: continuous_measurable_fun. added by the integral of `normal_pdf a s x` on ]a - e, a + e[ *) -Let normal_pdf0 m s x : R := normal_peak s * normal_fun m s x. +Let normal_peak_fun m sigma x : R := + normal_peak sigma * normal_fun m sigma x. -Let normal_pdf0_ge0 m x : 0 <= normal_pdf0 m s x. +Let normal_peak_fun_ge0 m x : 0 <= normal_peak_fun m s x. Proof. by rewrite mulr_ge0 ?normal_peak_ge0 ?expR_ge0. Qed. -Let continuous_normal_pdf0 m : continuous (normal_pdf0 m s). +Let continuous_normal_peak_fun m : continuous (normal_peak_fun m s). Proof. move=> x; apply: cvgM; first exact: cvg_cst. apply: (@cvg_comp _ R^o _ _ _ _ (nbhs (- (x - m) ^+ 2 / (s ^+ 2 *+ 2)))); last exact: continuous_expR. apply: cvgM; last exact: cvg_cst; apply: (@cvgN _ R^o). apply: (@cvg_comp _ _ _ _ (@GRing.exp R^~ 2) _ (nbhs (x - m))). - apply: (@cvgB _ R^o) => //; exact: cvg_cst. + by apply: (@cvgB _ R^o) => //; exact: cvg_cst. exact: sqr_continuous. Qed. -Let normal_pdf0_ub m x : normal_pdf0 m s x <= normal_peak s. +Let normal_peak_fun_ub m x : normal_peak_fun m s x <= normal_peak s. Proof. -rewrite /normal_pdf0 ler_piMr ?normal_peak_ge0//. +rewrite /normal_peak_fun ler_piMr ?normal_peak_ge0//. rewrite -[leRHS]expR0 ler_expR mulNr oppr_le0 mulr_ge0// ?sqr_ge0//. by rewrite invr_ge0 mulrn_wge0// sqr_ge0. Qed. -Let g' a e : R -> R := fun x => if x \in (ball a e : set R^o) then - normal_peak s else normal_pdf0 e s `|x - a|. - -Let ballFE_le (a e x : R) : x <= (a - e)%R -> - (x \in (ball a e : set R^o)) = false. +Let normal_peak_funxx e : normal_peak_fun e s e = normal_peak s. Proof. -move=> xae. -apply: memNset. -rewrite ball_itv/= in_itv/=; apply/negP/andP/not_andP; left. -by apply/negP; rewrite -leNgt. +rewrite /normal_peak_fun /normal_fun. +by rewrite subrr expr0n/= oppr0 mul0r expR0 mulr1. Qed. -Let ballFE_ge (a e x : R) : a + e <= x -> - (x \in (ball a e : set R^o)) = false. +Let g' a e : R -> R := fun x => if x \in (ball a e : set R^o) then + normal_peak s else normal_peak_fun e s `|x - a|. + +Let normal_fun0 x a : normal_fun 0 s `|x - a| = normal_fun a s x. Proof. -move=> xae. -apply: memNset. -rewrite ball_itv/= in_itv/=; apply/negP/andP/not_andP; right. -by apply/negP; rewrite -leNgt. +by rewrite /normal_peak_fun /normal_fun subr0 real_normK// num_real. Qed. -Let g'a0 (a : R) : g' a 0 = normal_pdf0 a s. +Let g'a0 (a : R) : g' a 0 = normal_peak_fun a s. Proof. apply/funext => x; rewrite /g'. -have /orP [x0|x0] := le_total x a. - rewrite ballFE_le; last by rewrite subr0. - by rewrite /normal_pdf0 /normal_fun subr0 real_normK// num_real. -rewrite ballFE_ge; last by rewrite addr0. -by rewrite /normal_pdf0 /normal_fun subr0 real_normK// num_real. +by rewrite /normal_peak_fun normal_fun0 memNset// le0_ball0. Qed. Let mg' a e : measurable_fun setT (g' a e). @@ -3519,7 +3412,19 @@ Qed. Let g'_ge0 a e x : 0 <= g' a e x. Proof. rewrite /g'; case: ifP => _; first by rewrite normal_peak_ge0. -exact: normal_pdf0_ge0. +exact: normal_peak_fun_ge0. +Qed. + +Let ballFE_le (a e x : R) : x <= (a - e)%R -> + x \notin (ball a e : set R^o). +Proof. +by move=> xae; rewrite memNset// ball_itv/= in_itv/= !ltNge xae/=. +Qed. + +Let ballFE_ge (a e x : R) : a + e <= x -> + x \notin (ball a e : set R^o). +Proof. +by move=> xae; rewrite memNset// ball_itv/= in_itv/= !ltNge xae andbF. Qed. Let continuous_g' (a e : R) : 0 <= e -> continuous (g' a e). @@ -3553,85 +3458,77 @@ apply: withinU_continuous. near=> t. have tae : t < a - e by near: t; exact: lt_nbhsl. rewrite /g'. - rewrite !ballFE_le ?(@ltW _ _ _ (a - e))//. - rewrite /normal_pdf0 /normal_fun !aNe//. - rewrite -!/(normal_fun _ _ _) -!/(normal_pdf0 _ _ _). + rewrite !(negbTE (ballFE_le (ltW _)))//. + rewrite /normal_peak_fun /normal_fun !aNe//. + rewrite -!/(normal_fun _ _ _) -!/(normal_peak_fun _ _ _). move=> {tae}; near: t. move: eps eps0. - apply/(@cvgrPdist_le R^o). - exact: continuous_normal_pdf0. + apply/cvgrPdist_le. + exact: continuous_normal_peak_fun. * apply/(@cvgrPdist_lt _ R^o) => eps eps0. near=> t. - rewrite /g' !ballFE_le//. + rewrite /g'. + rewrite !(negPf (ballFE_le _))//. rewrite -addrAC subrr sub0r normrN (ger0_norm e0)//. - rewrite /normal_pdf0 /normal_fun subrr -(subrr (a - e)) aNe//. + rewrite /normal_peak_fun /normal_fun subrr -(subrr (a - e)) aNe//. + rewrite -!/(normal_peak_fun _ _ _). near: t; move: eps eps0. - apply/(@cvgrPdist_lt _ R^o). - apply: cvg_at_left_filter. - exact: continuous_normal_pdf0. + apply/cvgrPdist_lt/cvg_at_left_filter. + exact: continuous_normal_peak_fun. move: e0; rewrite le_eqVlt => /predU1P[<-|e0]. rewrite g'a0. apply: continuous_subspaceT. - exact: continuous_normal_pdf0. + exact: continuous_normal_peak_fun. apply/continuous_within_itvP; first by rewrite -(opprK e) ler_ltB// opprK gtrN. split. + move=> x xae. - rewrite /continuous_at. - rewrite /g' ifT; last by rewrite ball_itv inE/=. - apply/(@cvgrPdist_le _ R^o) => eps eps0. + rewrite /continuous_at /g' ifT; last by rewrite ball_itv inE. + apply/cvgrPdist_le => eps eps0. near=> t. rewrite ifT; first by rewrite subrr normr0 ltW. rewrite ball_itv inE/= in_itv/=; apply/andP; split. - near: t. - apply: lt_nbhsr. + - near: t; apply: lt_nbhsr. + by move: xae; rewrite in_itv/= => /andP[]. + - near: t; apply: lt_nbhsl. by move: xae; rewrite in_itv/= => /andP[]. - near: t. - apply: lt_nbhsl. - by move: xae; rewrite in_itv/= => /andP[]. + apply/(@cvgrPdist_le _ R^o) => eps eps0. near=> t. - rewrite /g' ballFE_le// ifT; last first. + rewrite /g' (negPf (ballFE_le _))// ifT; last first. rewrite ball_itv inE/= in_itv/=; apply/andP => []; split => //. - near: t. - apply: nbhs_right_lt. - by rewrite -(opprK e) ler_ltB// opprK gtrN. - rewrite /normal_pdf0. - rewrite addrAC subrr sub0r normrN /normal_fun (gtr0_norm e0) subrr. - by rewrite expr0n/= oppr0 mul0r expR0 mulr1 subrr normr0 ltW. - + apply/(@cvgrPdist_le _ R^o) => eps eps0. + near: t; apply: nbhs_right_lt. + by rewrite ltrBlDr -addrA ltrDl// addr_gt0. + rewrite addrAC subrr sub0r normrN (gtr0_norm e0). + by rewrite normal_peak_funxx subrr normr0 ltW. + + apply/cvgrPdist_le => eps eps0. near=> t. - rewrite /g' ballFE_ge// ifT; last first. + rewrite /g' (negPf (ballFE_ge _))// ifT; last first. rewrite ball_itv inE/= in_itv/=; apply/andP => []; split => //. - near: t. - apply: nbhs_left_gt. - by rewrite -(opprK e) ler_ltB// opprK gtrN. - rewrite /normal_pdf0. - rewrite addrAC subrr add0r /normal_fun (gtr0_norm e0) subrr expr0n oppr0. - by rewrite mul0r expR0 mulr1 subrr normr0 ltW. + near: t; apply: nbhs_left_gt. + by rewrite ltrBlDr -addrA ltrDl// addr_gt0. + rewrite addrAC subrr add0r (gtr0_norm e0). + by rewrite normal_peak_funxx subrr normr0 ltW. - apply/continuous_within_itvcyP; split. + move=> x. rewrite in_itv/= andbT => aex. - apply/(@cvgrPdist_le _ R^o) => /= eps eps0. + apply/cvgrPdist_le => /= eps eps0. near=> t. have tae : a + e < t by near: t; exact: lt_nbhsr. - rewrite /g'. - rewrite !ballFE_ge ?(@ltW _ _ (a + e)%E)//. - rewrite /normal_pdf0 /normal_fun !aDe// ?(@ltW _ _ (a + e)). - move=> {tae}; near: t. - move: eps eps0. - apply/(@cvgrPdist_le _ R^o). - exact: continuous_normal_pdf0. - + apply/(@cvgrPdist_le _ R^o) => eps eps0. + rewrite /g' !(negPf (ballFE_ge (ltW _)))//. + rewrite /normal_peak_fun /normal_fun !aDe// -!/(normal_peak_fun _ _ _). + near: t. + apply/cvgrPdist_le : eps eps0. + exact: continuous_normal_peak_fun. + + apply/cvgrPdist_le => eps eps0. near=> t. - rewrite /g' !ballFE_ge//. + rewrite /g' !(negPf (ballFE_ge _))//. rewrite addrAC subrr add0r (ger0_norm e0)//. - rewrite /normal_pdf0 /normal_fun subrr -(subrr (a + e)). - rewrite aDe//. + rewrite /normal_peak_fun /normal_fun subrr -(subrr (a + e)) aDe//. + rewrite -!/(normal_peak_fun _ _ _). near: t. - move: eps eps0. - apply/cvgrPdist_le. + apply/cvgrPdist_le : eps eps0. apply: cvg_at_right_filter. - exact: continuous_normal_pdf0. + under eq_fun do rewrite -/(normal_peak_fun _ _ _). + exact: continuous_normal_peak_fun. Unshelve. all: end_near. Qed. Let gE_Ny a e : 0 <= e -> @@ -3647,10 +3544,10 @@ rewrite ge0_integration_by_substitution_shift_itvNy => /=; first last. under eq_integral. move=> x. rewrite inE/= in_itv/= => xae. - rewrite /g' ballFE_le//; last exact: lerB. + rewrite /g' (negPf (ballFE_le _)); last exact: lerB. rewrite -(normrN (x - e - a)) !opprB addrA. have /normr_idP -> : 0 <= a + e - x by rewrite subr_ge0 ler_wpDr. - rewrite /normal_pdf0 /normal_fun. + rewrite /normal_peak_fun /normal_fun. rewrite -(addrAC _ (- x)) addrK. rewrite -(sqrrN (a - x)) opprB. over. @@ -3670,9 +3567,9 @@ rewrite ge0_integration_by_substitution_shift_itvy => /=; first last. under eq_integral. move=> x. rewrite inE/= in_itv/= andbT => aex. - rewrite /g' ballFE_ge//; last exact: lerD. + rewrite /g' (negPf (ballFE_ge _))//; last exact: lerD. have /normr_idP -> : 0 <= x + e - a by rewrite subr_ge0 ler_wpDr. - rewrite /normal_pdf0 /normal_fun -(addrAC _ (- a)) addrK. + rewrite /normal_peak_fun /normal_fun -(addrAC _ (- a)) addrK. over. rewrite /=. by apply: eq_integral => /= x xay; rewrite /normal_pdf (negbTE s0). @@ -3685,17 +3582,16 @@ move=> mV a. near (0 : R)^'+ => e. set g := g' a e. have mg := mg' a e. -apply: (@continuity_under_integral _ _ _ mu _ _ _ _ (a - e) (a + e) _ _ _ g) => //=. -- rewrite in_itv/=. - by rewrite ltrDl gtrBl andbb. +apply: (@continuity_under_integral _ _ _ mu _ _ _ (a - e) (a + e) _ _ g) => //=. - move=> x _. by apply: (integrableS measurableT) => //=; exact: integrable_normal_pdf. -- apply/aeW => y _ x. +- apply/aeW => y Vy x. + rewrite inE/= in_itv/= => /andP[aex xae]. under [X in _ _ X]eq_fun. move=> x0. - rewrite normal_pdfE// /normal_pdf0 /normal_fun -(sqrrN (y - _)) opprB. + rewrite normal_pdfE// -(sqrrN (y - _)) opprB. over. - exact: continuous_normal_pdf0. + exact: continuous_normal_peak_fun. - apply: (integrableS measurableT) => //=. apply/integrableP; split; first exact/measurable_EFinP. rewrite -(setUv (ball a e)) ge0_integral_setU//=; last 4 first. @@ -3735,29 +3631,30 @@ apply: (@continuity_under_integral _ _ _ mu _ _ _ _ (a - e) (a + e) _ _ _ g) => apply/abse_integralP => //=. by apply/measurable_EFinP; exact: measurable_normal_pdf. by rewrite integral_normal_pdf ltry. -move=> x; rewrite in_itv/= => /andP[aex xae]. -apply: aeW => /= y Vy. -rewrite ger0_norm; last exact: normal_pdf_ge0. -rewrite normal_pdfE// /g /g'. -case: ifPn => [_|]; first exact: normal_pdf0_ub. -rewrite notin_setE/= ball_itv/= in_itv/= => aey. -rewrite /normal_pdf0 ler_pM//. -rewrite ler_expR !mulNr lerN2 ler_pM //. - exact: sqr_ge0. - by rewrite invr_ge0 mulrn_wge0// sqr_ge0. -move: aey; move/negP/nandP; rewrite -!leNgt => -[yae|aey]. - rewrite -normrN opprB ger0_norm; last first. - by rewrite subr_ge0 (le_trans yae)// gerBl. - rewrite -[leRHS]sqrrN opprB ler_sqr ?nnegrE; first last. - rewrite subr_ge0 ltW// (le_lt_trans yae)//. - by rewrite addrAC subr_ge0. - by rewrite addrAC lerD2r ltW. -rewrite ger0_norm; last first. - by rewrite subr_ge0 (le_trans _ aey)// lerDl. -rewrite ler_sqr ?nnegrE; last 2 first. - by rewrite -addrA -opprD subr_ge0. - by rewrite subr_ge0 (le_trans _ aey)// ltW. -by rewrite -addrA -opprD lerD2l lerN2 ltW. +- move=> x; rewrite in_itv/= => /andP[aex xae]. + apply: aeW => /= y Vy. + rewrite ger0_norm; last exact: normal_pdf_ge0. + rewrite normal_pdfE// /g /g'. + case: ifPn => [_|]; first exact: normal_peak_fun_ub. + rewrite notin_setE/= ball_itv/= in_itv/= => aey. + rewrite /normal_peak_fun ler_pM//. + rewrite ler_expR !mulNr lerN2 ler_pM //. + exact: sqr_ge0. + by rewrite invr_ge0 mulrn_wge0// sqr_ge0. + move: aey; move/negP/nandP; rewrite -!leNgt => -[yae|aey]. + rewrite -normrN opprB ger0_norm; last first. + by rewrite subr_ge0 (le_trans yae)// gerBl. + rewrite -[leRHS]sqrrN opprB ler_sqr ?nnegrE; first last. + rewrite subr_ge0 ltW// (le_lt_trans yae)//. + by rewrite addrAC subr_ge0. + by rewrite addrAC lerD2r ltW. + rewrite ger0_norm; last first. + by rewrite subr_ge0 (le_trans _ aey)// lerDl. + rewrite ler_sqr ?nnegrE; last 2 first. + by rewrite -addrA -opprD subr_ge0. + by rewrite subr_ge0 (le_trans _ aey)// ltW. + by rewrite -addrA -opprD lerD2l lerN2 ltW. +- by rewrite -ball_itv inE; exact: ballxx. Unshelve. end_near. Qed. Lemma measurable_normal_prob2 : @@ -3765,7 +3662,7 @@ Lemma measurable_normal_prob2 : Proof. apply: (@measurability _ _ _ _ _ _ (@pset _ _ _ : set (set (pprobability _ R)))) => //. -move=>_ -[_ [r r01] [Ys mYs <-]] <-. +move=>_ -[_ [r r01] [Ys mYs <-]] <-. apply: emeasurable_fun_infty_o => //=. under [X in _ _ X]eq_fun. move=> x. @@ -3775,6 +3672,7 @@ under [X in _ _ X]eq_fun. by rewrite le_measure ?inE. exact: (le_lt_trans (probability_le1 _ _) (ltey _)). over. +rewrite -/(measurable_fun _ _) {r r01}. apply/measurable_EFinP; apply: continuous_measurable_fun. exact: normal_prob_continuous. Qed.