Skip to content

Commit

Permalink
Merge branch 'master' of github.com:RemyDegenne/CLT into master
Browse files Browse the repository at this point in the history
  • Loading branch information
RemyDegenne committed Dec 29, 2023
2 parents 1a9c1c4 + 1b6ecbb commit 78bc0c7
Show file tree
Hide file tree
Showing 5 changed files with 114 additions and 2 deletions.
1 change: 1 addition & 0 deletions Clt.lean
Original file line number Diff line number Diff line change
@@ -1,5 +1,6 @@
-- This module serves as the root of the `Clt` library.
-- Import modules here that should be built as part of the library.
import Clt.Tight
import Clt.Separating
import Clt.CharFun
import Clt.ExpPoly
35 changes: 35 additions & 0 deletions Clt/Separating.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,35 @@
/-
Copyright (c) 2023 Rémy Degenne. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Rémy Degenne
-/
import Mathlib.Topology.ContinuousFunction.Compact
import Mathlib.MeasureTheory.Measure.ProbabilityMeasure

/-!
# Separating sets of functions
-/

open BoundedContinuousFunction

open scoped ENNReal

namespace MeasureTheory

variable {α E : Type*} [MeasurableSpace α] [TopologicalSpace α] [NormedRing E] [NormedAlgebra ℝ E]

/-- A subalgebra `A` of the bounded continuous function from `α` to `E` is separating in the
probability measures if for all probability measures `μ ≠ ν`, there exists `f ∈ A` such that
`∫ x, f x ∂μ ≠ ∫ x, f x ∂ν`. -/
structure Separating (A : Subalgebra ℝ (α →ᵇ E)) : Prop :=
(exists_ne : ∀ ⦃μ ν : ProbabilityMeasure α⦄, μ ≠ ν → ∃ f ∈ A, ∫ x, f x ∂μ ≠ ∫ x, f x ∂ν)

lemma ProbabilityMeasure.ext_of_Separating {μ ν : ProbabilityMeasure α} {A : Subalgebra ℝ (α →ᵇ E)}
(hA : Separating A) (h : ∀ f ∈ A, ∫ x, f x ∂μ = ∫ x, f x ∂ν) :
μ = ν := by
by_contra h_ne
obtain ⟨f, hfA, hf_ne⟩ := hA.exists_ne h_ne
exact hf_ne (h f hfA)

end MeasureTheory
74 changes: 74 additions & 0 deletions Clt/Tight.lean
Original file line number Diff line number Diff line change
@@ -0,0 +1,74 @@
/-
Copyright (c) 2023 Rémy Degenne. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Rémy Degenne
-/
import Mathlib.MeasureTheory.Measure.Regular

/-!
# Characteristic function of a measure
## Main definitions
* `Tight`: A set `S` of measures is tight if for all `0 < ε`, there exists `K` compact such that
for all `μ` in `S`, `μ Kᶜ ≤ ε`.
## Main statements
* `fooBar_unique`
## Notation
## Implementation details
-/

open scoped ENNReal

namespace MeasureTheory

variable {α : Type*} {mα : MeasurableSpace α} {μ : Measure α}
[TopologicalSpace α]

/-- A set `S` of measures is tight if for all `0 < ε`, there exists `K` compact such that for all
`μ` in `S`, `μ Kᶜ ≤ ε`. -/
def Tight (S : Set (Measure α)) : Prop :=
∀ ε : ℝ≥0∞, 0 < ε → ∃ K : Set α, IsCompact K ∧ ∀ μ ∈ S, μ Kᶜ ≤ ε

-- TODO: the T2Space hypothesis is here only to make compact sets closed. It could be removed if
-- InnerRegular was about compact and closed sets, and not only compact sets.
lemma tight_singleton [T2Space α] [OpensMeasurableSpace α]
(μ : Measure α) [IsFiniteMeasure μ] [μ.InnerRegular] :
Tight {μ} := by
cases eq_zero_or_neZero μ with
| inl hμ =>
rw [hμ]
refine fun _ _ ↦ ⟨∅, isCompact_empty, ?_⟩
simp
| inr hμ =>
let r := μ Set.univ
have hr : 0 < r := by simp [hμ.out]
intro ε hε
cases lt_or_ge ε r with
| inl hεr =>
have hεr' : r - ε < r := ENNReal.sub_lt_self (measure_ne_top μ _) hr.ne' hε.ne'
obtain ⟨K, _, hK_compact, hKμ⟩ :=
(MeasurableSet.univ : MeasurableSet (Set.univ : Set α)).exists_lt_isCompact hεr'
refine ⟨K, hK_compact, fun μ' hμ' ↦ ?_⟩
simp only [Set.mem_singleton_iff] at hμ'
rw [hμ', measure_compl hK_compact.isClosed.measurableSet (measure_ne_top μ _),
tsub_le_iff_right]
rw [ENNReal.sub_lt_iff_lt_right (ne_top_of_lt hεr) hεr.le, add_comm] at hKμ
exact hKμ.le
| inr hεr =>
refine ⟨∅, isCompact_empty, ?_⟩
intro μ' hμ'
simp only [Set.mem_singleton_iff] at hμ'
rw [hμ', Set.compl_empty]
exact hεr

end MeasureTheory
3 changes: 2 additions & 1 deletion blueprint/src/chapter/separating.tex
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@ \chapter{Separating subalgebras}
\end{definition}

\begin{definition}\label{def:separating}
\lean{MeasureTheory.Separating} \leanok
A set $\mathcal F$ of functions $E \to F$ is separating in $\mathcal P(E)$ if for all probability measures $\mu, \nu$ on $E$ with $\mu \ne \nu$, there exists $f \in \mathcal F$ with $\mu[f] \ne \nu[f]$.
\end{definition}

Expand Down Expand Up @@ -93,6 +94,7 @@ \chapter{Separating subalgebras}
\end{proof}

\begin{lemma}\label{lem:Cb_eq_of_separating}
\uses{def:separates_points}
Let $\mathcal A$ be a subalgebra of $C_b(E, \mathbb{C})$ that separates points. Let $\mu, \nu$ be two probability measures. If for all $f \in \mathcal A$, $\mu[f] = \nu[f]$, then for all $g \in C_b(E, \mathbb{C})$, $\mu[g] = \nu[g]$.
\end{lemma}

Expand All @@ -115,7 +117,6 @@ \chapter{Separating subalgebras}
\end{proof}

\begin{theorem}[Subalgebras separating points]\label{thm:separating_starSubalgebra}
\uses{def:separates_points}
Let $E$ be a complete separable pseudo-metric space. Let $\mathcal A \subseteq C_b(E, \mathbb{C})$ be a star-subalgebra that separates points. Then $\mathcal A$ is separating in $\mathcal P(E)$.
\end{theorem}

Expand Down
3 changes: 2 additions & 1 deletion blueprint/src/chapter/tight.tex
Original file line number Diff line number Diff line change
@@ -1,7 +1,8 @@
\chapter{Tight families of measures}

\begin{definition}\label{def:tight}
A set $S$ of measures on $\Omega$ is tight if for all $\varepsilon > 0$ there exists a compact set $K$ such that for all $\mu \in S$, $\mu(\Omega \setminus K) \le \varepsilon$.
\lean{MeasureTheory.Tight} \leanok
A set $S$ of measures on $\Omega$ is tight if for all $\varepsilon > 0$ there exists a compact set $K$ such that for all $\mu \in S$, $\mu(K^c) \le \varepsilon$.
\end{definition}

\begin{lemma}\label{lem:tight_of_cvg}
Expand Down

0 comments on commit 78bc0c7

Please sign in to comment.