Skip to content

Commit

Permalink
nonneg 2/7
Browse files Browse the repository at this point in the history
  • Loading branch information
AyumuSaito committed Aug 8, 2022
1 parent b5ed9ac commit aa653a6
Showing 1 changed file with 17 additions and 5 deletions.
22 changes: 17 additions & 5 deletions theories/kernel.v
Original file line number Diff line number Diff line change
Expand Up @@ -997,15 +997,20 @@ End discrete_measurable_bool.

Section nonneg_constants.
Variable R : realType.
(*
Let twoseven_proof : (0 <= 2 / 7 :> R)%R.
Proof. by rewrite divr_ge0// ler0n. Qed.
*)

Definition twoseven : {nonneg R} := NngNum twoseven_proof.
(* Check (2%:R / 7%:R)%:nng. *)

(* Definition twoseven : {nonneg R} := (2%:R / 7%:R)%:nng. *)
(*
Let fiveseven_proof : (0 <= 5 / 7 :> R)%R.
Proof. by rewrite divr_ge0// ler0n. Qed.
Definition fiveseven : {nonneg R} := NngNum fiveseven_proof.
*)

End nonneg_constants.

Expand All @@ -1020,13 +1025,20 @@ Proof. by rewrite /= diracE in_setT. Qed.
Section bernoulli27.
Variable R : realType.

Local Open Scope ring_scope.
Notation "'2/7'" := (2%:R / 7%:R)%:nng.
Definition twoseven : {nonneg R} := (2%:R / 7%:R)%:nng.
Definition fiveseven : {nonneg R} := (5%:R / 7%:R)%:nng.

Definition bernoulli27 : set _ -> \bar R :=
measure_add
[the measure _ _ of mscale (twoseven R) [the measure _ _ of dirac true]]
[the measure _ _ of mscale (fiveseven R) [the measure _ _ of dirac false]].
[the measure _ _ of mscale twoseven [the measure _ _ of dirac true]]
[the measure _ _ of mscale fiveseven [the measure _ _ of dirac false]].

HB.instance Definition _ := Measure.on bernoulli27.

Local Close Scope ring_scope.

Lemma bernoulli27_setT : bernoulli27 [set: _] = 1.
Proof.
rewrite /bernoulli27/= /measure_add/= /msum 2!big_ord_recr/= big_ord0 add0e/=.
Expand Down Expand Up @@ -1522,8 +1534,8 @@ Variables (R : realType) (d : _) (T : measurableType d).
let _ = score (1/4! r^4 e^-r) in
return x *)

Definition k3' : T * bool -> R := cst 3.
Definition k10' : T * bool -> R := cst 10.
Definition k3' : T * bool -> R := cst 3%:R.
Definition k10' : T * bool -> R := cst 10%:R.

Lemma mk3 : measurable_fun setT k3'.
exact: measurable_fun_cst.
Expand Down

0 comments on commit aa653a6

Please sign in to comment.