Documentation

Mathlib.MeasureTheory.PiSystem

Induction principles for measurable sets, related to π-systems and λ-systems. #

Main statements #

Implementation details #

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

A π-system is a collection of subsets of α that is closed under binary intersection of non-disjoint sets. Usually it is also required that the collection is nonempty, but we don't do that here.

Equations
theorem IsPiSystem.singleton {α : Type u_1} (S : Set α) :
theorem IsPiSystem.insert_empty {α : Type u_1} {S : Set (Set α)} (h_pi : IsPiSystem S) :
theorem IsPiSystem.insert_univ {α : Type u_1} {S : Set (Set α)} (h_pi : IsPiSystem S) :
theorem IsPiSystem.comap {α : Type u_3} {β : Type u_4} {S : Set (Set β)} (h_pi : IsPiSystem S) (f : αβ) :
IsPiSystem {s : Set α | tS, f ⁻¹' t = s}
theorem isPiSystem_iUnion_of_directed_le {α : Type u_3} {ι : Sort u_4} (p : ιSet (Set α)) (hp_pi : ∀ (n : ι), IsPiSystem (p n)) (hp_directed : Directed (fun (x1 x2 : Set (Set α)) => x1 x2) p) :
IsPiSystem (⋃ (n : ι), p n)
theorem isPiSystem_iUnion_of_monotone {α : Type u_3} {ι : Type u_4} [SemilatticeSup ι] (p : ιSet (Set α)) (hp_pi : ∀ (n : ι), IsPiSystem (p n)) (hp_mono : Monotone p) :
IsPiSystem (⋃ (n : ι), p n)
theorem IsPiSystem.prod {α : Type u_1} {β : Type u_2} {C : Set (Set α)} {D : Set (Set β)} (hC : IsPiSystem C) (hD : IsPiSystem D) :
IsPiSystem (Set.image2 (fun (x1 : Set α) (x2 : Set β) => x1 ×ˢ x2) C D)

Rectangles formed by π-systems form a π-system.

theorem isPiSystem_image_Iio {α : Type u_1} [LinearOrder α] (s : Set α) :
theorem isPiSystem_image_Ioi {α : Type u_1} [LinearOrder α] (s : Set α) :
theorem isPiSystem_image_Iic {α : Type u_1} [LinearOrder α] (s : Set α) :
theorem isPiSystem_image_Ici {α : Type u_1} [LinearOrder α] (s : Set α) :
theorem isPiSystem_Ixx_mem {α : Type u_1} [LinearOrder α] {Ixx : ααSet α} {p : ααProp} (Hne : ∀ {a b : α}, (Ixx a b).Nonemptyp a b) (Hi : ∀ {a₁ b₁ a₂ b₂ : α}, Ixx a₁ b₁ Ixx a₂ b₂ = Ixx (max a₁ a₂) (min b₁ b₂)) (s t : Set α) :
IsPiSystem {S : Set α | ls, ut, p l u Ixx l u = S}
theorem isPiSystem_Ixx {α : Type u_1} {ι : Sort u_3} {ι' : Sort u_4} [LinearOrder α] {Ixx : ααSet α} {p : ααProp} (Hne : ∀ {a b : α}, (Ixx a b).Nonemptyp a b) (Hi : ∀ {a₁ b₁ a₂ b₂ : α}, Ixx a₁ b₁ Ixx a₂ b₂ = Ixx (max a₁ a₂) (min b₁ b₂)) (f : ια) (g : ι'α) :
IsPiSystem {S : Set α | ∃ (i : ι) (j : ι'), p (f i) (g j) Ixx (f i) (g j) = S}
theorem isPiSystem_Ioo_mem {α : Type u_1} [LinearOrder α] (s t : Set α) :
IsPiSystem {S : Set α | ls, ut, l < u Set.Ioo l u = S}
theorem isPiSystem_Ioo {α : Type u_1} {ι : Sort u_3} {ι' : Sort u_4} [LinearOrder α] (f : ια) (g : ι'α) :
IsPiSystem {S : Set α | ∃ (l : ι) (u : ι'), f l < g u Set.Ioo (f l) (g u) = S}
theorem isPiSystem_Ioc_mem {α : Type u_1} [LinearOrder α] (s t : Set α) :
IsPiSystem {S : Set α | ls, ut, l < u Set.Ioc l u = S}
theorem isPiSystem_Ioc {α : Type u_1} {ι : Sort u_3} {ι' : Sort u_4} [LinearOrder α] (f : ια) (g : ι'α) :
IsPiSystem {S : Set α | ∃ (i : ι) (j : ι'), f i < g j Set.Ioc (f i) (g j) = S}
theorem isPiSystem_Ico_mem {α : Type u_1} [LinearOrder α] (s t : Set α) :
IsPiSystem {S : Set α | ls, ut, l < u Set.Ico l u = S}
theorem isPiSystem_Ico {α : Type u_1} {ι : Sort u_3} {ι' : Sort u_4} [LinearOrder α] (f : ια) (g : ι'α) :
IsPiSystem {S : Set α | ∃ (i : ι) (j : ι'), f i < g j Set.Ico (f i) (g j) = S}
theorem isPiSystem_Icc_mem {α : Type u_1} [LinearOrder α] (s t : Set α) :
IsPiSystem {S : Set α | ls, ut, l u Set.Icc l u = S}
theorem isPiSystem_Icc {α : Type u_1} {ι : Sort u_3} {ι' : Sort u_4} [LinearOrder α] (f : ια) (g : ι'α) :
IsPiSystem {S : Set α | ∃ (i : ι) (j : ι'), f i g j Set.Icc (f i) (g j) = S}
inductive generatePiSystem {α : Type u_1} (S : Set (Set α)) :
Set (Set α)

Given a collection S of subsets of α, then generatePiSystem S is the smallest π-system containing S.

theorem generatePiSystem_subset_self {α : Type u_1} {S : Set (Set α)} (h_S : IsPiSystem S) :
theorem generatePiSystem_eq {α : Type u_1} {S : Set (Set α)} (h_pi : IsPiSystem S) :
theorem generatePiSystem_mono {α : Type u_1} {S T : Set (Set α)} (hST : S T) :
theorem generatePiSystem_measurableSet {α : Type u_1} [M : MeasurableSpace α] {S : Set (Set α)} (h_meas_S : sS, MeasurableSet s) (t : Set α) (h_in_pi : t generatePiSystem S) :
theorem mem_generatePiSystem_iUnion_elim {α : Type u_3} {β : Type u_4} {g : βSet (Set α)} (h_pi : ∀ (b : β), IsPiSystem (g b)) (t : Set α) (h_t : t generatePiSystem (⋃ (b : β), g b)) :
∃ (T : Finset β) (f : βSet α), t = bT, f b bT, f b g b

Every element of the π-system generated by the union of a family of π-systems is a finite intersection of elements from the π-systems. For an indexed union version, see mem_generatePiSystem_iUnion_elim'.

theorem mem_generatePiSystem_iUnion_elim' {α : Type u_3} {β : Type u_4} {g : βSet (Set α)} {s : Set β} (h_pi : bs, IsPiSystem (g b)) (t : Set α) (h_t : t generatePiSystem (⋃ bs, g b)) :
∃ (T : Finset β) (f : βSet α), T s t = bT, f b bT, f b g b

Every element of the π-system generated by an indexed union of a family of π-systems is a finite intersection of elements from the π-systems. For a total union version, see mem_generatePiSystem_iUnion_elim.

π-system generated by finite intersections of sets of a π-system family #

def piiUnionInter {α : Type u_3} {ι : Type u_4} (π : ιSet (Set α)) (S : Set ι) :
Set (Set α)

From a set of indices S : Set ι and a family of sets of sets π : ι → Set (Set α), define the set of sets that can be written as ⋂ x ∈ t, f x for some finset t ⊆ S and sets f x ∈ π x. If π is a family of π-systems, then it is a π-system.

Equations
theorem piiUnionInter_singleton {α : Type u_3} {ι : Type u_4} (π : ιSet (Set α)) (i : ι) :
theorem piiUnionInter_singleton_left {α : Type u_3} {ι : Type u_4} (s : ιSet α) (S : Set ι) :
piiUnionInter (fun (i : ι) => {s i}) S = {s' : Set α | ∃ (t : Finset ι) (_ : t S), s' = it, s i}
theorem generateFrom_piiUnionInter_singleton_left {α : Type u_3} {ι : Type u_4} (s : ιSet α) (S : Set ι) :
MeasurableSpace.generateFrom (piiUnionInter (fun (k : ι) => {s k}) S) = MeasurableSpace.generateFrom {t : Set α | kS, s k = t}
theorem isPiSystem_piiUnionInter {α : Type u_3} {ι : Type u_4} (π : ιSet (Set α)) (hpi : ∀ (x : ι), IsPiSystem (π x)) (S : Set ι) :

If π is a family of π-systems, then piiUnionInter π S is a π-system.

theorem piiUnionInter_mono_left {α : Type u_3} {ι : Type u_4} {π π' : ιSet (Set α)} (h_le : ∀ (i : ι), π i π' i) (S : Set ι) :
theorem piiUnionInter_mono_right {α : Type u_3} {ι : Type u_4} {π : ιSet (Set α)} {S T : Set ι} (hST : S T) :
theorem generateFrom_piiUnionInter_le {α : Type u_3} {ι : Type u_4} {m : MeasurableSpace α} (π : ιSet (Set α)) (h : ∀ (n : ι), MeasurableSpace.generateFrom (π n) m) (S : Set ι) :
theorem subset_piiUnionInter {α : Type u_3} {ι : Type u_4} {π : ιSet (Set α)} {S : Set ι} {i : ι} (his : i S) :
π i piiUnionInter π S
theorem mem_piiUnionInter_of_measurableSet {α : Type u_3} {ι : Type u_4} (m : ιMeasurableSpace α) {S : Set ι} {i : ι} (hiS : i S) (s : Set α) (hs : MeasurableSet s) :
s piiUnionInter (fun (n : ι) => {s : Set α | MeasurableSet s}) S
theorem le_generateFrom_piiUnionInter {α : Type u_3} {ι : Type u_4} {π : ιSet (Set α)} (S : Set ι) {x : ι} (hxS : x S) :
theorem measurableSet_iSup_of_mem_piiUnionInter {α : Type u_3} {ι : Type u_4} (m : ιMeasurableSpace α) (S : Set ι) (t : Set α) (ht : t piiUnionInter (fun (n : ι) => {s : Set α | MeasurableSet s}) S) :
theorem generateFrom_piiUnionInter_measurableSet {α : Type u_3} {ι : Type u_4} (m : ιMeasurableSpace α) (S : Set ι) :
MeasurableSpace.generateFrom (piiUnionInter (fun (n : ι) => {s : Set α | MeasurableSet s}) S) = iS, m i

Dynkin systems and Π-λ theorem #

structure MeasurableSpace.DynkinSystem (α : Type u_4) :
Type u_4

A Dynkin system is a collection of subsets of a type α that contains the empty set, is closed under complementation and under countable union of pairwise disjoint sets. The disjointness condition is the only difference with σ-algebras.

The main purpose of Dynkin systems is to provide a powerful induction rule for σ-algebras generated by a collection of sets which is stable under intersection.

A Dynkin system is also known as a "λ-system" or a "d-system".

theorem MeasurableSpace.DynkinSystem.ext {α : Type u_3} {d₁ d₂ : DynkinSystem α} :
(∀ (s : Set α), d₁.Has s d₂.Has s)d₁ = d₂
theorem MeasurableSpace.DynkinSystem.ext_iff {α : Type u_3} {d₁ d₂ : DynkinSystem α} :
d₁ = d₂ ∀ (s : Set α), d₁.Has s d₂.Has s
theorem MeasurableSpace.DynkinSystem.has_iUnion {α : Type u_3} (d : DynkinSystem α) {β : Type u_4} [Countable β] {f : βSet α} (hd : Pairwise (Function.onFun Disjoint f)) (h : ∀ (i : β), d.Has (f i)) :
d.Has (⋃ (i : β), f i)
theorem MeasurableSpace.DynkinSystem.has_union {α : Type u_3} (d : DynkinSystem α) {s₁ s₂ : Set α} (h₁ : d.Has s₁) (h₂ : d.Has s₂) (h : Disjoint s₁ s₂) :
d.Has (s₁ s₂)
theorem MeasurableSpace.DynkinSystem.has_diff {α : Type u_3} (d : DynkinSystem α) {s₁ s₂ : Set α} (h₁ : d.Has s₁) (h₂ : d.Has s₂) (h : s₂ s₁) :
d.Has (s₁ \ s₂)
Equations
  • One or more equations did not get rendered due to their size.

Every measurable space (σ-algebra) forms a Dynkin system

Equations
inductive MeasurableSpace.DynkinSystem.GenerateHas {α : Type u_3} (s : Set (Set α)) :
Set αProp

The least Dynkin system containing a collection of basic sets. This inductive type gives the underlying collection of sets.

The least Dynkin system containing a collection of basic sets.

Equations
def MeasurableSpace.DynkinSystem.toMeasurableSpace {α : Type u_3} (d : DynkinSystem α) (h_inter : ∀ (s₁ s₂ : Set α), d.Has s₁d.Has s₂d.Has (s₁ s₂)) :

If a Dynkin system is closed under binary intersection, then it forms a σ-algebra.

Equations
  • d.toMeasurableSpace h_inter = { MeasurableSet' := d.Has, measurableSet_empty := , measurableSet_compl := , measurableSet_iUnion := }
theorem MeasurableSpace.DynkinSystem.ofMeasurableSpace_toMeasurableSpace {α : Type u_3} (d : DynkinSystem α) (h_inter : ∀ (s₁ s₂ : Set α), d.Has s₁d.Has s₂d.Has (s₁ s₂)) :
def MeasurableSpace.DynkinSystem.restrictOn {α : Type u_3} (d : DynkinSystem α) {s : Set α} (h : d.Has s) :

If s is in a Dynkin system d, we can form the new Dynkin system {s ∩ t | t ∈ d}.

Equations
  • d.restrictOn h = { Has := fun (t : Set α) => d.Has (t s), has_empty := , has_compl := , has_iUnion_nat := }
theorem MeasurableSpace.DynkinSystem.generate_le {α : Type u_3} (d : DynkinSystem α) {s : Set (Set α)} (h : ts, d.Has t) :
theorem MeasurableSpace.DynkinSystem.generate_inter {α : Type u_3} {s : Set (Set α)} (hs : IsPiSystem s) {t₁ t₂ : Set α} (ht₁ : (generate s).Has t₁) (ht₂ : (generate s).Has t₂) :
(generate s).Has (t₁ t₂)

Dynkin's π-λ theorem: Given a collection of sets closed under binary intersections, then the Dynkin system it generates is equal to the σ-algebra it generates. This result is known as the π-λ theorem. A collection of sets closed under binary intersection is called a π-system (often requiring additionally that it is non-empty, but we drop this condition in the formalization).

theorem MeasurableSpace.induction_on_inter {α : Type u_3} {m : MeasurableSpace α} {C : (s : Set α) → MeasurableSet sProp} {s : Set (Set α)} (h_eq : m = generateFrom s) (h_inter : IsPiSystem s) (empty : C ) (basic : ∀ (t : Set α) (ht : t s), C t ) (compl : ∀ (t : Set α) (htm : MeasurableSet t), C t htmC t ) (iUnion : ∀ (f : Set α), Pairwise (Function.onFun Disjoint f)∀ (hfm : ∀ (i : ), MeasurableSet (f i)), (∀ (i : ), C (f i) )C (⋃ (i : ), f i) ) (t : Set α) (ht : MeasurableSet t) :
C t ht

Induction principle for measurable sets. If s is a π-system that generates the product σ-algebra on α and a predicate C defined on measurable sets is true

  • on the empty set;
  • on each set t ∈ s;
  • on the complement of a measurable set that satisfies C;
  • on the union of a sequence of pairwise disjoint measurable sets that satisfy C,

then it is true on all measurable sets in α.