Documentation

Mathlib.MeasureTheory.OuterMeasure.Induced

Induced Outer Measure #

We can extend a function defined on a subset of Set α to an outer measure. The underlying function is called extend, and the measure it induces is called inducedOuterMeasure.

Some lemmas below are proven twice, once in the general case, and one where the function m is only defined on measurable sets (i.e. when P = MeasurableSet). In the latter cases, we can remove some hypotheses in the statement. The general version has the same name, but with a prime at the end.

Tags #

outer measure

def MeasureTheory.extend {α : Type u_1} {P : αProp} (m : (s : α) → P sENNReal) (s : α) :

We can trivially extend a function defined on a subclass of objects (with codomain ℝ≥0∞) to all objects by defining it to be on the objects not in the class.

Equations
Instances For
    theorem MeasureTheory.extend_eq {α : Type u_1} {P : αProp} (m : (s : α) → P sENNReal) {s : α} (h : P s) :
    theorem MeasureTheory.extend_eq_top {α : Type u_1} {P : αProp} (m : (s : α) → P sENNReal) {s : α} (h : ¬P s) :
    theorem MeasureTheory.smul_extend {α : Type u_1} {P : αProp} (m : (s : α) → P sENNReal) {R : Type u_2} [Zero R] [SMulWithZero R ENNReal] [IsScalarTower R ENNReal ENNReal] [NoZeroSMulDivisors R ENNReal] {c : R} (hc : c 0) :
    c MeasureTheory.extend m = MeasureTheory.extend fun (s : α) (h : P s) => c m s h
    theorem MeasureTheory.le_extend {α : Type u_1} {P : αProp} (m : (s : α) → P sENNReal) {s : α} (h : P s) :
    theorem MeasureTheory.extend_congr {α : Type u_1} {P : αProp} (m : (s : α) → P sENNReal) {β : Type u_2} {Pb : βProp} {mb : (s : β) → Pb sENNReal} {sa : α} {sb : β} (hP : P sa Pb sb) (hm : ∀ (ha : P sa) (hb : Pb sb), m sa ha = mb sb hb) :
    @[simp]
    theorem MeasureTheory.extend_top {α : Type u_2} {P : αProp} :
    (MeasureTheory.extend fun (x : α) (x : P x) => ) =
    theorem MeasureTheory.extend_empty {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} (P0 : P ) (m0 : m P0 = 0) :
    theorem MeasureTheory.extend_iUnion_nat {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) {f : Set α} (hm : ∀ (i : ), P (f i)) (mU : m (⋃ (i : ), f i) = ∑' (i : ), m (f i) ) :
    MeasureTheory.extend m (⋃ (i : ), f i) = ∑' (i : ), MeasureTheory.extend m (f i)
    theorem MeasureTheory.extend_iUnion_le_tsum_nat' {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) (msU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), P (f i)), m (⋃ (i : ), f i) ∑' (i : ), m (f i) ) (s : Set α) :
    MeasureTheory.extend m (⋃ (i : ), s i) ∑' (i : ), MeasureTheory.extend m (s i)
    theorem MeasureTheory.extend_mono' {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} (m_mono : ∀ ⦃s₁ s₂ : Set α⦄ (hs₁ : P s₁) (hs₂ : P s₂), s₁ s₂m s₁ hs₁ m s₂ hs₂) ⦃s₁ : Set α ⦃s₂ : Set α (h₁ : P s₁) (hs : s₁ s₂) :
    theorem MeasureTheory.extend_iUnion {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} (P0 : P ) (m0 : m P0 = 0) (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) (mU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), P (f i)), Pairwise (Disjoint on f)m (⋃ (i : ), f i) = ∑' (i : ), m (f i) ) {β : Type u_2} [Countable β] {f : βSet α} (hd : Pairwise (Disjoint on f)) (hm : ∀ (i : β), P (f i)) :
    MeasureTheory.extend m (⋃ (i : β), f i) = ∑' (i : β), MeasureTheory.extend m (f i)
    theorem MeasureTheory.extend_union {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} (P0 : P ) (m0 : m P0 = 0) (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) (mU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), P (f i)), Pairwise (Disjoint on f)m (⋃ (i : ), f i) = ∑' (i : ), m (f i) ) {s₁ : Set α} {s₂ : Set α} (hd : Disjoint s₁ s₂) (h₁ : P s₁) (h₂ : P s₂) :
    def MeasureTheory.inducedOuterMeasure {α : Type u_1} {P : Set αProp} (m : (s : Set α) → P sENNReal) (P0 : P ) (m0 : m P0 = 0) :

    Given an arbitrary function on a subset of sets, we can define the outer measure corresponding to it (this is the unique maximal outer measure that is at most m on the domain of m).

    Equations
    Instances For
      theorem MeasureTheory.le_inducedOuterMeasure {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} {P0 : P } {m0 : m P0 = 0} {μ : MeasureTheory.OuterMeasure α} :
      μ MeasureTheory.inducedOuterMeasure m P0 m0 ∀ (s : Set α) (hs : P s), μ s m s hs
      theorem MeasureTheory.inducedOuterMeasure_union_of_false_of_nonempty_inter {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} {P0 : P } {m0 : m P0 = 0} {s : Set α} {t : Set α} (h : ∀ (u : Set α), (s u).Nonempty(t u).Nonempty¬P u) :

      If P u is False for any set u that has nonempty intersection both with s and t, then μ (s ∪ t) = μ s + μ t, where μ = inducedOuterMeasure m P0 m0.

      E.g., if α is an (e)metric space and P u = diam u < r, then this lemma implies that μ (s ∪ t) = μ s + μ t on any two sets such that r ≤ edist x y for all x ∈ s and y ∈ t.

      theorem MeasureTheory.inducedOuterMeasure_eq_extend' {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} {P0 : P } {m0 : m P0 = 0} (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) (msU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), P (f i)), m (⋃ (i : ), f i) ∑' (i : ), m (f i) ) (m_mono : ∀ ⦃s₁ s₂ : Set α⦄ (hs₁ : P s₁) (hs₂ : P s₂), s₁ s₂m s₁ hs₁ m s₂ hs₂) {s : Set α} (hs : P s) :
      theorem MeasureTheory.inducedOuterMeasure_eq' {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} {P0 : P } {m0 : m P0 = 0} (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) (msU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), P (f i)), m (⋃ (i : ), f i) ∑' (i : ), m (f i) ) (m_mono : ∀ ⦃s₁ s₂ : Set α⦄ (hs₁ : P s₁) (hs₂ : P s₂), s₁ s₂m s₁ hs₁ m s₂ hs₂) {s : Set α} (hs : P s) :
      theorem MeasureTheory.inducedOuterMeasure_eq_iInf {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} {P0 : P } {m0 : m P0 = 0} (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) (msU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), P (f i)), m (⋃ (i : ), f i) ∑' (i : ), m (f i) ) (m_mono : ∀ ⦃s₁ s₂ : Set α⦄ (hs₁ : P s₁) (hs₂ : P s₂), s₁ s₂m s₁ hs₁ m s₂ hs₂) (s : Set α) :
      (MeasureTheory.inducedOuterMeasure m P0 m0) s = ⨅ (t : Set α), ⨅ (ht : P t), ⨅ (_ : s t), m t ht
      theorem MeasureTheory.inducedOuterMeasure_preimage {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} {P0 : P } {m0 : m P0 = 0} (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) (msU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), P (f i)), m (⋃ (i : ), f i) ∑' (i : ), m (f i) ) (m_mono : ∀ ⦃s₁ s₂ : Set α⦄ (hs₁ : P s₁) (hs₂ : P s₂), s₁ s₂m s₁ hs₁ m s₂ hs₂) (f : α α) (Pm : ∀ (s : Set α), P (f ⁻¹' s) P s) (mm : ∀ (s : Set α) (hs : P s), m (f ⁻¹' s) = m s hs) {A : Set α} :
      theorem MeasureTheory.inducedOuterMeasure_exists_set {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} {P0 : P } {m0 : m P0 = 0} (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) (msU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), P (f i)), m (⋃ (i : ), f i) ∑' (i : ), m (f i) ) (m_mono : ∀ ⦃s₁ s₂ : Set α⦄ (hs₁ : P s₁) (hs₂ : P s₂), s₁ s₂m s₁ hs₁ m s₂ hs₂) {s : Set α} (hs : (MeasureTheory.inducedOuterMeasure m P0 m0) s ) {ε : ENNReal} (hε : ε 0) :
      ∃ (t : Set α), P t s t (MeasureTheory.inducedOuterMeasure m P0 m0) t (MeasureTheory.inducedOuterMeasure m P0 m0) s + ε
      theorem MeasureTheory.inducedOuterMeasure_caratheodory {α : Type u_1} {P : Set αProp} {m : (s : Set α) → P sENNReal} {P0 : P } {m0 : m P0 = 0} (PU : ∀ ⦃f : Set α⦄, (∀ (i : ), P (f i))P (⋃ (i : ), f i)) (msU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), P (f i)), m (⋃ (i : ), f i) ∑' (i : ), m (f i) ) (m_mono : ∀ ⦃s₁ s₂ : Set α⦄ (hs₁ : P s₁) (hs₂ : P s₂), s₁ s₂m s₁ hs₁ m s₂ hs₂) (s : Set α) :

      To test whether s is Carathéodory-measurable we only need to check the sets t for which P t holds. See ofFunction_caratheodory for another way to show the Carathéodory-measurability of s.

      If P is MeasurableSet for some measurable space, then we can remove some hypotheses of the above lemmas.

      theorem MeasureTheory.extend_mono {α : Type u_1} [MeasurableSpace α] {m : (s : Set α) → MeasurableSet sENNReal} (m0 : m = 0) (mU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), MeasurableSet (f i)), Pairwise (Disjoint on f)m (⋃ (i : ), f i) = ∑' (i : ), m (f i) ) {s₁ : Set α} {s₂ : Set α} (h₁ : MeasurableSet s₁) (hs : s₁ s₂) :
      theorem MeasureTheory.extend_iUnion_le_tsum_nat {α : Type u_1} [MeasurableSpace α] {m : (s : Set α) → MeasurableSet sENNReal} (m0 : m = 0) (mU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), MeasurableSet (f i)), Pairwise (Disjoint on f)m (⋃ (i : ), f i) = ∑' (i : ), m (f i) ) (s : Set α) :
      MeasureTheory.extend m (⋃ (i : ), s i) ∑' (i : ), MeasureTheory.extend m (s i)
      theorem MeasureTheory.inducedOuterMeasure_eq_extend {α : Type u_1} [MeasurableSpace α] {m : (s : Set α) → MeasurableSet sENNReal} (m0 : m = 0) (mU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), MeasurableSet (f i)), Pairwise (Disjoint on f)m (⋃ (i : ), f i) = ∑' (i : ), m (f i) ) {s : Set α} (hs : MeasurableSet s) :
      theorem MeasureTheory.inducedOuterMeasure_eq {α : Type u_1} [MeasurableSpace α] {m : (s : Set α) → MeasurableSet sENNReal} (m0 : m = 0) (mU : ∀ ⦃f : Set α⦄ (hm : ∀ (i : ), MeasurableSet (f i)), Pairwise (Disjoint on f)m (⋃ (i : ), f i) = ∑' (i : ), m (f i) ) {s : Set α} (hs : MeasurableSet s) :

      Given an outer measure m we can forget its value on non-measurable sets, and then consider m.trim, the unique maximal outer measure less than that function.

      Equations
      Instances For
        theorem MeasureTheory.OuterMeasure.le_trim_iff {α : Type u_1} [MeasurableSpace α] {m₁ : MeasureTheory.OuterMeasure α} {m₂ : MeasureTheory.OuterMeasure α} :
        m₁ m₂.trim ∀ (s : Set α), MeasurableSet sm₁ s m₂ s
        @[simp]
        theorem MeasureTheory.OuterMeasure.trim_eq {α : Type u_1} [MeasurableSpace α] (m : MeasureTheory.OuterMeasure α) {s : Set α} (hs : MeasurableSet s) :
        m.trim s = m s
        theorem MeasureTheory.OuterMeasure.trim_congr {α : Type u_1} [MeasurableSpace α] {m₁ : MeasureTheory.OuterMeasure α} {m₂ : MeasureTheory.OuterMeasure α} (H : ∀ {s : Set α}, MeasurableSet sm₁ s = m₂ s) :
        m₁.trim = m₂.trim
        theorem MeasureTheory.OuterMeasure.trim_mono {α : Type u_1} [MeasurableSpace α] :
        Monotone MeasureTheory.OuterMeasure.trim
        theorem MeasureTheory.OuterMeasure.trim_anti_measurableSpace {α : Type u_1} (m : MeasureTheory.OuterMeasure α) {m0 : MeasurableSpace α} {m1 : MeasurableSpace α} (h : m0 m1) :
        m.trim m.trim

        OuterMeasure.trim is antitone in the σ-algebra.

        theorem MeasureTheory.OuterMeasure.trim_le_trim_iff {α : Type u_1} [MeasurableSpace α] {m₁ : MeasureTheory.OuterMeasure α} {m₂ : MeasureTheory.OuterMeasure α} :
        m₁.trim m₂.trim ∀ (s : Set α), MeasurableSet sm₁ s m₂ s
        theorem MeasureTheory.OuterMeasure.trim_eq_trim_iff {α : Type u_1} [MeasurableSpace α] {m₁ : MeasureTheory.OuterMeasure α} {m₂ : MeasureTheory.OuterMeasure α} :
        m₁.trim = m₂.trim ∀ (s : Set α), MeasurableSet sm₁ s = m₂ s
        theorem MeasureTheory.OuterMeasure.trim_eq_iInf {α : Type u_1} [MeasurableSpace α] (m : MeasureTheory.OuterMeasure α) (s : Set α) :
        m.trim s = ⨅ (t : Set α), ⨅ (_ : s t), ⨅ (_ : MeasurableSet t), m t
        theorem MeasureTheory.OuterMeasure.trim_eq_iInf' {α : Type u_1} [MeasurableSpace α] (m : MeasureTheory.OuterMeasure α) (s : Set α) :
        m.trim s = ⨅ (t : { t : Set α // s t MeasurableSet t }), m t
        theorem MeasureTheory.OuterMeasure.exists_measurable_superset_forall_eq_trim {α : Type u_1} [MeasurableSpace α] {ι : Sort u_2} [Countable ι] (μ : ιMeasureTheory.OuterMeasure α) (s : Set α) :
        ∃ (t : Set α), s t MeasurableSet t ∀ (i : ι), (μ i) t = (μ i).trim s

        If μ i is a countable family of outer measures, then for every set s there exists a measurable set t ⊇ s such that μ i t = (μ i).trim s for all i.

        theorem MeasureTheory.OuterMeasure.trim_binop {α : Type u_1} [MeasurableSpace α] {m₁ : MeasureTheory.OuterMeasure α} {m₂ : MeasureTheory.OuterMeasure α} {m₃ : MeasureTheory.OuterMeasure α} {op : ENNRealENNRealENNReal} (h : ∀ (s : Set α), m₁ s = op (m₂ s) (m₃ s)) (s : Set α) :
        m₁.trim s = op (m₂.trim s) (m₃.trim s)

        If m₁ s = op (m₂ s) (m₃ s) for all s, then the same is true for m₁.trim, m₂.trim, and m₃ s.

        theorem MeasureTheory.OuterMeasure.trim_op {α : Type u_1} [MeasurableSpace α] {m₁ : MeasureTheory.OuterMeasure α} {m₂ : MeasureTheory.OuterMeasure α} {op : ENNRealENNReal} (h : ∀ (s : Set α), m₁ s = op (m₂ s)) (s : Set α) :
        m₁.trim s = op (m₂.trim s)

        If m₁ s = op (m₂ s) for all s, then the same is true for m₁.trim and m₂.trim.

        theorem MeasureTheory.OuterMeasure.trim_add {α : Type u_1} [MeasurableSpace α] (m₁ : MeasureTheory.OuterMeasure α) (m₂ : MeasureTheory.OuterMeasure α) :
        (m₁ + m₂).trim = m₁.trim + m₂.trim

        trim is additive.

        theorem MeasureTheory.OuterMeasure.trim_smul {α : Type u_1} [MeasurableSpace α] {R : Type u_2} [SMul R ENNReal] [IsScalarTower R ENNReal ENNReal] (c : R) (m : MeasureTheory.OuterMeasure α) :
        (c m).trim = c m.trim

        trim respects scalar multiplication.

        theorem MeasureTheory.OuterMeasure.trim_sup {α : Type u_1} [MeasurableSpace α] (m₁ : MeasureTheory.OuterMeasure α) (m₂ : MeasureTheory.OuterMeasure α) :
        (m₁ m₂).trim = m₁.trim m₂.trim

        trim sends the supremum of two outer measures to the supremum of the trimmed measures.

        theorem MeasureTheory.OuterMeasure.trim_iSup {α : Type u_1} [MeasurableSpace α] {ι : Sort u_2} [Countable ι] (μ : ιMeasureTheory.OuterMeasure α) :
        (⨆ (i : ι), μ i).trim = ⨆ (i : ι), (μ i).trim

        trim sends the supremum of a countable family of outer measures to the supremum of the trimmed measures.

        The trimmed property of a measure μ states that μ.toOuterMeasure.trim = μ.toOuterMeasure. This theorem shows that a restricted trimmed outer measure is a trimmed outer measure.