Documentation

Mathlib.MeasureTheory.MeasurableSpace.Basic

Measurable spaces and measurable functions #

This file provides properties of measurable spaces and the functions and isomorphisms between them. The definition of a measurable space is in Mathlib/MeasureTheory/MeasurableSpace/Defs.lean.

A measurable space is a set equipped with a σ-algebra, a collection of subsets closed under complementation and countable union. A function between measurable spaces is measurable if the preimage of each measurable subset is measurable.

σ-algebras on a fixed set α form a complete lattice. Here we order σ-algebras by writing m₁ ≤ m₂ if every set which is m₁-measurable is also m₂-measurable (that is, m₁ is a subset of m₂). In particular, any collection of subsets of α generates a smallest σ-algebra which contains all of them. A function f : α → β induces a Galois connection between the lattices of σ-algebras on α and β.

Implementation notes #

Measurability of a function f : α → β between measurable spaces is defined in terms of the Galois connection induced by f.

References #

Tags #

measurable space, σ-algebra, measurable function, dynkin system, π-λ theorem, π-system

def MeasurableSpace.map {α : Type u_1} {β : Type u_2} (f : αβ) (m : MeasurableSpace α) :

The forward image of a measurable space under a function. map f m contains the sets s : Set β whose preimage under f is measurable.

