Skip to content

Commit

Permalink
instances for measures (#1419)
Browse files Browse the repository at this point in the history
* instances for measures

- restriction of finite measure is finite
- scaling of finite measure is finite
  • Loading branch information
affeldt-aist authored Dec 3, 2024
1 parent 61b7572 commit 7ced470
Showing 1 changed file with 33 additions and 0 deletions.
33 changes: 33 additions & 0 deletions theories/measure.v
Original file line number Diff line number Diff line change
Expand Up @@ -3394,6 +3394,39 @@ HB.instance Definition _ := @isFinite.Build d T R k finite.

HB.end.

Section finite_restr.
Context d (T : measurableType d) (R : realType).
Variables (mu : {finite_measure set T -> \bar R}) (D : set T).
Hypothesis mD : measurable D.

Local Notation restr := (mrestr mu mD).

Let fin_num_restr : fin_num_fun restr.
Proof.
move=> A mA; have := fin_num_measure mu A mA.
rewrite !ge0_fin_numE//=; apply: le_lt_trans.
by rewrite /mrestr; apply: le_measure => //; rewrite inE//=; exact: measurableI.
Qed.

HB.instance Definition _ := @Measure_isFinite.Build _ T _ restr fin_num_restr.

End finite_restr.

Section finite_mscale.
Context d (T : measurableType d) (R : realType).
Variables (mu : {finite_measure set T -> \bar R}) (r : {nonneg R}).

Local Notation scale := (mscale r mu).

Let fin_num_scale : fin_num_fun scale.
Proof.
by move=> A mA; have muA := fin_num_measure mu A mA; rewrite fin_numM.
Qed.

HB.instance Definition _ := @Measure_isFinite.Build _ T _ scale fin_num_scale.

End finite_mscale.

HB.factory Record Measure_isSFinite d (T : sigmaRingType d)
(R : realType) (k : set T -> \bar R) of isMeasure _ _ _ k := {
s_finite : exists s : {finite_measure set T -> \bar R}^nat,
Expand Down

0 comments on commit 7ced470

Please sign in to comment.