Riesz–Markov–Kakutani representation theorem #
This file prepares technical definitions and results for the Riesz-Markov-Kakutani representation
theorem on a locally compact T2 space X. As a special case, the statements about linear
functionals on bounded continuous functions follows. Actual theorems, depending on the
linearity (ℝ, ℝ≥0 or ℂ), are proven in separate files
(Mathlib/MeasureTheory/Integral/RieszMarkovKakutani/Real.lean,
Mathlib/MeasureTheory/Integral/RieszMarkovKakutani/NNReal.lean...)
To make use of the existing API, the measure is constructed from a content λ on the
compact subsets of a locally compact space X, rather than the usual construction of open sets in the
literature.
References #
- [Walter Rudin, Real and Complex Analysis.][Rud87]
Construction of the content: #
The positivity of a linear functional Λ implies that Λ is monotone.
Given a positive linear functional Λ on continuous compactly supported functions on X
with values in ℝ≥0, for K ⊆ X compact define λ(K) = inf {Λf | 1≤f on K}.
When X is a locally compact T2 space, this will be shown to be a
content, and will be shown to agree with the Riesz measure on the compact subsets K ⊆ X.
Equations
Instances For
For any compact subset K ⊆ X, there exist some compactly supported continuous nonnegative
functions f on X such that f ≥ 1 on K.
Riesz content λ (associated with a positive linear functional Λ) is
monotone: if K₁ ⊆ K₂ are compact subsets in X, then λ(K₁) ≤ λ(K₂).
Any compactly supported continuous nonnegative f such that f ≥ 1 on K gives an upper bound
on the content of K; namely λ(K) ≤ Λ f.
The Riesz content can be approximated arbitrarily well by evaluating the positive linear
functional on test functions: for any ε > 0, there exists a compactly supported continuous
nonnegative function f on X such that f ≥ 1 on K and such that λ(K) ≤ Λ f < λ(K) + ε.
The Riesz content λ associated to a given positive linear functional Λ is
finitely subadditive: λ(K₁ ∪ K₂) ≤ λ(K₁) + λ(K₂) for any compact subsets K₁, K₂ ⊆ X.
The content induced by the linear functional Λ.
Equations
- rieszContent Λ = { toFun := rieszContentAux Λ, mono' := ⋯, sup_disjoint' := ⋯, sup_le' := ⋯ }
Instances For
rieszContent gives a Content from Λ : C_c(X, ℝ≥0) →ₗ[ℝ≥0] ℝ≥0. Here rieszContent Λ is
promoted to a measure. It will be later shown that
∫ (x : X), f x ∂(rieszMeasure Λ hΛ) = Λ f for all f : C_c(X, ℝ≥0).