Skip to content

Commit

Permalink
feat: essSup of the uniform measure
Browse files Browse the repository at this point in the history
I don't care about how to spell the uniform measure, but I need *something*.

From LeanAPAP
  • Loading branch information
YaelDillies committed Oct 6, 2024
1 parent 7c4340d commit 24d6456
Showing 1 changed file with 9 additions and 3 deletions.
12 changes: 9 additions & 3 deletions Mathlib/MeasureTheory/Function/EssSup.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@ Authors: Rémy Degenne
import Mathlib.MeasureTheory.Constructions.BorelSpace.Order
import Mathlib.MeasureTheory.Measure.Count
import Mathlib.Order.Filter.ENNReal
import Mathlib.Probability.ConditionalProbability

/-!
# Essential supremum and infimum
Expand All @@ -28,9 +29,8 @@ sense). We do not define that quantity here, which is simply the supremum of a m
-/


open MeasureTheory Filter Set TopologicalSpace

open ENNReal MeasureTheory NNReal
open Filter MeasureTheory ProbabilityTheory Set TopologicalSpace
open scoped ENNReal NNReal

variable {α β : Type*} {m : MeasurableSpace α} {μ ν : Measure α}

Expand Down Expand Up @@ -94,6 +94,12 @@ variable [MeasurableSingletonClass α]
@[simp] lemma essInf_count_eq_ciInf (hf : BddBelow (Set.range f)) :
essInf f .count = ⨅ a, f a := essInf_eq_ciInf (by simp) hf

@[simp] lemma essSup_cond_count_eq_ciSup [Finite α] (hf : BddAbove (Set.range f)) :
essSup f .count[|.univ] = ⨆ a, f a := essSup_eq_ciSup (by simp [cond_apply, Set.finite_univ]) hf

@[simp] lemma essInf_cond_count_eq_ciInf [Finite α] (hf : BddBelow (Set.range f)) :
essInf f .count[|.univ] = ⨅ a, f a := essInf_eq_ciInf (by simp [cond_apply, Set.finite_univ]) hf

end ConditionallyCompleteLattice

section ConditionallyCompleteLinearOrder
Expand Down

0 comments on commit 24d6456

Please sign in to comment.