Skip to content

Commit

Permalink
feat: for an open-positive measure, an open/closed subset is almost e…
Browse files Browse the repository at this point in the history
…mpty/full iff it is actually empty/full (#5746)

Also invert the import order so that `MeasureTheory.Measure.OpenPos` imports `MeasureTheory.Constructions.BorelSpace.Basic` rather than the other way around.
  • Loading branch information
ocfnash committed Jul 14, 2023
1 parent 1ce44f1 commit a1a257d
Show file tree
Hide file tree
Showing 3 changed files with 37 additions and 10 deletions.
9 changes: 0 additions & 9 deletions Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean
Expand Up @@ -12,7 +12,6 @@ import Mathlib.Analysis.Normed.Group.Basic
import Mathlib.MeasureTheory.Function.AEMeasurableSequence
import Mathlib.MeasureTheory.Group.Arithmetic
import Mathlib.MeasureTheory.Lattice
import Mathlib.MeasureTheory.Measure.OpenPos
import Mathlib.Topology.Algebra.Order.LiminfLimsup
import Mathlib.Topology.ContinuousFunction.Basic
import Mathlib.Topology.Instances.EReal
Expand Down Expand Up @@ -839,14 +838,6 @@ theorem ClosedEmbedding.measurable {f : α → γ} (hf : ClosedEmbedding f) : Me
hf.continuous.measurable
#align closed_embedding.measurable ClosedEmbedding.measurable

theorem Continuous.isOpenPosMeasure_map {f : β → γ} (hf : Continuous f)
(hf_surj : Function.Surjective f) {μ : Measure β} [μ.IsOpenPosMeasure] :
(Measure.map f μ).IsOpenPosMeasure := by
refine' ⟨fun U hUo hUne => _⟩
rw [Measure.map_apply hf.measurable hUo.measurableSet]
exact (hUo.preimage hf).measure_ne_zero μ (hf_surj.nonempty_preimage.mpr hUne)
#align continuous.is_open_pos_measure_map Continuous.isOpenPosMeasure_map

/-- If a function is defined piecewise in terms of functions which are continuous on their
respective pieces, then it is measurable. -/
theorem ContinuousOn.measurable_piecewise {f g : α → γ} {s : Set α} [∀ j : α, Decidable (j ∈ s)]
Expand Down
1 change: 1 addition & 0 deletions Mathlib/MeasureTheory/Function/LpSpace.lean
Expand Up @@ -10,6 +10,7 @@ Authors: Rémy Degenne, Sébastien Gouëzel
-/
import Mathlib.Analysis.Normed.Group.Hom
import Mathlib.MeasureTheory.Function.LpSeminorm
import Mathlib.MeasureTheory.Measure.OpenPos
import Mathlib.Topology.ContinuousFunction.Compact

/-!
Expand Down
37 changes: 36 additions & 1 deletion Mathlib/MeasureTheory/Measure/OpenPos.lean
Expand Up @@ -9,6 +9,7 @@ Authors: Yury Kudryashov
! if you have ported upstream changes.
-/
import Mathlib.MeasureTheory.Measure.MeasureSpace
import Mathlib.MeasureTheory.Constructions.BorelSpace.Basic

/-!
# Measures positive on nonempty opens
Expand Down Expand Up @@ -39,7 +40,7 @@ class IsOpenPosMeasure : Prop where
open_pos : ∀ U : Set X, IsOpen U → U.Nonempty → μ U ≠ 0
#align measure_theory.measure.is_open_pos_measure MeasureTheory.Measure.IsOpenPosMeasure

variable [IsOpenPosMeasure μ] {s U : Set X} {x : X}
variable [IsOpenPosMeasure μ] {s U F : Set X} {x : X}

theorem _root_.IsOpen.measure_ne_zero (hU : IsOpen U) (hne : U.Nonempty) : μ U ≠ 0 :=
IsOpenPosMeasure.open_pos U hU hne
Expand Down Expand Up @@ -80,10 +81,33 @@ theorem _root_.LE.le.isOpenPosMeasure (h : μ ≤ ν) : IsOpenPosMeasure ν :=
h.absolutelyContinuous.isOpenPosMeasure
#align has_le.le.is_open_pos_measure LE.le.isOpenPosMeasure

theorem _root_.IsOpen.measure_zero_iff_eq_empty (hU : IsOpen U) :
μ U = 0 ↔ U = ∅ :=
fun h ↦ (hU.measure_eq_zero_iff μ).mp h, fun h ↦ by simp [h]⟩

theorem _root_.IsOpen.ae_eq_empty_iff_eq (hU : IsOpen U) :
U =ᵐ[μ] (∅ : Set X) ↔ U = ∅ := by
rw [ae_eq_empty, hU.measure_zero_iff_eq_empty]

theorem _root_.IsOpen.eq_empty_of_measure_zero (hU : IsOpen U) (h₀ : μ U = 0) : U = ∅ :=
(hU.measure_eq_zero_iff μ).mp h₀
#align is_open.eq_empty_of_measure_zero IsOpen.eq_empty_of_measure_zero

theorem _root_.IsClosed.ae_eq_univ_iff_eq (hF : IsClosed F) :
F =ᵐ[μ] univ ↔ F = univ := by
refine' ⟨fun h ↦ _, fun h ↦ by rw [h]⟩
rwa [ae_eq_univ, hF.isOpen_compl.measure_eq_zero_iff μ, compl_empty_iff] at h

theorem _root_.IsClosed.measure_eq_univ_iff_eq [OpensMeasurableSpace X] [IsFiniteMeasure μ]
(hF : IsClosed F) :
μ F = μ univ ↔ F = univ := by
rw [← ae_eq_univ_iff_measure_eq hF.measurableSet.nullMeasurableSet, hF.ae_eq_univ_iff_eq]

theorem _root_.IsClosed.measure_eq_one_iff_eq_univ [OpensMeasurableSpace X] [IsProbabilityMeasure μ]
(hF : IsClosed F) :
μ F = 1 ↔ F = univ := by
rw [← measure_univ (μ := μ), hF.measure_eq_univ_iff_eq]

theorem interior_eq_empty_of_null (hs : μ s = 0) : interior s = ∅ :=
isOpen_interior.eq_empty_of_measure_zero <| measure_mono_null interior_subset hs
#align measure_theory.measure.interior_eq_empty_of_null MeasureTheory.Measure.interior_eq_empty_of_null
Expand Down Expand Up @@ -125,6 +149,17 @@ theorem _root_.Continuous.ae_eq_iff_eq {f g : X → Y} (hf : Continuous f) (hg :
fun h => eq_of_ae_eq h hf hg, fun h => h ▸ EventuallyEq.rfl⟩
#align continuous.ae_eq_iff_eq Continuous.ae_eq_iff_eq

variable {μ}

theorem _root_.Continuous.isOpenPosMeasure_map [OpensMeasurableSpace X]
{Z : Type _} [TopologicalSpace Z] [MeasurableSpace Z] [BorelSpace Z]
{f : X → Z} (hf : Continuous f) (hf_surj : Function.Surjective f) :
(Measure.map f μ).IsOpenPosMeasure := by
refine' ⟨fun U hUo hUne => _⟩
rw [Measure.map_apply hf.measurable hUo.measurableSet]
exact (hUo.preimage hf).measure_ne_zero μ (hf_surj.nonempty_preimage.mpr hUne)
#align continuous.is_open_pos_measure_map Continuous.isOpenPosMeasure_map

end Basic

section LinearOrder
Expand Down

0 comments on commit a1a257d

Please sign in to comment.