Skip to content

Commit 60d731f

Browse files
committed
feat: define IsKolmogorovProcess (#27202)
A stochastic process `X : T → Ω → E` on an index space `T` and a measurable space `Ω` with measure `P` is said to satisfy the Kolmogorov condition with exponents `p, q` and constant `M` if for all `s, t : T`, the pair `(X s, X t)` is measurable for the Borel sigma-algebra on `E × E` and the following condition holds: `∫⁻ ω, edist (X s ω) (X t ω) ^ p ∂P ≤ M * edist s t ^ q`. This is the main assumption of the Kolmogorov-Chentsov theorem, which gives the existence of a continuous modification of the process. From the Brownian motion project.
1 parent 220be37 commit 60d731f

File tree

2 files changed

+184
-0
lines changed

2 files changed

+184
-0
lines changed

Mathlib.lean

Lines changed: 1 addition & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -5100,6 +5100,7 @@ import Mathlib.Probability.ProbabilityMassFunction.Monad
51005100
import Mathlib.Probability.Process.Adapted
51015101
import Mathlib.Probability.Process.Filtration
51025102
import Mathlib.Probability.Process.HittingTime
5103+
import Mathlib.Probability.Process.Kolmogorov
51035104
import Mathlib.Probability.Process.PartitionFiltration
51045105
import Mathlib.Probability.Process.Stopping
51055106
import Mathlib.Probability.ProductMeasure
Lines changed: 183 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,183 @@
1+
/-
2+
Copyright (c) 2025 Rémy Degenne. All rights reserved.
3+
Released under Apache 2.0 license as described in the file LICENSE.
4+
Authors: Rémy Degenne
5+
-/
6+
import Mathlib.MeasureTheory.Function.SpecialFunctions.Basic
7+
import Mathlib.MeasureTheory.Function.StronglyMeasurable.AEStronglyMeasurable
8+
import Mathlib.MeasureTheory.Integral.Lebesgue.Basic
9+
10+
/-!
11+
# Stochastic processes satisfying the Kolmogorov condition
12+
13+
A stochastic process `X : T → Ω → E` on an index space `T` and a measurable space `Ω`
14+
with measure `P` is said to satisfy the Kolmogorov condition with exponents `p, q` and constant `M`
15+
if for all `s, t : T`, the pair `(X s, X t)` is measurable for the Borel sigma-algebra on `E × E`
16+
and the following condition holds:
17+
`∫⁻ ω, edist (X s ω) (X t ω) ^ p ∂P ≤ M * edist s t ^ q`.
18+
19+
This condition is the main assumption of the Kolmogorov-Chentsov theorem, which gives the existence
20+
of a continuous modification of the process.
21+
22+
The measurability condition on pairs ensures that the distance `edist (X s ω) (X t ω)` is
23+
measurable in `ω` for fixed `s, t`. In a space with second-countable topology, the measurability
24+
of pairs can be obtained from measurability of each `X t`.
25+
26+
## Main definitions
27+
28+
* `IsKolmogorovProcess`: property of being a stochastic process that satisfies
29+
the Kolmogorov condition.
30+
* `IsAEKolmogorovProcess`: a stochastic process satisfies `IsAEKolmogorovProcess` if it is
31+
a modification of a process satisfying the Kolmogorov condition.
32+
33+
## Main statements
34+
35+
* `IsKolmogorovProcess.mk_of_secondCountableTopology`: in a space with second-countable topology,
36+
a process is a Kolmogorov process if each `X t` is measurable and the Kolmogorov condition holds.
37+
38+
-/
39+
40+
open MeasureTheory
41+
open scoped ENNReal NNReal
42+
43+
namespace ProbabilityTheory
44+
45+
variable {T Ω E : Type*} [PseudoEMetricSpace T] {mΩ : MeasurableSpace Ω} [PseudoEMetricSpace E]
46+
{p q : ℝ} {M : ℝ≥0} {P : Measure Ω} {X : T → Ω → E}
47+
48+
/-- A stochastic process `X : T → Ω → E` on an index space `T` and a measurable space `Ω`
49+
with measure `P` is said to satisfy the Kolmogorov condition with exponents `p, q` and constant `M`
50+
if for all `s, t : T`, the pair `(X s, X t)` is measurable for the Borel sigma-algebra on `E × E`
51+
and the following condition holds: `∫⁻ ω, edist (X s ω) (X t ω) ^ p ∂P ≤ M * edist s t ^ q`. -/
52+
structure IsKolmogorovProcess (X : T → Ω → E) (P : Measure Ω) (p q : ℝ) (M : ℝ≥0) : Prop where
53+
measurablePair : ∀ s t : T, Measurable[_, borel (E × E)] fun ω ↦ (X s ω, X t ω)
54+
kolmogorovCondition : ∀ s t : T, ∫⁻ ω, edist (X s ω) (X t ω) ^ p ∂P ≤ M * edist s t ^ q
55+
p_pos : 0 < p
56+
q_pos : 0 < q
57+
58+
/-- Property of being a modification of a stochastic process that satisfies the Kolmogorov
59+
condition (`IsKolmogorovProcess`). -/
60+
def IsAEKolmogorovProcess (X : T → Ω → E) (P : Measure Ω) (p q : ℝ) (M : ℝ≥0) : Prop :=
61+
∃ Y, IsKolmogorovProcess Y P p q M ∧ ∀ t, X t =ᵐ[P] Y t
62+
63+
lemma IsKolmogorovProcess.IsAEKolmogorovProcess (hX : IsKolmogorovProcess X P p q M) :
64+
IsAEKolmogorovProcess X P p q M := ⟨X, hX, by simp⟩
65+
66+
namespace IsAEKolmogorovProcess
67+
68+
/-- A process with the property `IsKolmogorovProcess` such that `∀ t, X t =ᵐ[P] h.mk X t`. -/
69+
protected noncomputable
70+
def mk (X : T → Ω → E) (h : IsAEKolmogorovProcess X P p q M) : T → Ω → E :=
71+
Classical.choose h
72+
73+
lemma IsKolmogorovProcess_mk (h : IsAEKolmogorovProcess X P p q M) :
74+
IsKolmogorovProcess (h.mk X) P p q M := (Classical.choose_spec h).1
75+
76+
lemma ae_eq_mk (h : IsAEKolmogorovProcess X P p q M) : ∀ t, X t =ᵐ[P] h.mk X t :=
77+
(Classical.choose_spec h).2
78+
79+
lemma kolmogorovCondition (hX : IsAEKolmogorovProcess X P p q M) (s t : T) :
80+
∫⁻ ω, edist (X s ω) (X t ω) ^ p ∂P ≤ M * edist s t ^ q := by
81+
convert hX.IsKolmogorovProcess_mk.kolmogorovCondition s t using 1
82+
refine lintegral_congr_ae ?_
83+
filter_upwards [hX.ae_eq_mk s, hX.ae_eq_mk t] with ω hω₁ hω₂
84+
simp_rw [hω₁, hω₂]
85+
86+
lemma p_pos (hX : IsAEKolmogorovProcess X P p q M) : 0 < p := hX.IsKolmogorovProcess_mk.p_pos
87+
88+
lemma q_pos (hX : IsAEKolmogorovProcess X P p q M) : 0 < q := hX.IsKolmogorovProcess_mk.q_pos
89+
90+
lemma congr {Y : T → Ω → E} (hX : IsAEKolmogorovProcess X P p q M)
91+
(h : ∀ t, X t =ᵐ[P] Y t) :
92+
IsAEKolmogorovProcess Y P p q M := by
93+
refine ⟨hX.mk X, hX.IsKolmogorovProcess_mk, fun t ↦ ?_⟩
94+
filter_upwards [hX.ae_eq_mk t, h t] with ω hX hY using hY.symm.trans hX
95+
96+
end IsAEKolmogorovProcess
97+
98+
section Measurability
99+
100+
lemma IsKolmogorovProcess.stronglyMeasurable_edist
101+
(hX : IsKolmogorovProcess X P p q M) {s t : T} :
102+
StronglyMeasurable (fun ω ↦ edist (X s ω) (X t ω)) := by
103+
borelize (E × E)
104+
exact continuous_edist.stronglyMeasurable.comp_measurable (hX.measurablePair s t)
105+
106+
lemma IsAEKolmogorovProcess.aestronglyMeasurable_edist
107+
(hX : IsAEKolmogorovProcess X P p q M) {s t : T} :
108+
AEStronglyMeasurable (fun ω ↦ edist (X s ω) (X t ω)) P := by
109+
refine ⟨(fun ω ↦ edist (hX.mk X s ω) (hX.mk X t ω)),
110+
hX.IsKolmogorovProcess_mk.stronglyMeasurable_edist, ?_⟩
111+
filter_upwards [hX.ae_eq_mk s, hX.ae_eq_mk t] with ω hω₁ hω₂ using by simp [hω₁, hω₂]
112+
113+
lemma IsKolmogorovProcess.measurable_edist (hX : IsKolmogorovProcess X P p q M) {s t : T} :
114+
Measurable (fun ω ↦ edist (X s ω) (X t ω)) := hX.stronglyMeasurable_edist.measurable
115+
116+
lemma IsAEKolmogorovProcess.aemeasurable_edist (hX : IsAEKolmogorovProcess X P p q M) {s t : T} :
117+
AEMeasurable (fun ω ↦ edist (X s ω) (X t ω)) P := hX.aestronglyMeasurable_edist.aemeasurable
118+
119+
variable [MeasurableSpace E] [BorelSpace E]
120+
121+
lemma IsKolmogorovProcess.measurable (hX : IsKolmogorovProcess X P p q M) (s : T) :
122+
Measurable (X s) :=
123+
(measurable_fst.mono prod_le_borel_prod le_rfl).comp (hX.measurablePair s s)
124+
125+
lemma IsAEKolmogorovProcess.aemeasurable (hX : IsAEKolmogorovProcess X P p q M) (s : T) :
126+
AEMeasurable (X s) P := by
127+
refine ⟨hX.mk X s, hX.IsKolmogorovProcess_mk.measurable s, ?_⟩
128+
filter_upwards [hX.ae_eq_mk s] with ω hω using hω
129+
130+
lemma IsKolmogorovProcess.mk_of_secondCountableTopology [SecondCountableTopology E]
131+
(h_meas : ∀ s, Measurable (X s))
132+
(h_kol : ∀ s t : T, ∫⁻ ω, (edist (X s ω) (X t ω)) ^ p ∂P ≤ M * edist s t ^ q)
133+
(hp : 0 < p) (hq : 0 < q) :
134+
IsKolmogorovProcess X P p q M where
135+
measurablePair s t := by
136+
suffices Measurable (fun ω ↦ (X s ω, X t ω)) by
137+
rwa [Prod.borelSpace.measurable_eq] at this
138+
fun_prop
139+
kolmogorovCondition := h_kol
140+
p_pos := hp
141+
q_pos := hq
142+
143+
end Measurability
144+
145+
section ZeroDist
146+
147+
lemma IsAEKolmogorovProcess.edist_eq_zero (hX : IsAEKolmogorovProcess X P p q M)
148+
{s t : T} (h : edist s t = 0) :
149+
∀ᵐ ω ∂P, edist (X s ω) (X t ω) = 0 := by
150+
suffices ∀ᵐ ω ∂P, edist (X s ω) (X t ω) ^ p = 0 by
151+
filter_upwards [this] with ω hω
152+
simpa [hX.p_pos, not_lt_of_gt hX.p_pos] using hω
153+
refine (lintegral_eq_zero_iff' (hX.aemeasurable_edist.pow_const p)).mp ?_
154+
refine le_antisymm ?_ zero_le'
155+
calc ∫⁻ ω, edist (X s ω) (X t ω) ^ p ∂P
156+
_ ≤ M * edist s t ^ q := hX.kolmogorovCondition s t
157+
_ = 0 := by simp [h, hX.q_pos]
158+
159+
lemma IsKolmogorovProcess.edist_eq_zero (hX : IsKolmogorovProcess X P p q M)
160+
{s t : T} (h : edist s t = 0) :
161+
∀ᵐ ω ∂P, edist (X s ω) (X t ω) = 0 :=
162+
hX.IsAEKolmogorovProcess.edist_eq_zero h
163+
164+
lemma IsAEKolmogorovProcess.edist_eq_zero_of_const_eq_zero (hX : IsAEKolmogorovProcess X P p q 0)
165+
(s t : T) :
166+
∀ᵐ ω ∂P, edist (X s ω) (X t ω) = 0 := by
167+
suffices ∀ᵐ ω ∂P, edist (X s ω) (X t ω) ^ p = 0 by
168+
filter_upwards [this] with ω hω
169+
simpa [hX.p_pos, not_lt_of_gt hX.p_pos] using hω
170+
refine (lintegral_eq_zero_iff' (hX.aemeasurable_edist.pow_const p)).mp ?_
171+
refine le_antisymm ?_ zero_le'
172+
calc ∫⁻ ω, edist (X s ω) (X t ω) ^ p ∂P
173+
_ ≤ 0 * edist s t ^ q := hX.kolmogorovCondition s t
174+
_ = 0 := by simp
175+
176+
lemma IsKolmogorovProcess.edist_eq_zero_of_const_eq_zero (hX : IsKolmogorovProcess X P p q 0)
177+
(s t : T) :
178+
∀ᵐ ω ∂P, edist (X s ω) (X t ω) = 0 :=
179+
hX.IsAEKolmogorovProcess.edist_eq_zero_of_const_eq_zero s t
180+
181+
end ZeroDist
182+
183+
end ProbabilityTheory

0 commit comments

Comments
 (0)