Filtrations #
This file defines filtrations of a measurable space and σ-finite filtrations.
Main definitions #
MeasureTheory.Filtration: a filtration on a measurable space. That is, a monotone sequence of sub-σ-algebras.MeasureTheory.SigmaFiniteFiltration: a filtrationfis σ-finite with respect to a measureμif for alli,μ.trim (f.le i)is σ-finite.MeasureTheory.Filtration.natural: the smallest filtration that makes a process adapted. That notionadaptedis not defined yet in this file. SeeMeasureTheory.adapted.MeasureTheory.Filtration.rightCont: the right-continuation of a filtration.MeasureTheory.Filtration.IsRightContinuous: a filtration is right-continuous if it is equal to its right-continuation.
Main results #
MeasureTheory.Filtration.instCompleteLattice: filtrations are a complete lattice.
Tags #
filtration, stochastic process
A Filtration on a measurable space Ω with σ-algebra m is a monotone
sequence of sub-σ-algebras of m.
- seq : ι → MeasurableSpace Ω
The sequence of sub-σ-algebras of
m - mono' : Monotone ↑self
Instances For
Equations
- MeasureTheory.instCoeFunFiltrationForallMeasurableSpace = { coe := fun (f : MeasureTheory.Filtration ι m) => ↑f }
The constant filtration which is equal to m for all i : ι.
Equations
- MeasureTheory.Filtration.const ι m' hm' = { seq := fun (x : ι) => m', mono' := ⋯, le' := ⋯ }
Instances For
Equations
- MeasureTheory.Filtration.instInhabited = { default := MeasureTheory.Filtration.const ι m ⋯ }
Equations
- MeasureTheory.Filtration.instLE = { le := fun (f g : MeasureTheory.Filtration ι m) => ∀ (i : ι), ↑f i ≤ ↑g i }
Equations
Equations
- MeasureTheory.Filtration.instTop = { top := MeasureTheory.Filtration.const ι m ⋯ }
Equations
- MeasureTheory.Filtration.instMax = { max := fun (f g : MeasureTheory.Filtration ι m) => { seq := fun (i : ι) => ↑f i ⊔ ↑g i, mono' := ⋯, le' := ⋯ } }
Equations
- MeasureTheory.Filtration.instMin = { min := fun (f g : MeasureTheory.Filtration ι m) => { seq := fun (i : ι) => ↑f i ⊓ ↑g i, mono' := ⋯, le' := ⋯ } }
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
Equations
- One or more equations did not get rendered due to their size.
A measure is σ-finite with respect to filtration if it is σ-finite with respect to all the sub-σ-algebra of the filtration.
- SigmaFinite (i : ι) : MeasureTheory.SigmaFinite (μ.trim ⋯)
Instances
Given an integrable function g, the conditional expectations of g with respect to a
filtration is uniformly integrable.
Given a sequence of measurable sets (sₙ), filtrationOfSet is the smallest filtration
such that sₙ is measurable with respect to the n-th sub-σ-algebra in filtrationOfSet.
Equations
Instances For
Given a filtration 𝓕, its right continuation is the filtration 𝓕₊ defined as follows:
- If
iis isolated on the right, then𝓕₊ i := 𝓕 i; - Otherwise,
𝓕₊ i := ⨅ j > i, 𝓕 j. It is sometimes simply defined as𝓕₊ i := ⨅ j > i, 𝓕 jwhen the index type isℝ. In the general case this is not ideal however. Ifiis maximal for instance, then𝓕₊ i = ⊤, which is inconvenient because𝓕₊is not aFiltration ι manymore. If the index type is discrete (such asℕ), then we would have𝓕 = 𝓕₊(i.e.𝓕is right-continuous) only if𝓕is constant.
To avoid requiring a TopologicalSpace instance on ι in the definition, we endow ι with
the order topology Preorder.topology inside the definition. Say you write a statement about
𝓕₊ which does not require a TopologicalSpace structure on ι,
but you wish to use a statement which requires a topology (such as rightCont_apply).
Then you can endow ι with the order topology by writing
letI := Preorder.topology ι
haveI : OrderTopology ι := ⟨rfl⟩
Equations
Instances For
Given a filtration 𝓕, its right continuation is the filtration 𝓕₊ defined as follows:
- If
iis isolated on the right, then𝓕₊ i := 𝓕 i; - Otherwise,
𝓕₊ i := ⨅ j > i, 𝓕 j. It is sometimes simply defined as𝓕₊ i := ⨅ j > i, 𝓕 jwhen the index type isℝ. In the general case this is not ideal however. Ifiis maximal for instance, then𝓕₊ i = ⊤, which is inconvenient because𝓕₊is not aFiltration ι manymore. If the index type is discrete (such asℕ), then we would have𝓕 = 𝓕₊(i.e.𝓕is right-continuous) only if𝓕is constant.
To avoid requiring a TopologicalSpace instance on ι in the definition, we endow ι with
the order topology Preorder.topology inside the definition. Say you write a statement about
𝓕₊ which does not require a TopologicalSpace structure on ι,
but you wish to use a statement which requires a topology (such as rightCont_apply).
Then you can endow ι with the order topology by writing
letI := Preorder.topology ι
haveI : OrderTopology ι := ⟨rfl⟩
Equations
- MeasureTheory.Filtration.«term_₊» = Lean.ParserDescr.trailingNode `MeasureTheory.Filtration.«term_₊» 1024 1024 (Lean.ParserDescr.symbol "₊")
Instances For
If the index type is a SuccOrder, then 𝓕₊ = 𝓕.
If i is not isolated on the right, then 𝓕₊ i = ⨅ j > i, 𝓕 j. This is for instance the case
when ι is a densely ordered linear order with no maximal elements and equipped with the order
topology, see rightCont_eq.
If ι is a densely ordered linear order with no maximal element, then no point is isolated
on the right, so that 𝓕₊ i = ⨅ j > i, 𝓕 j holds for all i. This is in particular the
case when ι := ℝ≥0.
A filtration 𝓕 is right continuous if it is equal to its right continuation 𝓕₊.
The right continuity property.
Instances
Given a sequence of functions, the natural filtration is the smallest sequence of σ-algebras such that the sequence of functions is measurable with respect to the filtration.
Equations
- MeasureTheory.Filtration.natural u hum = { seq := fun (i : ι) => ⨆ (j : ι), ⨆ (_ : j ≤ i), MeasurableSpace.comap (u j) (mβ j), mono' := ⋯, le' := ⋯ }
Instances For
Given a process f and a filtration ℱ, if f converges to some g almost everywhere and
g is ⨆ n, ℱ n-measurable, then limitProcess f ℱ μ chooses said g, else it returns 0.
This definition is used to phrase the a.e. martingale convergence theorem
Submartingale.ae_tendsto_limitProcess where an L¹-bounded submartingale f adapted to ℱ
converges to limitProcess f ℱ μ μ-almost everywhere.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Filtration of the first events #
The canonical filtration on the product space Π i, X i, where piLE i
consists of measurable sets depending only on coordinates ≤ i.
Equations
- MeasureTheory.Filtration.piLE = { seq := fun (i : ι) => MeasurableSpace.comap (Preorder.restrictLe i) MeasurableSpace.pi, mono' := ⋯, le' := ⋯ }
Instances For
The filtration of events which only depends on finitely many coordinates
on the product space Π i, X i, piFinset s consists of measurable sets depending only on
coordinates in s, where s : Finset ι.
Equations
- MeasureTheory.Filtration.piFinset = { seq := fun (s : Finset ι) => MeasurableSpace.comap s.restrict MeasurableSpace.pi, mono' := ⋯, le' := ⋯ }
Instances For
The exterior σ-algebras of finite sets of α form a cofiltration indexed by Finset α.
Equations
- MeasureTheory.Filtration.cylinderEventsCompl = { seq := fun (Λ : (Finset α)ᵒᵈ) => MeasureTheory.cylinderEvents (↑(OrderDual.ofDual Λ))ᶜ, mono' := ⋯, le' := ⋯ }