Projective measure families and projective limits #
A family of measures indexed by finite sets of ι is projective if, for finite sets J ⊆ I,
the projection from ∀ i : I, α i to ∀ i : J, α i maps P I to P J.
A measure μ is the projective limit of such a family of measures if for all I : Finset ι,
the projection from ∀ i, α i to ∀ i : I, α i maps μ to P I.
Main definitions #
- MeasureTheory.IsProjectiveMeasureFamily:- P : ∀ J : Finset ι, Measure (∀ j : J, α j)is projective if the projection from- ∀ i : I, α ito- ∀ i : J, α imaps- P Ito- P Jfor all- J ⊆ I.
- MeasureTheory.IsProjectiveLimit:- μis the projective limit of the measure family- Pif for all- I : Finset ι, the map of- μby the projection to- Iis- P I.
Main statements #
- MeasureTheory.IsProjectiveLimit.unique: the projective limit of a family of finite measures is unique.
def
MeasureTheory.IsProjectiveMeasureFamily
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
(P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j))
 :
A family of measures indexed by finite sets of ι is projective if, for finite sets J ⊆ I,
the projection from ∀ i : I, α i to ∀ i : J, α i maps P I to P J.
Equations
- MeasureTheory.IsProjectiveMeasureFamily P = ∀ (I J : Finset ι) (hJI : J ⊆ I), P J = MeasureTheory.Measure.map (Finset.restrict₂ hJI) (P I)
Instances For
theorem
MeasureTheory.IsProjectiveMeasureFamily.eq_zero_of_isEmpty
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
[h : IsEmpty ((i : ι) → α i)]
(hP : IsProjectiveMeasureFamily P)
(I : Finset ι)
 :
theorem
MeasureTheory.IsProjectiveMeasureFamily.measure_univ_eq_of_subset
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
{I J : Finset ι}
(hP : IsProjectiveMeasureFamily P)
(hJI : J ⊆ I)
 :
Auxiliary lemma for measure_univ_eq.
theorem
MeasureTheory.IsProjectiveMeasureFamily.measure_univ_eq
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
(hP : IsProjectiveMeasureFamily P)
(I J : Finset ι)
 :
theorem
MeasureTheory.IsProjectiveMeasureFamily.congr_cylinder_of_subset
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
{I J : Finset ι}
(hP : IsProjectiveMeasureFamily P)
{S : Set ((i : ↥I) → α ↑i)}
{T : Set ((i : ↥J) → α ↑i)}
(hT : MeasurableSet T)
(h_eq : cylinder I S = cylinder J T)
(hJI : J ⊆ I)
 :
theorem
MeasureTheory.IsProjectiveMeasureFamily.congr_cylinder
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
{I J : Finset ι}
(hP : IsProjectiveMeasureFamily P)
{S : Set ((i : ↥I) → α ↑i)}
{T : Set ((i : ↥J) → α ↑i)}
(hS : MeasurableSet S)
(hT : MeasurableSet T)
(h_eq : cylinder I S = cylinder J T)
 :
def
MeasureTheory.IsProjectiveLimit
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
(μ : Measure ((i : ι) → α i))
(P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j))
 :
A measure μ is the projective limit of a family of measures indexed by finite sets of ι if
for all I : Finset ι, the projection from ∀ i, α i to ∀ i : I, α i maps μ to P I.
Equations
- MeasureTheory.IsProjectiveLimit μ P = ∀ (I : Finset ι), MeasureTheory.Measure.map I.restrict μ = P I
Instances For
theorem
MeasureTheory.IsProjectiveLimit.measure_cylinder
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
{μ : Measure ((i : ι) → α i)}
(h : IsProjectiveLimit μ P)
(I : Finset ι)
{s : Set ((i : ↥I) → α ↑i)}
(hs : MeasurableSet s)
 :
theorem
MeasureTheory.IsProjectiveLimit.measure_univ_eq
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
{μ : Measure ((i : ι) → α i)}
(hμ : IsProjectiveLimit μ P)
(I : Finset ι)
 :
theorem
MeasureTheory.IsProjectiveLimit.isFiniteMeasure
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
{μ : Measure ((i : ι) → α i)}
[∀ (i : Finset ι), IsFiniteMeasure (P i)]
(hμ : IsProjectiveLimit μ P)
 :
theorem
MeasureTheory.IsProjectiveLimit.isProbabilityMeasure
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
{μ : Measure ((i : ι) → α i)}
[∀ (i : Finset ι), IsProbabilityMeasure (P i)]
(hμ : IsProjectiveLimit μ P)
 :
theorem
MeasureTheory.IsProjectiveLimit.measure_univ_unique
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
{μ ν : Measure ((i : ι) → α i)}
(hμ : IsProjectiveLimit μ P)
(hν : IsProjectiveLimit ν P)
 :
theorem
MeasureTheory.IsProjectiveLimit.unique
{ι : Type u_1}
{α : ι → Type u_2}
[(i : ι) → MeasurableSpace (α i)]
{P : (J : Finset ι) → Measure ((j : ↥J) → α ↑j)}
{μ ν : Measure ((i : ι) → α i)}
[∀ (i : Finset ι), IsFiniteMeasure (P i)]
(hμ : IsProjectiveLimit μ P)
(hν : IsProjectiveLimit ν P)
 :
The projective limit of a family of finite measures is unique.