Additive Haar measure constructed from a basis #
Given a basis of a finite-dimensional real vector space, we define the corresponding Lebesgue
measure, which gives measure 1 to the parallelepiped spanned by the basis.
Main definitions #
- parallelepiped vis the parallelepiped spanned by a finite family of vectors.
- Basis.parallelepipedis the parallelepiped associated to a basis, seen as a compact set with nonempty interior.
- Basis.addHaaris the Lebesgue measure associated to a basis, giving measure- 1to the corresponding parallelepiped.
In particular, we declare a MeasureSpace instance on any finite-dimensional inner product space,
by using the Lebesgue measure associated to some orthonormal basis (which is in fact independent
of the basis).
The closed parallelepiped spanned by a finite family of vectors.
Instances For
Reindexing a family of vectors does not change their parallelepiped.
A parallelepiped is the convex hull of its vertices
The axis aligned parallelepiped over ι → ℝ is a cuboid.
The parallelepiped spanned by a basis, as a compact set with nonempty interior.
Equations
- b.parallelepiped = { carrier := parallelepiped ⇑b, isCompact' := ⋯, interior_nonempty' := ⋯ }
Instances For
The Lebesgue measure associated to a basis, giving measure 1 to the parallelepiped spanned
by the basis.
Equations
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Let μ be a σ-finite left invariant measure on E. Then μ is equal to the Haar measure
defined by b iff the parallelepiped defined by b has measure 1 for μ.
A finite dimensional inner product space has a canonical measure, the Lebesgue measure giving
volume 1 to the parallelepiped spanned by any orthonormal basis. We define the measure using
some arbitrary choice of orthonormal basis. The fact that it works with any orthonormal basis
is proved in orthonormalBasis.volume_parallelepiped.
Equations
- measureSpaceOfInnerProductSpace = MeasureTheory.MeasureSpace.mk (stdOrthonormalBasis ℝ E).toBasis.addHaar
Equations
- ⋯ = ⋯
Equations
- Real.measureSpace = inferInstance
Miscellaneous instances for EuclideanSpace #
In combination with measureSpaceOfInnerProductSpace, these put a MeasureSpace structure
on EuclideanSpace.
Equations
- EuclideanSpace.instMeasurableSpaceReal ι = MeasurableSpace.pi
Equations
- ⋯ = ⋯
WithLp.equiv as a MeasurableEquiv.
Equations
- EuclideanSpace.measurableEquiv ι = { toEquiv := WithLp.equiv 2 ((i : ι) → (fun (x : ι) => ℝ) i), measurable_toFun := ⋯, measurable_invFun := ⋯ }