Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit e18fa7b

Browse files
committed
feat(probability/martingale/centering): uniqueness of Doob's decomposition (#16532)
1 parent f4210e9 commit e18fa7b

File tree

3 files changed

+79
-0
lines changed

3 files changed

+79
-0
lines changed

src/probability/martingale/basic.lean

Lines changed: 35 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -490,6 +490,41 @@ begin
490490
simpa only [pi.zero_apply, pi.neg_apply, zero_eq_neg],
491491
end
492492

493+
-- Note that one cannot use `submartingale.zero_le_of_predictable` to prove the other two
494+
-- corresponding lemmas without imposing more restrictions to the ordering of `E`
495+
/-- A predictable submartingale is a.e. greater equal than its initial state. -/
496+
lemma submartingale.zero_le_of_predictable [preorder E] [sigma_finite_filtration μ 𝒢]
497+
{f : ℕ → Ω → E} (hfmgle : submartingale f 𝒢 μ) (hfadp : adapted 𝒢 (λ n, f (n + 1))) (n : ℕ) :
498+
f 0 ≤ᵐ[μ] f n :=
499+
begin
500+
induction n with k ih,
501+
{ refl },
502+
{ exact ih.trans ((hfmgle.2.1 k (k + 1) k.le_succ).trans_eq $ germ.coe_eq.mp $ congr_arg coe $
503+
condexp_of_strongly_measurable (𝒢.le _) (hfadp _) $ hfmgle.integrable _) }
504+
end
505+
506+
/-- A predictable supermartingale is a.e. less equal than its initial state. -/
507+
lemma supermartingale.le_zero_of_predictable [preorder E] [sigma_finite_filtration μ 𝒢]
508+
{f : ℕ → Ω → E} (hfmgle : supermartingale f 𝒢 μ) (hfadp : adapted 𝒢 (λ n, f (n + 1))) (n : ℕ) :
509+
f n ≤ᵐ[μ] f 0 :=
510+
begin
511+
induction n with k ih,
512+
{ refl },
513+
{ exact ((germ.coe_eq.mp $ congr_arg coe $ condexp_of_strongly_measurable (𝒢.le _) (hfadp _) $
514+
hfmgle.integrable _).symm.trans_le (hfmgle.2.1 k (k + 1) k.le_succ)).trans ih }
515+
end
516+
517+
/-- A predictable martingale is a.e. equal to its initial state. -/
518+
lemma martingale.eq_zero_of_predicatable [sigma_finite_filtration μ 𝒢]
519+
{f : ℕ → Ω → E} (hfmgle : martingale f 𝒢 μ) (hfadp : adapted 𝒢 (λ n, f (n + 1))) (n : ℕ) :
520+
f n =ᵐ[μ] f 0 :=
521+
begin
522+
induction n with k ih,
523+
{ refl },
524+
{ exact ((germ.coe_eq.mp (congr_arg coe $ condexp_of_strongly_measurable (𝒢.le _) (hfadp _)
525+
(hfmgle.integrable _))).symm.trans (hfmgle.2 k (k + 1) k.le_succ)).trans ih }
526+
end
527+
493528
namespace submartingale
494529

495530
@[protected]

src/probability/martingale/centering.lean

Lines changed: 35 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -136,6 +136,41 @@ begin
136136
refl,
137137
end
138138

139+
-- The following two lemmas demonstrate the essential uniqueness of the decomposition
140+
lemma martingale_part_add_ae_eq [sigma_finite_filtration μ ℱ] {f g : ℕ → Ω → E}
141+
(hf : martingale f ℱ μ) (hg : adapted ℱ (λ n, g (n + 1))) (hg0 : g 0 = 0)
142+
(hgint : ∀ n, integrable (g n) μ) (n : ℕ) :
143+
martingale_part ℱ μ (f + g) n =ᵐ[μ] f n :=
144+
begin
145+
set h := f - martingale_part ℱ μ (f + g) with hhdef,
146+
have hh : h = predictable_part ℱ μ (f + g) - g,
147+
{ rw [hhdef, sub_eq_sub_iff_add_eq_add, add_comm (predictable_part ℱ μ (f + g)),
148+
martingale_part_add_predictable_part] },
149+
have hhpred : adapted ℱ (λ n, h (n + 1)),
150+
{ rw hh,
151+
exact adapted_predictable_part.sub hg },
152+
have hhmgle : martingale h ℱ μ := hf.sub (martingale_martingale_part
153+
(hf.adapted.add $ predictable.adapted hg $ hg0.symm ▸ strongly_measurable_zero) $
154+
λ n, (hf.integrable n).add $ hgint n),
155+
refine (eventually_eq_iff_sub.2 _).symm,
156+
filter_upwards [hhmgle.eq_zero_of_predicatable hhpred n] with ω hω,
157+
rw [hhdef, pi.sub_apply] at hω,
158+
rw [hω, pi.sub_apply, martingale_part],
159+
simp [hg0],
160+
end
161+
162+
lemma predictable_part_add_ae_eq [sigma_finite_filtration μ ℱ] {f g : ℕ → Ω → E}
163+
(hf : martingale f ℱ μ) (hg : adapted ℱ (λ n, g (n + 1))) (hg0 : g 0 = 0)
164+
(hgint : ∀ n, integrable (g n) μ) (n : ℕ) :
165+
predictable_part ℱ μ (f + g) n =ᵐ[μ] g n :=
166+
begin
167+
filter_upwards [martingale_part_add_ae_eq hf hg hg0 hgint n] with ω hω,
168+
rw ← add_right_inj (f n ω),
169+
conv_rhs { rw [← pi.add_apply, ← pi.add_apply,
170+
← martingale_part_add_predictable_part ℱ μ (f + g)] },
171+
rw [pi.add_apply, pi.add_apply, hω],
172+
end
173+
139174
section difference
140175

141176
lemma predictable_part_bdd_difference {R : ℝ≥0} {f : ℕ → Ω → ℝ}

src/probability/process/adapted.lean

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -216,4 +216,13 @@ begin
216216
{ exact strongly_measurable_const, },
217217
end
218218

219+
-- this dot notation will make more sense once we have a more general definition for predictable
220+
lemma predictable.adapted {f : filtration ℕ m} {u : ℕ → Ω → β}
221+
(hu : adapted f (λ n, u (n + 1))) (hu0 : strongly_measurable[f 0] (u 0)) :
222+
adapted f u :=
223+
λ n, match n with
224+
| 0 := hu0
225+
| n + 1 := (hu n).mono (f.mono n.le_succ)
226+
end
227+
219228
end measure_theory

0 commit comments

Comments
 (0)