Documentation

Mathlib.Probability.Process.Stopping

Stopping times, stopped processes and stopped values #

Definition and properties of stopping times.

Main definitions #

Main results #

Implementation notes #

For a filtration on a type ι, we define stopping times as functions from the measurable space Ω to WithTop ι, which allows stopping times that can take an infinite value, represented by ⊤ : WithTop ι.

This means that if we have a process X : ι → Ω → β and a stopping time τ : Ω → WithTop ι, then to consider the value of X at the stopping time τ ω, we need to write X (τ ω).untopA ω, in which (τ ω).untopA is the value of τ ω in ι if τ ω ≠ ⊤ and some arbitrary value if τ ω = ⊤.

While indexing would be more convenient if we defined stopping times as functions from Ω to ι, this would prevent us from using stopping times as in standard mathematical literature, where a typical example of stopping time is the first time an event occurs, which may never happen. Consider for example the first time a coin lands heads when flipping it infinitely many times: this is almost surely finite, but possibly infinite. We could also not use a function Ω → ι with arbitrary value for the infinite case, because this would be incompatible with the stopping time property.

Tags #

stopping time, stochastic process

Stopping times #

def MeasureTheory.IsStoppingTime {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] (f : Filtration ι m) (τ : ΩWithTop ι) :

A stopping time with respect to some filtration f is a function τ such that for all i, the preimage of {j | j ≤ i} along τ is measurable with respect to f i.

Intuitively, the stopping time τ describes some stopping rule such that at time i, we may determine it with the information we have at time i.

