diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean index 0d80172e45051..5ed34323264cf 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Real.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.MeasureTheory.Constructions.BorelSpace.Order +import Mathlib.MeasureTheory.MeasurableSpace.Prod /-! # Borel (measurable) spaces ℝ, ℝ≥0, ℝ≥0∞ @@ -465,6 +466,72 @@ _root_.measurable_of_tendsto_nnreal := NNReal.measurable_of_tendsto end NNReal +namespace EReal + +lemma measurableEmbedding_coe : MeasurableEmbedding Real.toEReal := + isOpenEmbedding_coe.measurableEmbedding + +instance : MeasurableAdd₂ EReal := ⟨EReal.lowerSemicontinuous_add.measurable⟩ + +section MeasurableMul + +variable {α β γ : Type*} {mα : MeasurableSpace α} {mβ : MeasurableSpace β} {mγ : MeasurableSpace γ} + +lemma measurable_of_real_prod {f : EReal × β → γ} + (h_real : Measurable fun p : ℝ × β ↦ f (p.1, p.2)) + (h_bot : Measurable fun x ↦ f (⊥, x)) (h_top : Measurable fun x ↦ f (⊤, x)) : Measurable f := + .of_union₃_range_cover (measurableEmbedding_prod_mk_left _) (measurableEmbedding_prod_mk_left _) + (measurableEmbedding_coe.prod_mk .id) (by simp [-univ_subset_iff, subset_def, EReal.forall]) + h_bot h_top h_real + +lemma measurable_of_real_real {f : EReal × EReal → β} + (h_real : Measurable fun p : ℝ × ℝ ↦ f (p.1, p.2)) + (h_bot_left : Measurable fun r : ℝ ↦ f (⊥, r)) + (h_top_left : Measurable fun r : ℝ ↦ f (⊤, r)) + (h_bot_right : Measurable fun r : ℝ ↦ f (r, ⊥)) + (h_top_right : Measurable fun r : ℝ ↦ f (r, ⊤)) : + Measurable f := by + refine measurable_of_real_prod ?_ ?_ ?_ + · refine measurable_swap_iff.mp <| measurable_of_real_prod ?_ h_bot_right h_top_right + exact h_real.comp measurable_swap + · exact measurable_of_measurable_real h_bot_left + · exact measurable_of_measurable_real h_top_left + +private lemma measurable_const_mul (c : EReal) : Measurable fun (x : EReal) ↦ c * x := by + refine measurable_of_measurable_real ?_ + have h1 : (fun (p : ℝ) ↦ (⊥ : EReal) * p) + = fun p ↦ if p = 0 then (0 : EReal) else (if p < 0 then ⊤ else ⊥) := by + ext p + split_ifs with h1 h2 + · simp [h1] + · rw [bot_mul_coe_of_neg h2] + · rw [bot_mul_coe_of_pos] + exact lt_of_le_of_ne (not_lt.mp h2) (Ne.symm h1) + have h2 : Measurable fun (p : ℝ) ↦ if p = 0 then (0 : EReal) else if p < 0 then ⊤ else ⊥ := by + refine Measurable.piecewise (measurableSet_singleton _) measurable_const ?_ + exact Measurable.piecewise measurableSet_Iio measurable_const measurable_const + induction c with + | h_bot => rwa [h1] + | h_real c => exact (measurable_id.const_mul _).coe_real_ereal + | h_top => + simp_rw [← neg_bot, neg_mul] + apply Measurable.neg + rwa [h1] + +instance : MeasurableMul₂ EReal := by + refine ⟨measurable_of_real_real ?_ ?_ ?_ ?_ ?_⟩ + · exact (measurable_fst.mul measurable_snd).coe_real_ereal + · exact (measurable_const_mul _).comp measurable_coe_real_ereal + · exact (measurable_const_mul _).comp measurable_coe_real_ereal + · simp_rw [mul_comm _ ⊥] + exact (measurable_const_mul _).comp measurable_coe_real_ereal + · simp_rw [mul_comm _ ⊤] + exact (measurable_const_mul _).comp measurable_coe_real_ereal + +end MeasurableMul + +end EReal + /-- If a function `f : α → ℝ≥0` is measurable and the measure is σ-finite, then there exists spanning measurable sets with finite measure on which `f` is bounded. See also `StronglyMeasurable.exists_spanning_measurableSet_norm_le` for functions into normed diff --git a/Mathlib/Topology/Instances/EReal.lean b/Mathlib/Topology/Instances/EReal.lean index 746c9fd365e5b..3af0b2ab4b18d 100644 --- a/Mathlib/Topology/Instances/EReal.lean +++ b/Mathlib/Topology/Instances/EReal.lean @@ -7,6 +7,7 @@ import Mathlib.Data.Rat.Encodable import Mathlib.Data.Real.EReal import Mathlib.Topology.Instances.ENNReal import Mathlib.Topology.Order.MonotoneContinuity +import Mathlib.Topology.Semicontinuous /-! # Topological structure on `EReal` @@ -441,4 +442,12 @@ theorem continuousAt_mul {p : EReal × EReal} (h₁ : p.1 ≠ 0 ∨ p.2 ≠ ⊥) exact continuousAt_mul_top_ne_zero h₄ · exact continuousAt_mul_top_top +lemma lowerSemicontinuous_add : LowerSemicontinuous fun p : EReal × EReal ↦ p.1 + p.2 := by + intro x y + by_cases hx₁ : x.1 = ⊥ + · simp [hx₁] + by_cases hx₂ : x.2 = ⊥ + · simp [hx₂] + · exact continuousAt_add (.inr hx₂) (.inl hx₁) |>.lowerSemicontinuousAt _ + end EReal