Skip to content

Commit

Permalink
Added essential_supremum_ereal
Browse files Browse the repository at this point in the history
Added documentation and changelog

Co-authored-by: @jmmarulang
  • Loading branch information
hoheinzollern committed Oct 22, 2024
1 parent 6dfceaa commit a6d72f6
Show file tree
Hide file tree
Showing 2 changed files with 43 additions and 0 deletions.
22 changes: 22 additions & 0 deletions CHANGELOG_PR_ESS_SUP.md
Original file line number Diff line number Diff line change
@@ -0,0 +1,22 @@
# Changelog (unreleased)

## [Unreleased]

### Added
- in `measure.v`:
+ definition `ess_supe`
+ lemma `ess_supe_ge0`

### Changed

### Renamed

### Generalized

### Deprecated

### Removed

### Infrastructure

### Misc
21 changes: 21 additions & 0 deletions theories/measure.v
Original file line number Diff line number Diff line change
Expand Up @@ -277,6 +277,8 @@ From HB Require Import structures.
(* m1 `<< m2 == m1 is absolutely continuous w.r.t. m2 or m2 dominates m1 *)
(* ess_sup f == essential supremum of the function f : T -> R where T is a *)
(* semiRingOfSetsType and R is a realType *)
(* ess_supe f == essential supremum of the function f : T -> \bar R where *)
(* T is a semiRingOfSetsType and R is a realType *)
(* ``` *)
(* *)
(******************************************************************************)
Expand Down Expand Up @@ -5264,3 +5266,22 @@ by apply/seteqP; split => // x _ /=; rewrite in_itv/= (lt_le_trans _ (f0 x)).
Qed.

End essential_supremum.

Section essential_supremum_ereal.
Context d {T : semiRingOfSetsType d} {R : realType}.
Variable mu : {measure set T -> \bar R}.
Implicit Types f : T -> \bar R.

Definition ess_supe f :=
ereal_inf ([set r | mu (f @^-1` `]r, +oo[) = 0]).

Lemma ess_supe_ge0 f : 0 < mu [set: T] -> (forall t, 0 <= f t) ->
0 <= ess_supe f.
Proof.
move=> muT f0; apply: lb_ereal_inf.
by case=> //= [r|] /eqP rf; rewrite leNgt;
apply/negP => r0; apply/negP : rf; rewrite gt_eqF// (_ : _ @^-1` _ = setT)//;
apply/seteqP; split => // x _ /=; rewrite in_itv/= (lt_le_trans _ (f0 x)).
Qed.

End essential_supremum_ereal.

0 comments on commit a6d72f6

Please sign in to comment.