Equations
Instances For
    theorem MeasureTheory.isStoppingTime_const {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] (f : Filtration ι m) (i : ι) :
    IsStoppingTime f fun (x : Ω) => i
    theorem MeasureTheory.IsStoppingTime.measurableSet_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (i : ι) :
    MeasurableSet {ω : Ω | τ ω i}
    theorem MeasureTheory.IsStoppingTime.measurableSet_lt_of_pred {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [PredOrder ι] ( : IsStoppingTime f τ) (i : ι) :
    MeasurableSet {ω : Ω | τ ω < i}
    theorem MeasureTheory.IsStoppingTime.measurableSet_eq_of_countable_range {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [PartialOrder ι] {τ : ΩWithTop ι} {f : Filtration ι m} ( : IsStoppingTime f τ) (h_countable : (Set.range τ).Countable) (i : ι) :
    MeasurableSet {ω : Ω | τ ω = i}
    theorem MeasureTheory.IsStoppingTime.measurableSet_eq_of_countable {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [PartialOrder ι] {τ : ΩWithTop ι} {f : Filtration ι m} [Countable ι] ( : IsStoppingTime f τ) (i : ι) :
    MeasurableSet {ω : Ω | τ ω = i}
    theorem MeasureTheory.IsStoppingTime.measurableSet_lt_of_countable_range {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [PartialOrder ι] {τ : ΩWithTop ι} {f : Filtration ι m} ( : IsStoppingTime f τ) (h_countable : (Set.range τ).Countable) (i : ι) :
    MeasurableSet {ω : Ω | τ ω < i}
    theorem MeasureTheory.IsStoppingTime.measurableSet_lt_of_countable {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [PartialOrder ι] {τ : ΩWithTop ι} {f : Filtration ι m} [Countable ι] ( : IsStoppingTime f τ) (i : ι) :
    MeasurableSet {ω : Ω | τ ω < i}
    theorem MeasureTheory.IsStoppingTime.measurableSet_ge_of_countable_range {Ω : Type u_1} {m : MeasurableSpace Ω} {ι : Type u_4} [LinearOrder ι] {τ : ΩWithTop ι} {f : Filtration ι m} ( : IsStoppingTime f τ) (h_countable : (Set.range τ).Countable) (i : ι) :
    MeasurableSet {ω : Ω | i τ ω}
    theorem MeasureTheory.IsStoppingTime.measurableSet_ge_of_countable {Ω : Type u_1} {m : MeasurableSpace Ω} {ι : Type u_4} [LinearOrder ι] {τ : ΩWithTop ι} {f : Filtration ι m} [Countable ι] ( : IsStoppingTime f τ) (i : ι) :
    MeasurableSet {ω : Ω | i τ ω}
    theorem MeasureTheory.IsStoppingTime.measurableSet_gt {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (i : ι) :
    MeasurableSet {ω : Ω | i < τ ω}
    theorem MeasureTheory.IsStoppingTime.measurableSet_lt_of_isLUB {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] ( : IsStoppingTime f τ) (i : ι) (h_lub : IsLUB (Set.Iio i) i) :
    MeasurableSet {ω : Ω | τ ω < i}

    Auxiliary lemma for MeasureTheory.IsStoppingTime.measurableSet_lt.

    theorem MeasureTheory.IsStoppingTime.measurableSet_lt {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] ( : IsStoppingTime f τ) (i : ι) :
    MeasurableSet {ω : Ω | τ ω < i}
    theorem MeasureTheory.IsStoppingTime.measurableSet_ge {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] ( : IsStoppingTime f τ) (i : ι) :
    MeasurableSet {ω : Ω | i τ ω}
    theorem MeasureTheory.IsStoppingTime.measurableSet_eq {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] ( : IsStoppingTime f τ) (i : ι) :
    MeasurableSet {ω : Ω | τ ω = i}
    theorem MeasureTheory.IsStoppingTime.measurableSet_eq_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] ( : IsStoppingTime f τ) {i j : ι} (hle : i j) :
    MeasurableSet {ω : Ω | τ ω = i}
    theorem MeasureTheory.IsStoppingTime.measurableSet_lt_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] ( : IsStoppingTime f τ) {i j : ι} (hle : i j) :
    MeasurableSet {ω : Ω | τ ω < i}
    theorem MeasureTheory.isStoppingTime_of_measurableSet_eq {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] [Countable ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : ∀ (i : ι), MeasurableSet {ω : Ω | τ ω = i}) :
    theorem MeasureTheory.IsStoppingTime.max {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
    IsStoppingTime f fun (ω : Ω) => max (τ ω) (π ω)
    theorem MeasureTheory.IsStoppingTime.max_const {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (i : ι) :
    IsStoppingTime f fun (ω : Ω) => max (τ ω) i
    theorem MeasureTheory.IsStoppingTime.min {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
    IsStoppingTime f fun (ω : Ω) => min (τ ω) (π ω)
    theorem MeasureTheory.IsStoppingTime.min_const {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (i : ι) :
    IsStoppingTime f fun (ω : Ω) => min (τ ω) i
    theorem MeasureTheory.IsStoppingTime.add_const {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [AddGroup ι] [Preorder ι] [AddRightMono ι] [AddLeftMono ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) {i : ι} (hi : 0 i) :
    IsStoppingTime f fun (ω : Ω) => τ ω + i
    theorem MeasureTheory.IsStoppingTime.add_const_nat {Ω : Type u_1} {m : MeasurableSpace Ω} {f : Filtration m} {τ : ΩWithTop } ( : IsStoppingTime f τ) {i : } :
    IsStoppingTime f fun (ω : Ω) => τ ω + i
    theorem MeasureTheory.IsStoppingTime.add {Ω : Type u_1} {m : MeasurableSpace Ω} {f : Filtration m} {τ π : ΩWithTop } ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
    IsStoppingTime f (τ + π)
    def MeasureTheory.IsStoppingTime.measurableSpace {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) :

    The associated σ-algebra with a stopping time.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      theorem MeasureTheory.IsStoppingTime.measurableSet {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (s : Set Ω) :
      MeasurableSet s MeasurableSet s ∀ (i : ι), MeasurableSet (s {ω : Ω | τ ω i})
      theorem MeasureTheory.IsStoppingTime.measurableSpace_mono {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} ( : IsStoppingTime f τ) ( : IsStoppingTime f π) (hle : τ π) :
      theorem MeasureTheory.IsStoppingTime.measurableSpace_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) :
      @[deprecated MeasureTheory.IsStoppingTime.measurableSpace_le (since := "2025-09-08")]
      theorem MeasureTheory.IsStoppingTime.measurableSpace_le_of_countable {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) :

      Alias of MeasureTheory.IsStoppingTime.measurableSpace_le.

      @[deprecated MeasureTheory.IsStoppingTime.measurableSpace_le (since := "2025-09-08")]
      theorem MeasureTheory.IsStoppingTime.measurableSpace_le_of_countable_range {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) :

      Alias of MeasureTheory.IsStoppingTime.measurableSpace_le.

      @[simp]
      theorem MeasureTheory.IsStoppingTime.measurableSpace_const {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] (f : Filtration ι m) (i : ι) :
      .measurableSpace = f i
      theorem MeasureTheory.IsStoppingTime.measurableSet_inter_eq_iff {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (s : Set Ω) (i : ι) :
      MeasurableSet (s {ω : Ω | τ ω = i}) MeasurableSet (s {ω : Ω | τ ω = i})
      theorem MeasureTheory.IsStoppingTime.measurableSpace_le_of_le_const {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) {i : ι} (hτ_le : ∀ (ω : Ω), τ ω i) :
      .measurableSpace f i
      theorem MeasureTheory.IsStoppingTime.measurableSpace_le_of_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) {n : ι} (hτ_le : ∀ (ω : Ω), τ ω n) :
      theorem MeasureTheory.IsStoppingTime.le_measurableSpace_of_const_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) {i : ι} (hτ_le : ∀ (ω : Ω), i τ ω) :
      f i .measurableSpace
      instance MeasureTheory.IsStoppingTime.sigmaFinite_stopping_time {Ω : Type u_1} {m : MeasurableSpace Ω} {ι : Type u_4} [SemilatticeSup ι] [OrderBot ι] {μ : Measure Ω} {f : Filtration ι m} {τ : ΩWithTop ι} [SigmaFiniteFiltration μ f] ( : IsStoppingTime f τ) :
      SigmaFinite (μ.trim )
      instance MeasureTheory.IsStoppingTime.sigmaFinite_stopping_time_of_le {Ω : Type u_1} {m : MeasurableSpace Ω} {ι : Type u_4} [SemilatticeSup ι] [OrderBot ι] {μ : Measure Ω} {f : Filtration ι m} {τ : ΩWithTop ι} [SigmaFiniteFiltration μ f] ( : IsStoppingTime f τ) {n : ι} (hτ_le : ∀ (ω : Ω), τ ω n) :
      SigmaFinite (μ.trim )
      theorem MeasureTheory.IsStoppingTime.measurableSet_le' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (i : ι) :
      MeasurableSet {ω : Ω | τ ω i}
      theorem MeasureTheory.IsStoppingTime.measurableSet_gt' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (i : ι) :
      MeasurableSet {ω : Ω | i < τ ω}
      theorem MeasureTheory.IsStoppingTime.measurableSet_eq' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] ( : IsStoppingTime f τ) (i : ι) :
      MeasurableSet {ω : Ω | τ ω = i}
      theorem MeasureTheory.IsStoppingTime.measurableSet_ge' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] ( : IsStoppingTime f τ) (i : ι) :
      MeasurableSet {ω : Ω | i τ ω}
      theorem MeasureTheory.IsStoppingTime.measurableSet_lt' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] ( : IsStoppingTime f τ) (i : ι) :
      MeasurableSet {ω : Ω | τ ω < i}
      theorem MeasureTheory.IsStoppingTime.measurableSet_eq_of_countable_range' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (h_countable : (Set.range τ).Countable) (i : ι) :
      MeasurableSet {ω : Ω | τ ω = i}
      theorem MeasureTheory.IsStoppingTime.measurableSet_eq_of_countable' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [Countable ι] ( : IsStoppingTime f τ) (i : ι) :
      MeasurableSet {ω : Ω | τ ω = i}
      theorem MeasureTheory.IsStoppingTime.measurableSet_ge_of_countable_range' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (h_countable : (Set.range τ).Countable) (i : ι) :
      MeasurableSet {ω : Ω | i τ ω}
      theorem MeasureTheory.IsStoppingTime.measurableSet_ge_of_countable' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [Countable ι] ( : IsStoppingTime f τ) (i : ι) :
      MeasurableSet {ω : Ω | i τ ω}
      theorem MeasureTheory.IsStoppingTime.measurableSet_lt_of_countable_range' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (h_countable : (Set.range τ).Countable) (i : ι) :
      MeasurableSet {ω : Ω | τ ω < i}
      theorem MeasureTheory.IsStoppingTime.measurableSet_lt_of_countable' {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [Countable ι] ( : IsStoppingTime f τ) (i : ι) :
      MeasurableSet {ω : Ω | τ ω < i}
      theorem MeasureTheory.IsStoppingTime.measurable {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] ( : IsStoppingTime f τ) :
      theorem MeasureTheory.IsStoppingTime.measurable_of_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] ( : IsStoppingTime f τ) {i : ι} (hτ_le : ∀ (ω : Ω), τ ω i) :
      theorem MeasureTheory.IsStoppingTime.measurableSpace_min {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
      theorem MeasureTheory.IsStoppingTime.measurableSet_min_iff {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} ( : IsStoppingTime f τ) ( : IsStoppingTime f π) (s : Set Ω) :
      theorem MeasureTheory.IsStoppingTime.measurableSpace_min_const {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) {i : ι} :
      .measurableSpace = .measurableSpacef i
      theorem MeasureTheory.IsStoppingTime.measurableSet_min_const_iff {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (s : Set Ω) {i : ι} :
      theorem MeasureTheory.IsStoppingTime.measurableSet_inter_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} [TopologicalSpace ι] [SecondCountableTopology ι] [OrderTopology ι] ( : IsStoppingTime f τ) ( : IsStoppingTime f π) (s : Set Ω) (hs : MeasurableSet s) :
      MeasurableSet (s {ω : Ω | τ ω π ω})
      theorem MeasureTheory.IsStoppingTime.measurableSet_inter_le_iff {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} [TopologicalSpace ι] [SecondCountableTopology ι] [OrderTopology ι] ( : IsStoppingTime f τ) ( : IsStoppingTime f π) (s : Set Ω) :
      MeasurableSet (s {ω : Ω | τ ω π ω}) MeasurableSet (s {ω : Ω | τ ω π ω})
      theorem MeasureTheory.IsStoppingTime.measurableSet_inter_le_const_iff {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ : ΩWithTop ι} ( : IsStoppingTime f τ) (s : Set Ω) (i : ι) :
      MeasurableSet (s {ω : Ω | τ ω i}) MeasurableSet (s {ω : Ω | τ ω i})
      theorem MeasureTheory.IsStoppingTime.measurableSet_le_stopping_time {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} [TopologicalSpace ι] [SecondCountableTopology ι] [OrderTopology ι] ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
      MeasurableSet {ω : Ω | τ ω π ω}
      theorem MeasureTheory.IsStoppingTime.measurableSet_stopping_time_le_min {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} [TopologicalSpace ι] [SecondCountableTopology ι] [OrderTopology ι] ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
      MeasurableSet {ω : Ω | τ ω π ω}
      theorem MeasureTheory.IsStoppingTime.measurableSet_stopping_time_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} [TopologicalSpace ι] [SecondCountableTopology ι] [OrderTopology ι] ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
      MeasurableSet {ω : Ω | τ ω π ω}
      theorem MeasureTheory.IsStoppingTime.measurableSet_eq_stopping_time_min {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
      MeasurableSet {ω : Ω | τ ω = π ω}
      theorem MeasureTheory.IsStoppingTime.measurableSet_eq_stopping_time {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
      MeasurableSet {ω : Ω | τ ω = π ω}
      @[deprecated MeasureTheory.IsStoppingTime.measurableSet_eq_stopping_time (since := "2025-09-08")]
      theorem MeasureTheory.IsStoppingTime.measurableSet_eq_stopping_time_of_countable {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {f : Filtration ι m} {τ π : ΩWithTop ι} [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] ( : IsStoppingTime f τ) ( : IsStoppingTime f π) :
      MeasurableSet {ω : Ω | τ ω = π ω}

      Alias of MeasureTheory.IsStoppingTime.measurableSet_eq_stopping_time.

      Stopped value and stopped process #

      noncomputable def MeasureTheory.stoppedValue {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} [Nonempty ι] (u : ιΩβ) (τ : ΩWithTop ι) :
      Ωβ

      Given a map u : ι → Ω → E, its stopped value with respect to the stopping time τ is the map x ↦ u (τ ω) ω.

      Equations
      Instances For
        theorem MeasureTheory.stoppedValue_const {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} [Nonempty ι] (u : ιΩβ) (i : ι) :
        (stoppedValue u fun (x : Ω) => i) = u i
        noncomputable def MeasureTheory.stoppedProcess {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} [Nonempty ι] [LinearOrder ι] (u : ιΩβ) (τ : ΩWithTop ι) :
        ιΩβ

        Given a map u : ι → Ω → E, the stopped process with respect to τ is u i ω if i ≤ τ ω, and u (τ ω) ω otherwise.

        Intuitively, the stopped process stops evolving once the stopping time has occurred.

        Equations
        Instances For
          theorem MeasureTheory.stoppedProcess_eq_stoppedValue {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} [Nonempty ι] [LinearOrder ι] {u : ιΩβ} {τ : ΩWithTop ι} :
          stoppedProcess u τ = fun (i : ι) => stoppedValue u fun (ω : Ω) => min (↑i) (τ ω)
          theorem MeasureTheory.stoppedValue_stoppedProcess {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} [Nonempty ι] [LinearOrder ι] {u : ιΩβ} {τ σ : ΩWithTop ι} :
          stoppedValue (stoppedProcess u τ) σ = fun (ω : Ω) => if σ ω then stoppedValue u (fun (ω : Ω) => min (σ ω) (τ ω)) ω else stoppedValue u (fun (ω : Ω) => min (↑(Classical.arbitrary ι)) (τ ω)) ω
          theorem MeasureTheory.stoppedValue_stoppedProcess_ae_eq {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] [LinearOrder ι] {u : ιΩβ} {τ σ : ΩWithTop ι} {μ : Measure Ω} ( : ∀ᵐ (ω : Ω) μ, σ ω ) :
          stoppedValue (stoppedProcess u τ) σ =ᶠ[ae μ] stoppedValue u fun (ω : Ω) => min (σ ω) (τ ω)
          theorem MeasureTheory.stoppedProcess_eq_of_le {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} [Nonempty ι] [LinearOrder ι] {u : ιΩβ} {τ : ΩWithTop ι} {i : ι} {ω : Ω} (h : i τ ω) :
          stoppedProcess u τ i ω = u i ω
          theorem MeasureTheory.stoppedProcess_eq_of_ge {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} [Nonempty ι] [LinearOrder ι] {u : ιΩβ} {τ : ΩWithTop ι} {i : ι} {ω : Ω} (h : τ ω i) :
          stoppedProcess u τ i ω = u (τ ω).untopA ω
          theorem MeasureTheory.progMeasurable_min_stopping_time {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] [LinearOrder ι] [MeasurableSpace ι] [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] [BorelSpace ι] {τ : ΩWithTop ι} {f : Filtration ι m} [TopologicalSpace.PseudoMetrizableSpace ι] ( : IsStoppingTime f τ) :
          ProgMeasurable f fun (i : ι) (ω : Ω) => (min (↑i) (τ ω)).untopA
          theorem MeasureTheory.stronglyMeasurable_stoppedValue_of_le {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] [LinearOrder ι] [MeasurableSpace ι] [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] [BorelSpace ι] [TopologicalSpace β] {u : ιΩβ} {τ : ΩWithTop ι} {f : Filtration ι m} (h : ProgMeasurable f u) ( : IsStoppingTime f τ) {n : ι} (hτ_le : ∀ (ω : Ω), τ ω n) :
          theorem MeasureTheory.measurableSet_preimage_stoppedValue_inter {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] [LinearOrder ι] [MeasurableSpace ι] [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] [BorelSpace ι] [TopologicalSpace β] {u : ιΩβ} {τ : ΩWithTop ι} {f : Filtration ι m} [TopologicalSpace.PseudoMetrizableSpace β] [MeasurableSpace β] [BorelSpace β] (hf_prog : ProgMeasurable f u) ( : IsStoppingTime f τ) {t : Set β} (ht : MeasurableSet t) (i : ι) :
          MeasurableSet (stoppedValue u τ ⁻¹' t {ω : Ω | τ ω i})
          theorem MeasureTheory.measurable_stoppedValue {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] [LinearOrder ι] [MeasurableSpace ι] [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] [BorelSpace ι] [TopologicalSpace β] {u : ιΩβ} {τ : ΩWithTop ι} {f : Filtration ι m} [TopologicalSpace.PseudoMetrizableSpace β] [MeasurableSpace β] [BorelSpace β] (hf_prog : ProgMeasurable f u) ( : IsStoppingTime f τ) :
          theorem MeasureTheory.stoppedValue_eq_of_mem_finset {Ω : Type u_1} {ι : Type u_3} [Nonempty ι] {τ : ΩWithTop ι} {E : Type u_4} {u : ιΩE} [AddCommMonoid E] {s : Finset ι} (hbdd : ∀ (ω : Ω), τ ω WithTop.some '' s) :
          stoppedValue u τ = is, {ω : Ω | τ ω = i}.indicator (u i)
          theorem MeasureTheory.stoppedValue_eq' {Ω : Type u_1} {ι : Type u_3} [Nonempty ι] {τ : ΩWithTop ι} {E : Type u_4} {u : ιΩE} [Preorder ι] [LocallyFiniteOrderBot ι] [AddCommMonoid E] {N : ι} (hbdd : ∀ (ω : Ω), τ ω N) :
          stoppedValue u τ = iFinset.Iic N, {ω : Ω | τ ω = i}.indicator (u i)
          theorem MeasureTheory.stoppedProcess_eq_of_mem_finset {Ω : Type u_1} {ι : Type u_3} [Nonempty ι] {τ : ΩWithTop ι} {E : Type u_4} {u : ιΩE} [LinearOrder ι] [AddCommMonoid E] {s : Finset ι} (n : ι) (hbdd : ∀ (ω : Ω), τ ω < nτ ω WithTop.some '' s) :
          stoppedProcess u τ n = {a : Ω | n τ a}.indicator (u n) + is with i < n, {ω : Ω | τ ω = i}.indicator (u i)
          theorem MeasureTheory.stoppedProcess_eq'' {Ω : Type u_1} {ι : Type u_3} [Nonempty ι] {τ : ΩWithTop ι} {E : Type u_4} {u : ιΩE} [LinearOrder ι] [LocallyFiniteOrderBot ι] [AddCommMonoid E] (n : ι) :
          stoppedProcess u τ n = {a : Ω | n τ a}.indicator (u n) + iFinset.Iio n, {ω : Ω | τ ω = i}.indicator (u i)
          theorem MeasureTheory.memLp_stoppedValue_of_mem_finset {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] {μ : Measure Ω} {τ : ΩWithTop ι} {E : Type u_4} {p : ENNReal} {u : ιΩE} [PartialOrder ι] { : Filtration ι m} [NormedAddCommGroup E] ( : IsStoppingTime τ) (hu : ∀ (n : ι), MemLp (u n) p μ) {s : Finset ι} (hbdd : ∀ (ω : Ω), τ ω WithTop.some '' s) :
          MemLp (stoppedValue u τ) p μ
          theorem MeasureTheory.memLp_stoppedValue {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] {μ : Measure Ω} {τ : ΩWithTop ι} {E : Type u_4} {p : ENNReal} {u : ιΩE} [PartialOrder ι] { : Filtration ι m} [NormedAddCommGroup E] [LocallyFiniteOrderBot ι] ( : IsStoppingTime τ) (hu : ∀ (n : ι), MemLp (u n) p μ) {N : ι} (hbdd : ∀ (ω : Ω), τ ω N) :
          MemLp (stoppedValue u τ) p μ
          theorem MeasureTheory.integrable_stoppedValue_of_mem_finset {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] {μ : Measure Ω} {τ : ΩWithTop ι} {E : Type u_4} {u : ιΩE} [PartialOrder ι] { : Filtration ι m} [NormedAddCommGroup E] ( : IsStoppingTime τ) (hu : ∀ (n : ι), Integrable (u n) μ) {s : Finset ι} (hbdd : ∀ (ω : Ω), τ ω WithTop.some '' s) :
          theorem MeasureTheory.integrable_stoppedValue {Ω : Type u_1} (ι : Type u_3) {m : MeasurableSpace Ω} [Nonempty ι] {μ : Measure Ω} {τ : ΩWithTop ι} {E : Type u_4} {u : ιΩE} [PartialOrder ι] { : Filtration ι m} [NormedAddCommGroup E] [LocallyFiniteOrderBot ι] ( : IsStoppingTime τ) (hu : ∀ (n : ι), Integrable (u n) μ) {N : ι} (hbdd : ∀ (ω : Ω), τ ω N) :
          theorem MeasureTheory.memLp_stoppedProcess_of_mem_finset {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] {μ : Measure Ω} {τ : ΩWithTop ι} {E : Type u_4} {p : ENNReal} {u : ιΩE} [LinearOrder ι] [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] { : Filtration ι m} [NormedAddCommGroup E] ( : IsStoppingTime τ) (hu : ∀ (n : ι), MemLp (u n) p μ) (n : ι) {s : Finset ι} (hbdd : ∀ (ω : Ω), τ ω < nτ ω WithTop.some '' s) :
          MemLp (stoppedProcess u τ n) p μ
          theorem MeasureTheory.memLp_stoppedProcess {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] {μ : Measure Ω} {τ : ΩWithTop ι} {E : Type u_4} {p : ENNReal} {u : ιΩE} [LinearOrder ι] [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] { : Filtration ι m} [NormedAddCommGroup E] [LocallyFiniteOrderBot ι] ( : IsStoppingTime τ) (hu : ∀ (n : ι), MemLp (u n) p μ) (n : ι) :
          MemLp (stoppedProcess u τ n) p μ
          theorem MeasureTheory.integrable_stoppedProcess_of_mem_finset {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] {μ : Measure Ω} {τ : ΩWithTop ι} {E : Type u_4} {u : ιΩE} [LinearOrder ι] [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] { : Filtration ι m} [NormedAddCommGroup E] ( : IsStoppingTime τ) (hu : ∀ (n : ι), Integrable (u n) μ) (n : ι) {s : Finset ι} (hbdd : ∀ (ω : Ω), τ ω < nτ ω WithTop.some '' s) :
          theorem MeasureTheory.integrable_stoppedProcess {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Nonempty ι] {μ : Measure Ω} {τ : ΩWithTop ι} {E : Type u_4} {u : ιΩE} [LinearOrder ι] [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] { : Filtration ι m} [NormedAddCommGroup E] [LocallyFiniteOrderBot ι] ( : IsStoppingTime τ) (hu : ∀ (n : ι), Integrable (u n) μ) (n : ι) :
          theorem MeasureTheory.Adapted.stoppedProcess {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} {m : MeasurableSpace Ω} [TopologicalSpace β] [TopologicalSpace.PseudoMetrizableSpace β] [Nonempty ι] [LinearOrder ι] [TopologicalSpace ι] [SecondCountableTopology ι] [OrderTopology ι] [MeasurableSpace ι] [BorelSpace ι] {f : Filtration ι m} {u : ιΩβ} {τ : ΩWithTop ι} [TopologicalSpace.MetrizableSpace ι] (hu : Adapted f u) (hu_cont : ∀ (ω : Ω), Continuous fun (i : ι) => u i ω) ( : IsStoppingTime f τ) :

          The stopped process of an adapted process with continuous paths is adapted.

          If the indexing order has the discrete topology, then the stopped process of an adapted process is adapted.

          theorem MeasureTheory.Adapted.stronglyMeasurable_stoppedProcess {Ω : Type u_1} {β : Type u_2} {ι : Type u_3} {m : MeasurableSpace Ω} [TopologicalSpace β] [TopologicalSpace.PseudoMetrizableSpace β] [Nonempty ι] [LinearOrder ι] [TopologicalSpace ι] [SecondCountableTopology ι] [OrderTopology ι] [MeasurableSpace ι] [BorelSpace ι] {f : Filtration ι m} {u : ιΩβ} {τ : ΩWithTop ι} [TopologicalSpace.MetrizableSpace ι] (hu : Adapted f u) (hu_cont : ∀ (ω : Ω), Continuous fun (i : ι) => u i ω) ( : IsStoppingTime f τ) (n : ι) :

          Filtrations indexed by #

          theorem MeasureTheory.stoppedValue_sub_eq_sum {Ω : Type u_1} {β : Type u_2} {u : Ωβ} {τ π : Ωℕ∞} [AddCommGroup β] (hle : τ π) ( : ∀ (ω : Ω), (π ω) ) :
          stoppedValue u π - stoppedValue u τ = fun (ω : Ω) => (∑ iFinset.Ico (WithTop.untopA (τ ω)) (WithTop.untopA (π ω)), (u (i + 1) - u i)) ω
          theorem MeasureTheory.stoppedValue_sub_eq_sum' {Ω : Type u_1} {β : Type u_2} {u : Ωβ} {τ π : Ωℕ∞} [AddCommGroup β] (hle : τ π) {N : } (hbdd : ∀ (ω : Ω), π ω N) :
          stoppedValue u π - stoppedValue u τ = fun (ω : Ω) => (∑ iFinset.range (N + 1), {ω : Ω | τ ω i i < π ω}.indicator (u (i + 1) - u i)) ω
          theorem MeasureTheory.stoppedValue_eq {Ω : Type u_1} {β : Type u_2} {u : Ωβ} {τ : Ωℕ∞} [AddCommMonoid β] {N : } (hbdd : ∀ (ω : Ω), τ ω N) :
          stoppedValue u τ = fun (x : Ω) => (∑ iFinset.range (N + 1), {ω : Ω | τ ω = i}.indicator (u i)) x
          theorem MeasureTheory.stoppedProcess_eq {Ω : Type u_1} {β : Type u_2} {u : Ωβ} {τ : Ωℕ∞} [AddCommMonoid β] (n : ) :
          stoppedProcess u τ n = {a : Ω | n τ a}.indicator (u n) + iFinset.range n, {ω : Ω | τ ω = i}.indicator (u i)
          theorem MeasureTheory.stoppedProcess_eq' {Ω : Type u_1} {β : Type u_2} {u : Ωβ} {τ : Ωℕ∞} [AddCommMonoid β] (n : ) :
          stoppedProcess u τ n = {a : Ω | n + 1 τ a}.indicator (u n) + iFinset.range (n + 1), {a : Ω | τ a = i}.indicator (u i)
          theorem MeasureTheory.IsStoppingTime.piecewise_of_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {𝒢 : Filtration ι m} {τ η : ΩWithTop ι} {i : ι} {s : Set Ω} [DecidablePred fun (x : Ω) => x s] (hτ_st : IsStoppingTime 𝒢 τ) (hη_st : IsStoppingTime 𝒢 η) ( : ∀ (ω : Ω), i τ ω) ( : ∀ (ω : Ω), i η ω) (hs : MeasurableSet s) :
          IsStoppingTime 𝒢 (s.piecewise τ η)

          Given stopping times τ and η which are bounded below, Set.piecewise s τ η is also a stopping time with respect to the same filtration.

          theorem MeasureTheory.isStoppingTime_piecewise_const {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [Preorder ι] {𝒢 : Filtration ι m} {i j : ι} {s : Set Ω} [DecidablePred fun (x : Ω) => x s] (hij : i j) (hs : MeasurableSet s) :
          IsStoppingTime 𝒢 (s.piecewise (fun (x : Ω) => i) fun (x : Ω) => j)
          theorem MeasureTheory.stoppedValue_piecewise_const {Ω : Type u_1} {s : Set Ω} [DecidablePred fun (x : Ω) => x s] {ι' : Type u_4} [Nonempty ι'] {i j : ι'} {f : ι'Ω} :
          stoppedValue f (s.piecewise (fun (x : Ω) => i) fun (x : Ω) => j) = s.piecewise (f i) (f j)
          theorem MeasureTheory.stoppedValue_piecewise_const' {Ω : Type u_1} {s : Set Ω} [DecidablePred fun (x : Ω) => x s] {ι' : Type u_4} [Nonempty ι'] {i j : ι'} {f : ι'Ω} :
          stoppedValue f (s.piecewise (fun (x : Ω) => i) fun (x : Ω) => j) = s.indicator (f i) + s.indicator (f j)

          Conditional expectation with respect to the σ-algebra generated by a stopping time #

          theorem MeasureTheory.condExp_stopping_time_ae_eq_restrict_eq_of_countable_range {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {μ : Measure Ω} { : Filtration ι m} {τ : ΩWithTop ι} {E : Type u_4} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] {f : ΩE} [SigmaFiniteFiltration μ ] ( : IsStoppingTime τ) (h_countable : (Set.range τ).Countable) [SigmaFinite (μ.trim )] (i : ι) :
          μ[f|.measurableSpace] =ᶠ[ae (μ.restrict {x : Ω | τ x = i})] μ[f| i]
          theorem MeasureTheory.condExp_stopping_time_ae_eq_restrict_eq_of_countable {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {μ : Measure Ω} { : Filtration ι m} {τ : ΩWithTop ι} {E : Type u_4} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] {f : ΩE} [Countable ι] [SigmaFiniteFiltration μ ] ( : IsStoppingTime τ) [SigmaFinite (μ.trim )] (i : ι) :
          μ[f|.measurableSpace] =ᶠ[ae (μ.restrict {x : Ω | τ x = i})] μ[f| i]
          theorem MeasureTheory.condExp_min_stopping_time_ae_eq_restrict_le_const {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {μ : Measure Ω} { : Filtration ι m} {τ : ΩWithTop ι} {E : Type u_4} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] {f : ΩE} ( : IsStoppingTime τ) (i : ι) [SigmaFinite (μ.trim )] :
          μ[f|.measurableSpace] =ᶠ[ae (μ.restrict {x : Ω | τ x i})] μ[f|.measurableSpace]
          theorem MeasureTheory.condExp_stopping_time_ae_eq_restrict_eq {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {μ : Measure Ω} { : Filtration ι m} {τ : ΩWithTop ι} {E : Type u_4} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] {f : ΩE} [TopologicalSpace ι] [OrderTopology ι] [FirstCountableTopology ι] [SigmaFiniteFiltration μ ] ( : IsStoppingTime τ) [SigmaFinite (μ.trim )] (i : ι) :
          μ[f|.measurableSpace] =ᶠ[ae (μ.restrict {x : Ω | τ x = i})] μ[f| i]
          theorem MeasureTheory.condExp_min_stopping_time_ae_eq_restrict_le {Ω : Type u_1} {ι : Type u_3} {m : MeasurableSpace Ω} [LinearOrder ι] {μ : Measure Ω} { : Filtration ι m} {τ σ : ΩWithTop ι} {E : Type u_4} [NormedAddCommGroup E] [NormedSpace E] [CompleteSpace E] {f : ΩE} [TopologicalSpace ι] [OrderTopology ι] [SecondCountableTopology ι] ( : IsStoppingTime τ) ( : IsStoppingTime σ) [SigmaFinite (μ.trim )] :
          μ[f|.measurableSpace] =ᶠ[ae (μ.restrict {x : Ω | τ x σ x})] μ[f|.measurableSpace]