Skip to content

Commit

Permalink
bernoulli probability measure
Browse files Browse the repository at this point in the history
  • Loading branch information
affeldt-aist committed Apr 26, 2023
1 parent 5a0c68f commit 87540b5
Show file tree
Hide file tree
Showing 4 changed files with 33 additions and 2 deletions.
7 changes: 7 additions & 0 deletions CHANGELOG_UNRELEASED.md
Original file line number Diff line number Diff line change
Expand Up @@ -4,6 +4,13 @@

### Added

- in `signed.v`:
+ definition `onem_NngNum`
- in `measure.v`:
+ definition `bernoulli`, declared as a probability measure instance
- in `itv.v`:
+ canonical `onem_itv01`

### Changed

### Renamed
Expand Down
5 changes: 3 additions & 2 deletions theories/itv.v
Original file line number Diff line number Diff line change
Expand Up @@ -843,6 +843,9 @@ Lemma inum_lt : {mono inum : x y / (x < y)%O}. Proof. by []. Qed.

End Morph.

Canonical onem_itv01 {R : realDomainType} (p : {i01 R}) : {i01 R} :=
@Itv.mk _ _ (onem p%:inum) [itv of 1 - p%:inum].

Section Test1.

Variable R : numDomainType.
Expand Down Expand Up @@ -882,8 +885,6 @@ apply/val_inj => /=.
by rewrite subr0 mulr1 opprB addrCA subrr addr0.
Qed.

Canonical onem_itv01 (p : {i01 R}) : {i01 R} :=
@Itv.mk _ _ (onem p%:inum) [itv of 1 - p%:inum].

Definition s_of_pq' (p q : {i01 R}) : {i01 R} :=
(`1- (`1-(p%:inum) * `1-(q%:inum)))%:i01.
Expand Down
22 changes: 22 additions & 0 deletions theories/measure.v
Original file line number Diff line number Diff line change
Expand Up @@ -3,6 +3,7 @@ From mathcomp Require Import all_ssreflect all_algebra finmap.
From mathcomp.classical Require Import boolp classical_sets functions.
From mathcomp.classical Require Import cardinality fsbigop mathcomp_extra.
Require Import reals ereal signed topology normedtype sequences esum numfun.
Require Import itv.
From HB Require Import structures.

(******************************************************************************)
Expand Down Expand Up @@ -137,6 +138,7 @@ From HB Require Import structures.
(* probability T R == probability measure over the measurableType T with *)
(* value in R : realType *)
(* Measure_isProbability == factor for probability measures *)
(* bernoulli p == the p-Bernoulli probability where p : {i01 R} *)
(* *)
(* mu.-negligible A == A is mu negligible *)
(* {ae mu, forall x, P x} == P holds almost everywhere for the measure mu, *)
Expand Down Expand Up @@ -2826,6 +2828,26 @@ HB.instance Definition _ x :=

End pdirac.

Section bernoulli.
Variables (R : realType) (p : {i01 R}).

Definition bernoulli : set _ -> \bar R := measure_add
(mscale (NngNum (itv_ge0 p)) (dirac true))
(mscale (NngNum (itv_ge0 ((`1-(p%:inum))%:i01))) (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 sigma_finite_counting (R : realType) :
sigma_finite [set: nat] (@counting _ R).
Proof.
Expand Down
1 change: 1 addition & 0 deletions theories/signed.v
Original file line number Diff line number Diff line change
Expand Up @@ -106,6 +106,7 @@ From mathcomp.classical Require Import mathcomp_extra.
(* --> Canonical instances are also provided according to types, as a *)
(* fallback when no known operator appears in the expression. Look to *)
(* nat_snum below for an example on how to add your favorite type. *)
(* *)
(******************************************************************************)

Reserved Notation "{ 'compare' x0 & nz & cond }"
Expand Down

0 comments on commit 87540b5

Please sign in to comment.