Skip to content

Commit

Permalink
feat: remove the diamond for Complex.measureSpace (#6832)
Browse files Browse the repository at this point in the history
We remove the instance 
```lean
instance measureSpace : MeasureSpace ℂ :=  ⟨Measure.map basisOneI.equivFun.symm volume⟩
```
in `MeasureTheory.Measure.Lebesgue.Complex` since `ℂ` has already a `measureSpace` instance coming from its `InnerProductSpace` structure over `ℝ`, and fix the proof of `volume_preserving_equiv_pi`.


Co-authored-by: Eric Wieser <[email protected]>
  • Loading branch information
xroblot and eric-wieser committed Sep 16, 2023
1 parent d703031 commit 67b54e2
Showing 1 changed file with 14 additions and 14 deletions.
28 changes: 14 additions & 14 deletions Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean
Original file line number Diff line number Diff line change
Expand Up @@ -4,19 +4,19 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Yury Kudryashov
-/
import Mathlib.MeasureTheory.Constructions.BorelSpace.Complex
import Mathlib.MeasureTheory.Measure.Lebesgue.Basic
import Mathlib.MeasureTheory.Measure.Haar.OfBasis
import Mathlib.MeasureTheory.Measure.Haar.InnerProductSpace

#align_import measure_theory.measure.lebesgue.complex from "leanprover-community/mathlib"@"fd5edc43dc4f10b85abfe544b88f82cf13c5f844"

/-!
# Lebesgue measure on `ℂ`
In this file we define Lebesgue measure on `ℂ`. Since `ℂ` is defined as a `structure` as the
push-forward of the volume on `ℝ²` under the natural isomorphism. There are (at least) two
frequently used ways to represent `ℝ²` in `mathlib`: `ℝ × ℝ` and `Fin 2 → ℝ`. We define measurable
equivalences (`MeasurableEquiv`) to both types and prove that both of them are volume preserving
(in the sense of `MeasureTheory.measurePreserving`).
In this file, we consider the Lebesgue measure on `ℂ` defined as the push-forward of the volume
on `ℝ²` under the natural isomorphism and prove that it is equal to the measure `volume` of `ℂ`
coming from its `InnerProductSpace` structure over `ℝ`. For that, we consider the two frequently
used ways to represent `ℝ²` in `mathlib`: `ℝ × ℝ` and `Fin 2 → ℝ`, define measurable equivalences
(`MeasurableEquiv`) to both types and prove that both of them are volume preserving (in the sense
of `MeasureTheory.measurePreserving`).
-/


Expand All @@ -26,11 +26,6 @@ noncomputable section

namespace Complex

/-- Lebesgue measure on `ℂ`. -/
instance measureSpace : MeasureSpace ℂ :=
⟨Measure.map basisOneI.equivFun.symm volume⟩
#align complex.measure_space Complex.measureSpace

/-- Measurable equivalence between `ℂ` and `ℝ² = Fin 2 → ℝ`. -/
def measurableEquivPi : ℂ ≃ᵐ (Fin 2 → ℝ) :=
basisOneI.equivFun.toContinuousLinearEquiv.toHomeomorph.toMeasurableEquiv
Expand All @@ -41,8 +36,13 @@ def measurableEquivRealProd : ℂ ≃ᵐ ℝ × ℝ :=
equivRealProdClm.toHomeomorph.toMeasurableEquiv
#align complex.measurable_equiv_real_prod Complex.measurableEquivRealProd

theorem volume_preserving_equiv_pi : MeasurePreserving measurableEquivPi :=
(measurableEquivPi.symm.measurable.measurePreserving _).symm _
theorem volume_preserving_equiv_pi : MeasurePreserving measurableEquivPi := by
convert (measurableEquivPi.symm.measurable.measurePreserving volume).symm
rw [← addHaarMeasure_eq_volume_pi, ← Basis.parallelepiped_basisFun, ← Basis.addHaar,
measurableEquivPi, Homeomorph.toMeasurableEquiv_symm_coe,
ContinuousLinearEquiv.symm_toHomeomorph, ContinuousLinearEquiv.coe_toHomeomorph,
Basis.map_addHaar, eq_comm]
exact (Basis.addHaar_eq_iff _ _).mpr Complex.orthonormalBasisOneI.volume_parallelepiped
#align complex.volume_preserving_equiv_pi Complex.volume_preserving_equiv_pi

theorem volume_preserving_equiv_real_prod : MeasurePreserving measurableEquivRealProd :=
Expand Down

0 comments on commit 67b54e2

Please sign in to comment.