Equations
Instances For
    theorem MeasurableSpace.map_def {α : Type u_1} {β : Type u_2} {m : MeasurableSpace α} {f : αβ} {s : Set β} :
    @[simp]
    @[simp]
    theorem MeasurableSpace.map_comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} {m : MeasurableSpace α} {f : αβ} {g : βγ} :
    def MeasurableSpace.comap {α : Type u_1} {β : Type u_2} (f : αβ) (m : MeasurableSpace β) :

    The reverse image of a measurable space under a function. comap f m contains the sets s : Set α such that s is the f-preimage of a measurable set in β.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      theorem MeasurableSpace.measurableSet_comap {α : Type u_1} {β : Type u_2} {s : Set α} {f : αβ} {m : MeasurableSpace β} :
      MeasurableSet s ∃ (s' : Set β), MeasurableSet s' f ⁻¹' s' = s
      theorem MeasurableSpace.comap_eq_generateFrom {α : Type u_1} {β : Type u_2} (m : MeasurableSpace β) (f : αβ) :
      @[simp]
      theorem MeasurableSpace.comap_comp {α : Type u_1} {β : Type u_2} {γ : Type u_3} {m : MeasurableSpace α} {f : βα} {g : γβ} :
      theorem MeasurableSpace.comap_le_iff_le_map {α : Type u_1} {β : Type u_2} {m : MeasurableSpace α} {m' : MeasurableSpace β} {f : αβ} :
      theorem MeasurableSpace.map_mono {α : Type u_1} {β : Type u_2} {m₁ m₂ : MeasurableSpace α} {f : αβ} (h : m₁ m₂) :
      theorem MeasurableSpace.monotone_map {α : Type u_1} {β : Type u_2} {f : αβ} :
      theorem MeasurableSpace.comap_mono {α : Type u_1} {β : Type u_2} {m₁ m₂ : MeasurableSpace α} {g : βα} (h : m₁ m₂) :
      theorem MeasurableSpace.monotone_comap {α : Type u_1} {β : Type u_2} {g : βα} :
      @[simp]
      theorem MeasurableSpace.comap_bot {α : Type u_1} {β : Type u_2} {g : βα} :
      @[simp]
      theorem MeasurableSpace.comap_sup {α : Type u_1} {β : Type u_2} {m₁ m₂ : MeasurableSpace α} {g : βα} :
      @[simp]
      theorem MeasurableSpace.comap_iSup {α : Type u_1} {β : Type u_2} {ι : Sort uι} {g : βα} {m : ιMeasurableSpace α} :
      MeasurableSpace.comap g (⨆ (i : ι), m i) = ⨆ (i : ι), MeasurableSpace.comap g (m i)
      @[simp]
      theorem MeasurableSpace.map_top {α : Type u_1} {β : Type u_2} {f : αβ} :
      @[simp]
      theorem MeasurableSpace.map_inf {α : Type u_1} {β : Type u_2} {m₁ m₂ : MeasurableSpace α} {f : αβ} :
      @[simp]
      theorem MeasurableSpace.map_iInf {α : Type u_1} {β : Type u_2} {ι : Sort uι} {f : αβ} {m : ιMeasurableSpace α} :
      MeasurableSpace.map f (⨅ (i : ι), m i) = ⨅ (i : ι), MeasurableSpace.map f (m i)
      theorem MeasurableSpace.comap_map_le {α : Type u_1} {β : Type u_2} {m : MeasurableSpace α} {f : αβ} :
      theorem MeasurableSpace.le_map_comap {α : Type u_1} {β : Type u_2} {m : MeasurableSpace α} {g : βα} :
      @[simp]
      theorem MeasurableSpace.map_const {α : Type u_1} {β : Type u_2} {m : MeasurableSpace α} (b : β) :
      MeasurableSpace.map (fun (_a : α) => b) m =
      @[simp]
      theorem MeasurableSpace.comap_const {α : Type u_1} {β : Type u_2} {m : MeasurableSpace β} (b : β) :
      MeasurableSpace.comap (fun (_a : α) => b) m =
      theorem MeasurableSpace.comap_generateFrom {α : Type u_1} {β : Type u_2} {f : αβ} {s : Set (Set β)} :
      theorem measurable_iff_le_map {α : Type u_1} {β : Type u_2} {m₁ : MeasurableSpace α} {m₂ : MeasurableSpace β} {f : αβ} :
      theorem Measurable.of_le_map {α : Type u_1} {β : Type u_2} {m₁ : MeasurableSpace α} {m₂ : MeasurableSpace β} {f : αβ} :

      Alias of the reverse direction of measurable_iff_le_map.

      theorem Measurable.le_map {α : Type u_1} {β : Type u_2} {m₁ : MeasurableSpace α} {m₂ : MeasurableSpace β} {f : αβ} :

      Alias of the forward direction of measurable_iff_le_map.

      theorem measurable_iff_comap_le {α : Type u_1} {β : Type u_2} {m₁ : MeasurableSpace α} {m₂ : MeasurableSpace β} {f : αβ} :
      theorem Measurable.comap_le {α : Type u_1} {β : Type u_2} {m₁ : MeasurableSpace α} {m₂ : MeasurableSpace β} {f : αβ} :

      Alias of the forward direction of measurable_iff_comap_le.

      theorem Measurable.of_comap_le {α : Type u_1} {β : Type u_2} {m₁ : MeasurableSpace α} {m₂ : MeasurableSpace β} {f : αβ} :

      Alias of the reverse direction of measurable_iff_comap_le.

      theorem comap_measurable {α : Type u_1} {β : Type u_2} {m : MeasurableSpace β} (f : αβ) :
      theorem Measurable.mono {α : Type u_1} {β : Type u_2} {ma ma' : MeasurableSpace α} {mb mb' : MeasurableSpace β} {f : αβ} (hf : Measurable f) (ha : ma ma') (hb : mb' mb) :
      theorem Measurable.iSup' {α : Type u_1} {β : Type u_2} {ι : Sort uι} { : ιMeasurableSpace α} {x✝ : MeasurableSpace β} {f : αβ} (i₀ : ι) (h : Measurable f) :
      theorem Measurable.sup_of_left {α : Type u_1} {β : Type u_2} {mα' : MeasurableSpace α} {x✝ : MeasurableSpace β} {f : αβ} (h : Measurable f) :
      theorem Measurable.sup_of_right {α : Type u_1} {β : Type u_2} {mα' : MeasurableSpace α} {x✝ : MeasurableSpace β} {f : αβ} (h : Measurable f) :
      theorem measurable_id'' {α : Type u_1} {m : MeasurableSpace α} (hm : m ) :
      theorem measurable_from_top {α : Type u_1} {β : Type u_2} [MeasurableSpace β] {f : αβ} :
      theorem measurable_generateFrom {α : Type u_1} {β : Type u_2} [MeasurableSpace α] {s : Set (Set β)} {f : αβ} (h : ts, MeasurableSet (f ⁻¹' t)) :
      theorem Subsingleton.measurable {α : Type u_1} {β : Type u_2} {f : αβ} [MeasurableSpace α] [MeasurableSpace β] [Subsingleton α] :
      theorem measurable_of_subsingleton_codomain {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] [Subsingleton β] (f : αβ) :
      theorem measurable_one {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] [One α] :
      theorem measurable_zero {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] [Zero α] :
      theorem measurable_of_empty {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] [IsEmpty α] (f : αβ) :
      theorem measurable_of_empty_codomain {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] [IsEmpty β] (f : αβ) :
      theorem measurable_const' {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] {f : βα} (hf : ∀ (x y : β), f x = f y) :

      A version of measurable_const that assumes f x = f y for all x, y. This version works for functions between empty types.

      theorem measurable_natCast {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] [NatCast α] (n : ) :
      theorem measurable_intCast {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] [IntCast α] (n : ) :
      theorem measurable_of_countable {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] [Countable α] [MeasurableSingletonClass α] (f : αβ) :
      theorem measurable_of_finite {α : Type u_1} {β : Type u_2} [MeasurableSpace α] [MeasurableSpace β] [Finite α] [MeasurableSingletonClass α] (f : αβ) :
      theorem Measurable.iterate {α : Type u_1} {m : MeasurableSpace α} {f : αα} (hf : Measurable f) (n : ) :
      theorem measurableSet_preimage {α : Type u_1} {β : Type u_2} {f : αβ} {m : MeasurableSpace α} { : MeasurableSpace β} {t : Set β} (hf : Measurable f) (ht : MeasurableSet t) :
      theorem MeasurableSet.preimage {α : Type u_1} {β : Type u_2} {f : αβ} {m : MeasurableSpace α} { : MeasurableSpace β} {t : Set β} (ht : MeasurableSet t) (hf : Measurable f) :
      theorem Measurable.piecewise {α : Type u_1} {β : Type u_2} {s : Set α} {f g : αβ} {m : MeasurableSpace α} { : MeasurableSpace β} {x✝ : DecidablePred fun (x : α) => x s} (hs : MeasurableSet s) (hf : Measurable f) (hg : Measurable g) :
      theorem Measurable.ite {α : Type u_1} {β : Type u_2} {f g : αβ} {m : MeasurableSpace α} { : MeasurableSpace β} {p : αProp} {x✝ : DecidablePred p} (hp : MeasurableSet {a : α | p a}) (hf : Measurable f) (hg : Measurable g) :
      Measurable fun (x : α) => if p x then f x else g x

      This is slightly different from Measurable.piecewise. It can be used to show Measurable (ite (x=0) 0 1) by exact Measurable.ite (measurableSet_singleton 0) measurable_const measurable_const, but replacing Measurable.ite by Measurable.piecewise in that example proof does not work.

      theorem Measurable.indicator {α : Type u_1} {β : Type u_2} {s : Set α} {f : αβ} {m : MeasurableSpace α} { : MeasurableSpace β} [Zero β] (hf : Measurable f) (hs : MeasurableSet s) :
      theorem measurable_indicator_const_iff {α : Type u_1} {β : Type u_2} {s : Set α} {m : MeasurableSpace α} { : MeasurableSpace β} [Zero β] [MeasurableSingletonClass β] (b : β) [NeZero b] :
      Measurable (s.indicator fun (x : α) => b) MeasurableSet s

      The measurability of a set A is equivalent to the measurability of the indicator function which takes a constant value b ≠ 0 on a set A and 0 elsewhere.

      theorem measurableSet_mulSupport {α : Type u_1} {β : Type u_2} {f : αβ} {m : MeasurableSpace α} { : MeasurableSpace β} [One β] [MeasurableSingletonClass β] (hf : Measurable f) :
      theorem measurableSet_support {α : Type u_1} {β : Type u_2} {f : αβ} {m : MeasurableSpace α} { : MeasurableSpace β} [Zero β] [MeasurableSingletonClass β] (hf : Measurable f) :
      theorem Measurable.measurable_of_countable_ne {α : Type u_1} {β : Type u_2} {f g : αβ} {m : MeasurableSpace α} { : MeasurableSpace β} [MeasurableSingletonClass α] (hf : Measurable f) (h : {x : α | f x g x}.Countable) :

      If a function coincides with a measurable function outside of a countable set, it is measurable.

      def IsCountablySpanning {α : Type u_1} (C : Set (Set α)) :

      We say that a collection of sets is countably spanning if a countable subset spans the whole type. This is a useful condition in various parts of measure theory. For example, it is a needed condition to show that the product of two collections generate the product sigma algebra, see generateFrom_prod_eq.

      Equations
      Instances For
        theorem IsCountablySpanning.prod {α : Type u_1} {β : Type u_2} {C : Set (Set α)} {D : Set (Set β)} (hC : IsCountablySpanning C) (hD : IsCountablySpanning D) :
        IsCountablySpanning (Set.image2 (fun (x1 : Set α) (x2 : Set β) => x1 ×ˢ x2) C D)

        Rectangles of countably spanning sets are countably spanning.