From aa348f4f72de085754c91257a8ec91abb5fb5019 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Fri, 18 Oct 2024 15:52:14 +0200 Subject: [PATCH 1/4] 1st commit --- Mathlib.lean | 5 + Mathlib/Algebra/Group/Indicator.lean | 5 + .../GroupWithZero/NonZeroDivisors.lean | 6 + Mathlib/Algebra/Module/ZLattice/Basic.lean | 114 +- Mathlib/Algebra/Module/ZLattice/Covolume.lean | 254 ++- .../Asymptotics/AsymptoticEquivalent.lean | 3 + Mathlib/Analysis/BoxIntegral/Box/Basic.lean | 7 + .../Analysis/BoxIntegral/UnitPartition.lean | 477 +++++ Mathlib/Analysis/Complex/Basic.lean | 9 +- .../Analysis/SpecialFunctions/PolarCoord.lean | 90 +- .../Analysis/SpecialFunctions/Pow/Real.lean | 4 + Mathlib/Data/Complex/Order.lean | 26 + Mathlib/Data/ENNReal/Real.lean | 3 + Mathlib/LinearAlgebra/Basis/Basic.lean | 4 + Mathlib/LinearAlgebra/Determinant.lean | 14 + .../Constructions/BorelSpace/Basic.lean | 4 + Mathlib/MeasureTheory/Constructions/Pi.lean | 58 +- .../Constructions/Prod/Basic.lean | 32 + Mathlib/MeasureTheory/Integral/Marginal.lean | 16 + Mathlib/MeasureTheory/Integral/Pi.lean | 70 +- .../MeasurableSpace/Embedding.lean | 26 + Mathlib/NumberTheory/LSeries/Basic.lean | 5 + Mathlib/NumberTheory/LSeries/Residue.lean | 330 ++++ Mathlib/NumberTheory/LSeries/RiemannZeta.lean | 12 + .../NumberField/CanonicalEmbedding/Basic.lean | 378 +++- .../CanonicalEmbedding/ConvexBody.lean | 11 +- .../CanonicalEmbedding/FundamentalCone.lean | 219 ++- .../CanonicalEmbedding/NormLessThanOne.lean | 1612 +++++++++++++++++ .../NumberTheory/NumberField/ClassNumber.lean | 4 + .../NumberField/DedekindZeta.lean | 70 + .../NumberField/Discriminant.lean | 19 + .../NumberTheory/NumberField/Embeddings.lean | 5 +- Mathlib/NumberTheory/NumberField/Ideal.lean | 139 ++ .../NumberTheory/NumberField/Units/Basic.lean | 8 + .../NumberField/Units/DirichletTheorem.lean | 63 +- .../NumberField/Units/Regulator.lean | 50 + Mathlib/RingTheory/Ideal/Norm.lean | 74 +- .../Algebra/InfiniteSum/Constructions.lean | 11 + Mathlib/Topology/Algebra/Module/Basic.lean | 50 + Mathlib/Topology/Basic.lean | 19 + Mathlib/Topology/ContinuousOn.lean | 12 + Mathlib/Topology/PartialHomeomorph.lean | 2 +- 42 files changed, 4182 insertions(+), 138 deletions(-) create mode 100644 Mathlib/Analysis/BoxIntegral/UnitPartition.lean create mode 100644 Mathlib/NumberTheory/LSeries/Residue.lean create mode 100644 Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean create mode 100644 Mathlib/NumberTheory/NumberField/DedekindZeta.lean create mode 100644 Mathlib/NumberTheory/NumberField/Ideal.lean diff --git a/Mathlib.lean b/Mathlib.lean index 14f4a5f1fa41b..a0f914f6d98e3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -964,6 +964,7 @@ import Mathlib.Analysis.BoxIntegral.Partition.Measure import Mathlib.Analysis.BoxIntegral.Partition.Split import Mathlib.Analysis.BoxIntegral.Partition.SubboxInduction import Mathlib.Analysis.BoxIntegral.Partition.Tagged +import Mathlib.Analysis.BoxIntegral.UnitPartition import Mathlib.Analysis.CStarAlgebra.Basic import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Basic import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances @@ -3506,6 +3507,7 @@ import Mathlib.NumberTheory.LSeries.HurwitzZetaOdd import Mathlib.NumberTheory.LSeries.HurwitzZetaValues import Mathlib.NumberTheory.LSeries.Linearity import Mathlib.NumberTheory.LSeries.MellinEqDirichlet +import Mathlib.NumberTheory.LSeries.Residue import Mathlib.NumberTheory.LSeries.RiemannZeta import Mathlib.NumberTheory.LSeries.ZMod import Mathlib.NumberTheory.LegendreSymbol.AddCharacter @@ -3546,12 +3548,15 @@ import Mathlib.NumberTheory.NumberField.Basic import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.Basic import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.ConvexBody import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.FundamentalCone +import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.NormLessThanOne import Mathlib.NumberTheory.NumberField.ClassNumber +import Mathlib.NumberTheory.NumberField.DedekindZeta import Mathlib.NumberTheory.NumberField.Discriminant import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.NumberTheory.NumberField.EquivReindex import Mathlib.NumberTheory.NumberField.FractionalIdeal import Mathlib.NumberTheory.NumberField.House +import Mathlib.NumberTheory.NumberField.Ideal import Mathlib.NumberTheory.NumberField.Norm import Mathlib.NumberTheory.NumberField.Units.Basic import Mathlib.NumberTheory.NumberField.Units.DirichletTheorem diff --git a/Mathlib/Algebra/Group/Indicator.lean b/Mathlib/Algebra/Group/Indicator.lean index fc2d14cdfb140..0563e3366ce74 100644 --- a/Mathlib/Algebra/Group/Indicator.lean +++ b/Mathlib/Algebra/Group/Indicator.lean @@ -122,6 +122,11 @@ theorem mem_of_mulIndicator_ne_one (h : mulIndicator s f a ≠ 1) : a ∈ s := @[to_additive] theorem eqOn_mulIndicator : EqOn (mulIndicator s f) f s := fun _ hx => mulIndicator_of_mem hx f +/-- Docstring. -/ +@[to_additive] +theorem eqOn_mulIndicator' : EqOn (mulIndicator s f) 1 sᶜ := + fun _ hx => mulIndicator_of_not_mem hx f + @[to_additive] theorem mulSupport_mulIndicator_subset : mulSupport (s.mulIndicator f) ⊆ s := fun _ hx => hx.imp_symm fun h => mulIndicator_of_not_mem h f diff --git a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean index f28e2198589e2..945a8d385b93b 100644 --- a/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean +++ b/Mathlib/Algebra/GroupWithZero/NonZeroDivisors.lean @@ -180,6 +180,12 @@ theorem mul_mem_nonZeroDivisors {a b : M₁} : a * b ∈ M₁⁰ ↔ a ∈ M₁ apply hb rw [mul_assoc, hx] +theorem nonZeroDivisors_dvd_iff_dvd_coe {a b : M₁⁰} : + a ∣ b ↔ (a : M₁) ∣ (b : M₁) := + ⟨fun ⟨c, hc⟩ ↦ by simp_rw [hc, Submonoid.coe_mul, dvd_mul_right], + fun ⟨c, hc⟩ ↦ ⟨⟨c, (mul_mem_nonZeroDivisors.mp (hc ▸ b.prop)).2⟩, + by simp_rw [Subtype.ext_iff, Submonoid.coe_mul, hc]⟩⟩ + theorem isUnit_of_mem_nonZeroDivisors {G₀ : Type*} [GroupWithZero G₀] {x : G₀} (hx : x ∈ nonZeroDivisors G₀) : IsUnit x := ⟨⟨x, x⁻¹, mul_inv_cancel₀ (nonZeroDivisors.ne_zero hx), diff --git a/Mathlib/Algebra/Module/ZLattice/Basic.lean b/Mathlib/Algebra/Module/ZLattice/Basic.lean index f5d850e0a404e..74da690f908cd 100644 --- a/Mathlib/Algebra/Module/ZLattice/Basic.lean +++ b/Mathlib/Algebra/Module/ZLattice/Basic.lean @@ -58,6 +58,17 @@ variable (b : Basis ι K E) theorem span_top : span K (span ℤ (Set.range b) : Set E) = ⊤ := by simp [span_span_of_tower] +theorem map {F : Type*} [NormedAddCommGroup F] [NormedSpace K F] (f : E ≃ₗ[K] F) : + Submodule.map (f.restrictScalars ℤ) (span ℤ (Set.range b)) = span ℤ (Set.range (b.map f)) := by + simp_rw [Submodule.map_span, LinearEquiv.restrictScalars_apply, Basis.coe_map, Set.range_comp] + +open scoped Pointwise in +theorem smul {c : K} (hc : c ≠ 0) : + c • span ℤ (Set.range b) = span ℤ (Set.range (b.isUnitSMul (fun _ ↦ Ne.isUnit hc))) := by + rw [smul_span, Set.smul_set_range] + congr! + rw [Basis.isUnitSMul_apply] + /-- The fundamental domain of the ℤ-lattice spanned by `b`. See `ZSpan.isAddFundamentalDomain` for the proof that it is a fundamental domain. -/ def fundamentalDomain : Set E := {m | ∀ i, b.repr m i ∈ Set.Ico (0 : K) 1} @@ -297,9 +308,16 @@ instance [Finite ι] : DiscreteTopology (span ℤ (Set.range b)) := by · exact discreteTopology_pi_basisFun · refine Subtype.map_injective _ (Basis.equivFun b).injective -instance [Finite ι] : DiscreteTopology (span ℤ (Set.range b)).toAddSubgroup := +instance instDiscreteTopology [Finite ι] : DiscreteTopology (span ℤ (Set.range b)).toAddSubgroup := inferInstanceAs <| DiscreteTopology (span ℤ (Set.range b)) +theorem setFinite_inter [ProperSpace E] [Finite ι] {s : Set E} (hs : Bornology.IsBounded s) : + Set.Finite (s ∩ (span ℤ (Set.range b))) := by + have : DiscreteTopology (span ℤ (Set.range b)) := by infer_instance + refine Metric.finite_isBounded_inter_isClosed hs ?_ + change IsClosed (span ℤ (Set.range b)).toAddSubgroup + exact inferInstance + @[measurability] theorem fundamentalDomain_measurableSet [MeasurableSpace E] [OpensMeasurableSpace E] [Finite ι] : MeasurableSet (fundamentalDomain b) := by @@ -399,6 +417,8 @@ theorem _root_.ZSpan.isZLattice {E ι : Type*} [NormedAddCommGroup E] [NormedSpa IsZLattice ℝ (span ℤ (Set.range b)) where span_top := ZSpan.span_top b +section NormedLinearOrderedField + variable (K : Type*) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] variable [ProperSpace E] (L : Submodule ℤ E) [DiscreteTopology L] @@ -607,4 +627,96 @@ instance instCountable_of_discrete_submodule {E : Type*} [NormedAddCommGroup E] simp_rw [← (Module.Free.chooseBasis ℤ L).ofZLatticeBasis_span ℝ] infer_instance +end NormedLinearOrderedField + +section comap + +variable (K : Type*) [NormedField K] {E F : Type*} [NormedAddCommGroup E] [NormedSpace K E] + [NormedAddCommGroup F] [NormedSpace K F] (L : Submodule ℤ E) + +/--- The coimage of a `ZLattice` by a linear map. -/ +protected def ZLattice.comap (e : F →ₗ[K] E) := L.comap (e.restrictScalars ℤ) + +@[simp] +theorem ZLattice.coe_comap (e : F →ₗ[K] E) : + (ZLattice.comap K L e : Set F) = e⁻¹' L := rfl + +theorem ZLattice.comap_refl : + ZLattice.comap K L (1 : E →ₗ[K] E)= L := Submodule.comap_id L + +theorem ZLattice.comap_discreteTopology [hL : DiscreteTopology L] {e : F →ₗ[K] E} + (he₁ : Continuous e) (he₂ : Function.Injective e) : + DiscreteTopology (ZLattice.comap K L e) := by + exact DiscreteTopology.preimage_of_continuous_injective L he₁ he₂ + +instance [DiscreteTopology L] (e : F ≃L[K] E) : + DiscreteTopology (ZLattice.comap K L e.toLinearMap) := + ZLattice.comap_discreteTopology K L e.continuous e.injective + +theorem ZLattice.comap_span_top (hL : span K (L : Set E) = ⊤) {e : F →ₗ[K] E} + (he : (L : Set E) ⊆ LinearMap.range e) : + span K (ZLattice.comap K L e : Set F) = ⊤ := by + rw [ZLattice.coe_comap, Submodule.span_preimage_eq (Submodule.nonempty L) he, hL, comap_top] + +instance [DiscreteTopology L] [IsZLattice K L] (e : F ≃L[K] E) : + IsZLattice K (ZLattice.comap K L e.toLinearMap) where + span_top := by + rw [ZLattice.coe_comap, LinearEquiv.coe_coe, e.coe_toLinearEquiv, ← e.image_symm_eq_preimage, + ← Submodule.map_span, IsZLattice.span_top, Submodule.map_top, LinearEquivClass.range] + +theorem ZLattice.comap_comp {G : Type*} [NormedAddCommGroup G] [NormedSpace K G] + (e : F →ₗ[K] E) (e' : G →ₗ[K] F) : + (ZLattice.comap K (ZLattice.comap K L e) e') = ZLattice.comap K L (e ∘ₗ e') := + (Submodule.comap_comp _ _ L).symm + +/-- If `e` is a linear equivalence, it induces a `ℤ`-linear equivalence between +`L` and `ZLattice.comap K L e`. -/ +def ZLattice.comap_equiv (e : F ≃ₗ[K] E) : + L ≃ₗ[ℤ] (ZLattice.comap K L e.toLinearMap) := + LinearEquiv.ofBijective + ((e.symm.toLinearMap.restrictScalars ℤ).restrict + (fun _ h ↦ by simpa [← SetLike.mem_coe] using h)) + ⟨fun _ _ h ↦ Subtype.ext_iff_val.mpr (e.symm.injective (congr_arg Subtype.val h)), + fun ⟨x, hx⟩ ↦ ⟨⟨e x, by rwa [← SetLike.mem_coe, ZLattice.coe_comap] at hx⟩, + by simp [Subtype.ext_iff_val]⟩⟩ + +@[simp] +theorem ZLattice.comap_equiv_apply (e : F ≃ₗ[K] E) (x : L) : + ZLattice.comap_equiv K L e x = e.symm x := rfl + +/-- The basis of `ZLattice.comap K L e` given by the image of a basis `b` of `L` by `e.symm`. -/ +def Basis.ofZLatticeComap (e : F ≃ₗ[K] E) {ι : Type*} + (b : Basis ι ℤ L) : + Basis ι ℤ (ZLattice.comap K L e.toLinearMap) := b.map (ZLattice.comap_equiv K L e) + +@[simp] +theorem Basis.ofZLatticeComap_apply (e : F ≃ₗ[K] E) {ι : Type*} + (b : Basis ι ℤ L) (i : ι) : + b.ofZLatticeComap K L e i = e.symm (b i) := by simp [Basis.ofZLatticeComap] + +@[simp] +theorem Basis.ofZLatticeComap_repr_apply (e : F ≃ₗ[K] E) {ι : Type*} (b : Basis ι ℤ L) (x : L) + (i : ι) : + (b.ofZLatticeComap K L e).repr (ZLattice.comap_equiv K L e x) i = b.repr x i := by + simp [Basis.ofZLatticeComap] + +end comap + +section NormedLinearOrderedField_comap + +variable (K : Type*) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K] +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] + [ProperSpace E] +variable {F : Type*} [NormedAddCommGroup F] [NormedSpace K F] [FiniteDimensional K F] + [ProperSpace F] +variable (L : Submodule ℤ E) [DiscreteTopology L] [IsZLattice K L] + +theorem Basis.ofZLatticeBasis_comap (e : F ≃L[K] E) {ι : Type*} (b : Basis ι ℤ L) : + (b.ofZLatticeComap K L e.toLinearEquiv).ofZLatticeBasis K (ZLattice.comap K L e.toLinearMap) = + (b.ofZLatticeBasis K L).map e.symm.toLinearEquiv := by + ext + simp + +end NormedLinearOrderedField_comap + end ZLattice diff --git a/Mathlib/Algebra/Module/ZLattice/Covolume.lean b/Mathlib/Algebra/Module/ZLattice/Covolume.lean index 04e559d80a384..cb2e72e866c1a 100644 --- a/Mathlib/Algebra/Module/ZLattice/Covolume.lean +++ b/Mathlib/Algebra/Module/ZLattice/Covolume.lean @@ -4,11 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ import Mathlib.Algebra.Module.ZLattice.Basic +import Mathlib.Analysis.BoxIntegral.UnitPartition +import Mathlib.MeasureTheory.Measure.Haar.InnerProductSpace +import Mathlib.MeasureTheory.Measure.Haar.OfBasis /-! # Covolume of ℤ-lattices -Let `E` be a finite dimensional real vector space with an inner product. +Let `E` be a finite dimensional real vector space. Let `L` be a `ℤ`-lattice `L` defined as a discrete `ℤ`-submodule of `E` that spans `E` over `ℝ`. @@ -23,13 +26,32 @@ choice of the fundamental domain of `L`. * `ZLattice.covolume_eq_det`: if `L` is a lattice in `ℝ^n`, then its covolume is the absolute value of the determinant of any `ℤ`-basis of `L`. +* `Zlattice.covolume.tendsto_card_div_pow`: Let `s` be a bounded measurable set of `ι → ℝ`, then +the number of points in `s ∩ n⁻¹ • L` divided by `n ^ card ι` tends to `volume s / covolume L` +when `n : ℕ` tends to infinity. See also `Zlattice.covolume.tendsto_card_div_pow'` for a version +for `InnerProductSpace ℝ E` and `Zlattice.covolume.tendsto_card_div_pow''` for the general version. + +* `Zlattice.covolume.tendsto_card_le_div`: Let `X` be a cone in `ι → ℝ` and let `F : (ι → ℝ) → ℝ` +be a function such that `F (c • x) = c ^ card ι * F x`. Then the number of points `x ∈ X` such that +`F x ≤ c` divided by `c` tends to `volume {x ∈ X | F x ≤ 1} / covolume L` when `c : ℝ` tends to +infinity. See also `Zlattice.covolume.tendsto_card_le_div'` for a version for +`InnerProductSpace ℝ E` and `Zlattice.covolume.tendsto_card_le_div''` for the general version. + +## Naming conventions + +Many of the same results are true in the pi case `E` is `ι → ℝ` and in the case `E` is an +`InnerProductSpace`. We use the following convention: the plain name is for the pi case, for eg. +`volume_image_eq_volume_div_covolume`. For the same result in the `InnerProductSpace` case, we add +a `prime`, for eg. `volume_image_eq_volume_div_covolume'`. When the same result exists in the +general case, we had two primes, eg. `Zlattice.covolume.tendsto_card_div_pow''`. + -/ noncomputable section namespace ZLattice -open Submodule MeasureTheory Module MeasureTheory Module +open Submodule MeasureTheory Module MeasureTheory Module ZSpan section General @@ -45,7 +67,7 @@ def covolume (μ : Measure E := by volume_tac) : ℝ := (addCovolume L E μ).toR end General -section Real +section Basic variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] variable [MeasurableSpace E] [BorelSpace E] @@ -61,18 +83,30 @@ theorem covolume_eq_measure_fundamentalDomain {F : Set E} (h : IsAddFundamentalD theorem covolume_ne_zero : covolume L μ ≠ 0 := by rw [covolume_eq_measure_fundamentalDomain L μ (isAddFundamentalDomain (Free.chooseBasis ℤ L) μ), ENNReal.toReal_ne_zero] - refine ⟨ZSpan.measure_fundamentalDomain_ne_zero _, ne_of_lt ?_⟩ - exact Bornology.IsBounded.measure_lt_top (ZSpan.fundamentalDomain_isBounded _) + refine ⟨measure_fundamentalDomain_ne_zero _, ne_of_lt ?_⟩ + exact Bornology.IsBounded.measure_lt_top (fundamentalDomain_isBounded _) theorem covolume_pos : 0 < covolume L μ := lt_of_le_of_ne ENNReal.toReal_nonneg (covolume_ne_zero L μ).symm +theorem covolume_comap {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] [FiniteDimensional ℝ F] + [MeasurableSpace F] [BorelSpace F] (ν : Measure F := by volume_tac) [Measure.IsAddHaarMeasure ν] + {e : F ≃L[ℝ] E} (he : MeasurePreserving e ν μ) : + covolume (ZLattice.comap ℝ L e.toLinearMap) ν = covolume L μ := by + rw [covolume_eq_measure_fundamentalDomain _ _ (isAddFundamentalDomain (Free.chooseBasis ℤ L) μ), + covolume_eq_measure_fundamentalDomain _ _ ((isAddFundamentalDomain + ((Free.chooseBasis ℤ L).ofZLatticeComap ℝ L e.toLinearEquiv) ν)), ← he.measure_preimage + (fundamentalDomain_measurableSet _).nullMeasurableSet, ← e.image_symm_eq_preimage, + ← e.symm.coe_toLinearEquiv, map_fundamentalDomain] + congr! + ext; simp + theorem covolume_eq_det_mul_measure {ι : Type*} [Fintype ι] [DecidableEq ι] (b : Basis ι ℤ L) (b₀ : Basis ι ℝ E) : - covolume L μ = |b₀.det ((↑) ∘ b)| * (μ (ZSpan.fundamentalDomain b₀)).toReal := by + covolume L μ = |b₀.det ((↑) ∘ b)| * (μ (fundamentalDomain b₀)).toReal := by rw [covolume_eq_measure_fundamentalDomain L μ (isAddFundamentalDomain b μ), - ZSpan.measure_fundamentalDomain _ _ b₀, - measure_congr (ZSpan.fundamentalDomain_ae_parallelepiped b₀ μ), ENNReal.toReal_mul, + measure_fundamentalDomain _ _ b₀, + measure_congr (fundamentalDomain_ae_parallelepiped b₀ μ), ENNReal.toReal_mul, ENNReal.toReal_ofReal (by positivity)] congr ext @@ -82,11 +116,211 @@ theorem covolume_eq_det {ι : Type*} [Fintype ι] [DecidableEq ι] (L : Submodul [DiscreteTopology L] [IsZLattice ℝ L] (b : Basis ι ℤ L) : covolume L = |(Matrix.of ((↑) ∘ b)).det| := by rw [covolume_eq_measure_fundamentalDomain L volume (isAddFundamentalDomain b volume), - ZSpan.volume_fundamentalDomain, ENNReal.toReal_ofReal (by positivity)] + volume_fundamentalDomain, ENNReal.toReal_ofReal (by positivity)] congr ext1 exact b.ofZLatticeBasis_apply ℝ L _ -end Real +theorem covolume_eq_det_inv {ι : Type*} [Fintype ι] [DecidableEq ι] (L : Submodule ℤ (ι → ℝ)) + [DiscreteTopology L] [IsZLattice ℝ L] (b : Basis ι ℤ L) : + covolume L = |(LinearEquiv.det (b.ofZLatticeBasis ℝ L).equivFun : ℝ)|⁻¹ := by + rw [covolume_eq_det L b, ← Pi.basisFun_det_apply, show (((↑) : L → _) ∘ ⇑b) = + (b.ofZLatticeBasis ℝ) by ext; simp, ← Basis.det_inv, ← abs_inv, Units.val_inv_eq_inv_val, + IsUnit.unit_spec, Basis.det_basis, LinearEquiv.coe_det] + rfl + +theorem volume_image_eq_volume_div_covolume {ι : Type*} [Fintype ι] [DecidableEq ι] + (L : Submodule ℤ (ι → ℝ)) [DiscreteTopology L] [IsZLattice ℝ L] (b : Basis ι ℤ L) + {s : Set (ι → ℝ)} : + volume ((b.ofZLatticeBasis ℝ L).equivFun '' s) = volume s / ENNReal.ofReal (covolume L) := by + rw [LinearEquiv.image_eq_preimage, Measure.addHaar_preimage_linearEquiv, LinearEquiv.symm_symm, + covolume_eq_det_inv L b, ENNReal.div_eq_inv_mul, ENNReal.ofReal_inv_of_pos + (abs_pos.mpr (LinearEquiv.det _).ne_zero), inv_inv, LinearEquiv.coe_det] + +/-- Docstring. -/ +theorem volume_image_eq_volume_div_covolume' {E : Type*} [NormedAddCommGroup E] + [InnerProductSpace ℝ E] [FiniteDimensional ℝ E] [MeasurableSpace E] [BorelSpace E] + (L : Submodule ℤ E) [DiscreteTopology L] [IsZLattice ℝ L] {ι : Type*} [Fintype ι] + (b : Basis ι ℤ L) {s : Set E} (hs : NullMeasurableSet s) : + volume ((b.ofZLatticeBasis ℝ).equivFun '' s) = volume s / ENNReal.ofReal (covolume L) := by + classical + let e : Fin (finrank ℝ E) ≃ ι := + Fintype.equivOfCardEq (by rw [Fintype.card_fin, finrank_eq_card_basis (b.ofZLatticeBasis ℝ)]) + let f := (EuclideanSpace.equiv ι ℝ).symm.trans + ((stdOrthonormalBasis ℝ E).reindex e).repr.toContinuousLinearEquiv.symm + have hf : MeasurePreserving f := + ((stdOrthonormalBasis ℝ E).reindex e).measurePreserving_repr_symm.comp + (EuclideanSpace.volume_preserving_measurableEquiv ι).symm + rw [← hf.measure_preimage hs, ← (covolume_comap L volume volume hf), + ← volume_image_eq_volume_div_covolume (ZLattice.comap ℝ L f.toLinearMap) + (b.ofZLatticeComap ℝ L f.toLinearEquiv), Basis.ofZLatticeBasis_comap, + ← f.image_symm_eq_preimage, ← Set.image_comp] + simp + +end Basic + +namespace covolume + +section General + +open Filter Fintype Pointwise Topology BoxIntegral Bornology + +variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] +variable {L : Submodule ℤ E} [DiscreteTopology L] [IsZLattice ℝ L] +variable {ι : Type*} [Fintype ι] (b : Basis ι ℤ L) + +/-- Docstring. -/ +theorem tendsto_card_div_pow'' [FiniteDimensional ℝ E] [MeasurableSpace E] [BorelSpace E] + {s : Set E} (hs₁ : IsBounded s) (hs₂ : MeasurableSet s) + (hs₃ : volume (frontier ((b.ofZLatticeBasis ℝ).equivFun '' s)) = 0): + Tendsto (fun n : ℕ ↦ (Nat.card (s ∩ (n : ℝ)⁻¹ • L : Set E) : ℝ) / n ^ card ι) + atTop (𝓝 (volume ((b.ofZLatticeBasis ℝ).equivFun '' s)).toReal) := by + refine Tendsto.congr' ?_ + (unitPartition.tendsto_card_div_pow' ((b.ofZLatticeBasis ℝ).equivFun '' s) ?_ ?_ hs₃) + · filter_upwards [eventually_gt_atTop 0] with n hn + congr + refine Nat.card_congr <| ((b.ofZLatticeBasis ℝ).equivFun.toEquiv.subtypeEquiv fun x ↦ ?_).symm + simp_rw [Set.mem_inter_iff, ← b.ofZLatticeBasis_span ℝ, LinearEquiv.coe_toEquiv, + Basis.equivFun_apply, Set.mem_image, DFunLike.coe_fn_eq, EmbeddingLike.apply_eq_iff_eq, + exists_eq_right, and_congr_right_iff, Set.mem_inv_smul_set_iff₀ (by aesop : (n:ℝ) ≠ 0), + ← Finsupp.coe_smul, ← LinearEquiv.map_smul, SetLike.mem_coe, Basis.mem_span_iff_repr_mem, + Pi.basisFun_repr, implies_true] + · rw [← NormedSpace.isVonNBounded_iff ℝ] at hs₁ ⊢ + exact Bornology.IsVonNBounded.image hs₁ ((b.ofZLatticeBasis ℝ).equivFunL : E →L[ℝ] ι → ℝ) + · exact (b.ofZLatticeBasis ℝ).equivFunL.toHomeomorph.toMeasurableEquiv.measurableSet_image.mpr hs₂ + +private theorem tendsto_card_le_div''_aux {X : Set E} (hX : ∀ ⦃x⦄ ⦃r:ℝ⦄, x ∈ X → 0 < r → r • x ∈ X) + {F : E → ℝ} (hF₁ : ∀ x ⦃r : ℝ⦄, 0 ≤ r → F (r • x) = r ^ card ι * (F x)) {c : ℝ} (hc : 0 < c) : + c • {x ∈ X | F x ≤ 1} = {x ∈ X | F x ≤ c ^ card ι} := by + ext x + simp_rw [Set.mem_smul_set_iff_inv_smul_mem₀ hc.ne', Set.mem_setOf_eq, hF₁ _ + (inv_pos_of_pos hc).le, inv_pow, inv_mul_le_iff₀ (pow_pos hc _), mul_one, and_congr_left_iff] + exact fun _ ↦ ⟨fun h ↦ (smul_inv_smul₀ hc.ne' x) ▸ hX h hc, fun h ↦ hX h (inv_pos_of_pos hc)⟩ + +/-- Docstring. -/ +theorem tendsto_card_le_div'' [FiniteDimensional ℝ E] [MeasurableSpace E] [BorelSpace E] + [Nonempty ι] {X : Set E} (hX : ∀ ⦃x⦄ ⦃r : ℝ⦄, x ∈ X → 0 < r → r • x ∈ X) + {F : E → ℝ} (h₁ : ∀ x ⦃r : ℝ⦄, 0 ≤ r → F (r • x) = r ^ card ι * (F x)) + (h₂ : IsBounded {x ∈ X | F x ≤ 1}) (h₃ : MeasurableSet {x ∈ X | F x ≤ 1}) + (h₄ : volume (frontier ((b.ofZLatticeBasis ℝ L).equivFun '' {x | x ∈ X ∧ F x ≤ 1})) = 0) : + Tendsto (fun c : ℝ ↦ + Nat.card ({x ∈ X | F x ≤ c} ∩ L : Set E) / (c : ℝ)) + atTop (𝓝 (volume ((b.ofZLatticeBasis ℝ).equivFun '' {x ∈ X | F x ≤ 1})).toReal) := by + have h : (card ι : ℝ) ≠ 0 := Nat.cast_ne_zero.mpr card_ne_zero + refine Tendsto.congr' ?_ <| (unitPartition.tendsto_card_div_pow + ((b.ofZLatticeBasis ℝ).equivFun '' {x ∈ X | F x ≤ 1}) ?_ ?_ h₄ fun x y hx hy ↦ ?_).comp + (tendsto_rpow_atTop <| inv_pos.mpr + (Nat.cast_pos.mpr card_pos) : Tendsto (fun x ↦ x ^ (card ι : ℝ)⁻¹) atTop atTop) + · filter_upwards [eventually_gt_atTop 0] with c hc + obtain ⟨hc₁, hc₂⟩ := lt_iff_le_and_ne.mp hc + rw [Function.comp_apply, ← Real.rpow_natCast, Real.rpow_inv_rpow hc₁ h, eq_comm] + congr + refine Nat.card_congr <| Equiv.subtypeEquiv ((b.ofZLatticeBasis ℝ).equivFun.toEquiv.trans + (Equiv.smulRight (a := c ^ (- (card ι : ℝ)⁻¹)) (by aesop))) fun _ ↦ ?_ + rw [Set.mem_inter_iff, Set.mem_inter_iff, Equiv.trans_apply, LinearEquiv.coe_toEquiv, + Equiv.smulRight_apply, Real.rpow_neg hc₁, Set.smul_mem_smul_set_iff₀ (by aesop), + ← Set.mem_smul_set_iff_inv_smul_mem₀ (by aesop), ← image_smul_set, + tendsto_card_le_div''_aux hX h₁ (by positivity), ← Real.rpow_natCast, ← Real.rpow_mul hc₁, + inv_mul_cancel₀ h, Real.rpow_one] + simp_rw [SetLike.mem_coe, Set.mem_image, EmbeddingLike.apply_eq_iff_eq, exists_eq_right, + and_congr_right_iff, ← b.ofZLatticeBasis_span ℝ, Basis.mem_span_iff_repr_mem, + Pi.basisFun_repr, Basis.equivFun_apply, implies_true] + · rw [← NormedSpace.isVonNBounded_iff ℝ] at h₂ ⊢ + exact Bornology.IsVonNBounded.image h₂ ((b.ofZLatticeBasis ℝ).equivFunL : E →L[ℝ] ι → ℝ) + · exact (b.ofZLatticeBasis ℝ).equivFunL.toHomeomorph.toMeasurableEquiv.measurableSet_image.mpr h₃ + · simp_rw [← image_smul_set] + apply Set.image_mono + rw [tendsto_card_le_div''_aux hX h₁ hx, + tendsto_card_le_div''_aux hX h₁ (lt_of_lt_of_le hx hy)] + exact fun a ⟨ha₁, ha₂⟩ ↦ ⟨ha₁, le_trans ha₂ <| pow_le_pow_left (le_of_lt hx) hy _⟩ + +end General + +section Pi + +open Filter Fintype Pointwise Topology Bornology + +private theorem frontier_equivFun {E : Type*} [AddCommGroup E] [Module ℝ E] {ι : Type*} [Fintype ι] + [TopologicalSpace E] [TopologicalAddGroup E] [ContinuousSMul ℝ E] [T2Space E] + (b : Basis ι ℝ E) (s : Set E) : + frontier (b.equivFun '' s) = b.equivFun '' (frontier s) := by + rw [LinearEquiv.image_eq_preimage, LinearEquiv.image_eq_preimage] + exact (Homeomorph.preimage_frontier b.equivFunL.toHomeomorph.symm s).symm + +variable {ι : Type*} [Fintype ι] +variable (L : Submodule ℤ (ι → ℝ)) [DiscreteTopology L] [IsZLattice ℝ L] + +theorem tendsto_card_div_pow (b : Basis ι ℤ L) {s : Set (ι → ℝ)} (hs₁ : IsBounded s) + (hs₂ : MeasurableSet s) (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (Nat.card (s ∩ (n : ℝ)⁻¹ • L : Set (ι → ℝ)) : ℝ) / n ^ card ι) + atTop (𝓝 ((volume s).toReal / covolume L)) := by + classical + convert tendsto_card_div_pow'' b hs₁ hs₂ ?_ + · rw [volume_image_eq_volume_div_covolume L b, ENNReal.toReal_div, + ENNReal.toReal_ofReal (covolume_pos L volume).le] + · rw [frontier_equivFun, volume_image_eq_volume_div_covolume, hs₃, ENNReal.zero_div] + +theorem tendsto_card_le_div {X : Set (ι → ℝ)} (hX : ∀ ⦃x⦄ ⦃r : ℝ⦄, x ∈ X → 0 < r → r • x ∈ X) + {F : (ι → ℝ) → ℝ} (h₁ : ∀ x ⦃r : ℝ⦄, 0 ≤ r → F (r • x) = r ^ card ι * (F x)) + (h₂ : IsBounded {x ∈ X | F x ≤ 1}) (h₃ : MeasurableSet {x ∈ X | F x ≤ 1}) + (h₄ : volume (frontier {x | x ∈ X ∧ F x ≤ 1}) = 0) [Nonempty ι] : + Tendsto (fun c : ℝ ↦ + Nat.card ({x ∈ X | F x ≤ c} ∩ L : Set (ι → ℝ)) / (c : ℝ)) + atTop (𝓝 ((volume {x ∈ X | F x ≤ 1}).toReal / covolume L)) := by + classical + let e : Free.ChooseBasisIndex ℤ ↥L ≃ ι := by + refine Fintype.equivOfCardEq ?_ + rw [← finrank_eq_card_chooseBasisIndex, ZLattice.rank ℝ, finrank_fintype_fun_eq_card] + let b := (Module.Free.chooseBasis ℤ L).reindex e + convert tendsto_card_le_div'' b hX h₁ h₂ h₃ ?_ + · rw [volume_image_eq_volume_div_covolume L b, ENNReal.toReal_div, + ENNReal.toReal_ofReal (covolume_pos L volume).le] + · rw [frontier_equivFun, volume_image_eq_volume_div_covolume, h₄, ENNReal.zero_div] + +end Pi + +section InnerProductSpace + +open Filter Pointwise Topology Bornology + +variable {E : Type*} [NormedAddCommGroup E] [InnerProductSpace ℝ E] [FiniteDimensional ℝ E] + [MeasurableSpace E] [BorelSpace E] +variable (L : Submodule ℤ E) [DiscreteTopology L] [IsZLattice ℝ L] + +/-- Docstring. -/ +theorem tendsto_card_div_pow' {s : Set E} (hs₁ : IsBounded s) (hs₂ : MeasurableSet s) + (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (Nat.card (s ∩ (n : ℝ)⁻¹ • L : Set E) : ℝ) / n ^ finrank ℝ E) + atTop (𝓝 ((volume s).toReal / covolume L)) := by + let b := Module.Free.chooseBasis ℤ L + convert tendsto_card_div_pow'' b hs₁ hs₂ ?_ + · rw [← finrank_eq_card_chooseBasisIndex, ZLattice.rank ℝ L] + · rw [volume_image_eq_volume_div_covolume' L b hs₂.nullMeasurableSet, ENNReal.toReal_div, + ENNReal.toReal_ofReal (covolume_pos L volume).le] + · rw [frontier_equivFun, volume_image_eq_volume_div_covolume', hs₃, ENNReal.zero_div] + exact NullMeasurableSet.of_null hs₃ + +/-- Docstring. -/ +theorem tendsto_card_le_div' [Nontrivial E] {X : Set E} {F : E → ℝ} + (hX : ∀ ⦃x⦄ ⦃r : ℝ⦄, x ∈ X → 0 < r → r • x ∈ X) + (h₁ : ∀ x ⦃r : ℝ⦄, 0 ≤ r → F (r • x) = r ^ finrank ℝ E * (F x)) + (h₂ : IsBounded {x ∈ X | F x ≤ 1}) (h₃ : MeasurableSet {x ∈ X | F x ≤ 1}) + (h₄ : volume (frontier {x ∈ X | F x ≤ 1}) = 0) : + Tendsto (fun c : ℝ ↦ + Nat.card ({x ∈ X | F x ≤ c} ∩ L : Set E) / (c : ℝ)) + atTop (𝓝 ((volume {x ∈ X | F x ≤ 1}).toReal / covolume L)) := by + let b := Module.Free.chooseBasis ℤ L + convert tendsto_card_le_div'' b hX ?_ h₂ h₃ ?_ + · rw [volume_image_eq_volume_div_covolume' L b h₃.nullMeasurableSet, ENNReal.toReal_div, + ENNReal.toReal_ofReal (covolume_pos L volume).le] + · have : Nontrivial L := nontrivial_of_finrank_pos <| (ZLattice.rank ℝ L).symm ▸ finrank_pos + infer_instance + · rwa [← finrank_eq_card_chooseBasisIndex, ZLattice.rank ℝ L] + · rw [frontier_equivFun, volume_image_eq_volume_div_covolume', h₄, ENNReal.zero_div] + exact NullMeasurableSet.of_null h₄ + +end InnerProductSpace + +end covolume end ZLattice diff --git a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean index 42aba8ff24d51..b62b95a972c27 100644 --- a/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean +++ b/Mathlib/Analysis/Asymptotics/AsymptoticEquivalent.lean @@ -104,6 +104,9 @@ theorem IsEquivalent.trans {l : Filter α} {u v w : α → β} (huv : u ~[l] v) u ~[l] w := (huv.isLittleO.trans_isBigO hvw.isBigO).triangle hvw.isLittleO +theorem isEquivalent_comm : u ~[l] v ↔ v ~[l] u := + ⟨fun h ↦ IsEquivalent.symm h, fun h ↦ IsEquivalent.symm h⟩ + theorem IsEquivalent.congr_left {u v w : α → β} {l : Filter α} (huv : u ~[l] v) (huw : u =ᶠ[l] w) : w ~[l] v := huv.congr' (huw.sub (EventuallyEq.refl _ _)) (EventuallyEq.refl _ _) diff --git a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean index 551320c2187d8..f49c7df397731 100644 --- a/Mathlib/Analysis/BoxIntegral/Box/Basic.lean +++ b/Mathlib/Analysis/BoxIntegral/Box/Basic.lean @@ -208,6 +208,13 @@ theorem monotone_upper : Monotone fun I : Box ι ↦ I.upper := theorem coe_subset_Icc : ↑I ⊆ Box.Icc I := fun _ hx ↦ ⟨fun i ↦ (hx i).1.le, fun i ↦ (hx i).2⟩ +theorem isBounded_Icc [Finite ι] (I : Box ι) : Bornology.IsBounded (Box.Icc I) := by + cases nonempty_fintype ι + exact Metric.isBounded_Icc _ _ + +theorem isBounded [Finite ι] (I : Box ι) : Bornology.IsBounded I.toSet := + Bornology.IsBounded.subset I.isBounded_Icc coe_subset_Icc + /-! ### Supremum of two boxes -/ diff --git a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean new file mode 100644 index 0000000000000..fdd2ba91a9caf --- /dev/null +++ b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean @@ -0,0 +1,477 @@ +/- +Copyright (c) 2024 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.Algebra.Module.ZLattice.Basic +import Mathlib.Analysis.BoxIntegral.Integrability + +/-! +# Unit Partition + +Fix `n` a positive integer. `BoxIntegral.unitPartition.Box` are boxes in `ι → ℝ` obtained by +dividing the unit box uniformly into boxes of side length `1 / n` and translating the boxes by +the lattice `ι → ℤ` so that they cover the whole space. For fixed `n`, there are indexed by +vectors `ν : ι → ℤ`. + +Let `B` be a `BoxIntegral`. A `unitPartition.Box` is admissible for `B` (more precisely its index is +admissible) if it is contained in `B`. There are finitely many admissible `unitPartition.Box` for +`B` and thus we can form the corresponing tagged prepartition, see +`BoxIntegral.unitPartition.prepartition` (note that each `unitPartition.Box` coming with its +tag situed at its "upper most" corner). If `B` satifies `BoxIntegral.hasIntegralVertices`, that +is its corners are in `ι → ℤ`, then the corresponding prepartition is actually a partition. + +## Main results + + +* `BoxIntegral.hasIntegralCorners`: a `Prop` that states that the corners of the box has +coordinates in `ℤ` + +* `BoxIntegral.unitPartition.box`: a `BoxIntegral`, indexed by `ν : ι → ℤ`, with corners +`ν i / n` and of side length `1 / n`. + +* `BoxIntegral.unitPartition.admissibleIndex`: For `B : BoxIntegral.Box`, the set of indices of +`unitPartition.Box` that are subset of `B`. It is a finite set. + +* `BoxIntegral.unitPartition.prepartition_isPartition`: For `B : BoxIntegral.Box`, if `B` +has integral corners, the prepartition of `unitPartition.Box` admissible for `B` is a partition +of `B`. + +* `BoxIntegral.unitPartition.tendsto_tsum_div_pow`: let `s` be a bounded, measurable set of `ι → ℝ` +whose frontier has zero volume and let `F` be a continuous function. Then the limit as `n → ∞` +of `∑ F x / n ^ card ι`, where the sum is over the points in `s ∩ n⁻¹ • (ι → ℤ)`, tends to the +integral of `F` over `s`. + +* `BoxIntegral.unitPartition.tendsto_card_div_pow'`: let `s` be a bounded, measurable set of +`ι → ℝ` whose frontier has zero volume. Then the limit as `n → ∞` of +`card (s ∩ n⁻¹ • (ι → ℤ)) / n ^ card ι` tends to the volume of `s`. + +* `BoxIntegral.unitPartition.tendsto_card_div_pow`: a version of `tendsto_card_div_pow` where we +assume in addition that `x • s ⊆ y • s` whenever `0 < x ≤ y`. Then we get the same limit +`card (s ∩ x⁻¹ • (ι → ℤ)) / x ^ card ι → volume s` but the limit is over a real variable `x`. + +-/ + +noncomputable section + +variable {ι : Type*} + +section hasIntegralCorners + +open Bornology + +/-- A `BoxIntegral.Box` has integral devices if its corners has coordinates in `ℤ`. -/ +def BoxIntegral.hasIntegralCorners (B : Box ι) : Prop := + ∃ l u : ι → ℤ, (∀ i, B.lower i = l i) ∧ (∀ i, B.upper i = u i) + +/-- Any bounded set is contained in a `BoxIntegral.Box` with integral corners. -/ +theorem BoxIntegral.le_hasIntegralVertices_of_isBounded [Finite ι] {s : Set (ι → ℝ)} + (h : IsBounded s) : + ∃ B : BoxIntegral.Box ι, hasIntegralCorners B ∧ s ≤ B := by + have := Fintype.ofFinite ι + obtain ⟨R, hR₁, hR₂⟩ := IsBounded.subset_ball_lt h 0 0 + let C : ℕ+ := ⟨⌈R⌉₊, Nat.ceil_pos.mpr hR₁⟩ + let I : Box ι := BoxIntegral.Box.mk (fun _ ↦ - C) (fun _ ↦ C ) (fun _ ↦ by norm_num [hR₁]) + refine ⟨I, ⟨fun _ ↦ - C, fun _ ↦ C, by aesop⟩, le_trans hR₂ ?_⟩ + suffices Metric.ball (0 : ι → ℝ) C ≤ I from + le_trans (Metric.ball_subset_ball (Nat.le_ceil R)) this + intro x hx + simp_rw [mem_ball_zero_iff, pi_norm_lt_iff (by aesop : (0:ℝ) < C), Real.norm_eq_abs, abs_lt] at hx + exact fun i ↦ ⟨(hx i).1, le_of_lt (hx i).2⟩ + +end hasIntegralCorners + +namespace BoxIntegral.unitPartition + +open Bornology MeasureTheory Fintype BoxIntegral + +variable (n : ℕ+) + +/-- A `BoxIntegral`, indexed by a positive integer `n` and `ν : ι → ℤ`, with corners `ν i / n` +and of side length `1 / n`. -/ +def box (ν : ι → ℤ) : Box ι where + lower := fun i ↦ ν i / n + upper := fun i ↦ ν i / n + 1 / n + lower_lt_upper := fun _ ↦ by norm_num + +@[simp] +theorem box_lower (ν : ι → ℤ) : + (box n ν).lower = fun i ↦ (ν i / n : ℝ) := rfl + +@[simp] +theorem box_upper (ν : ι → ℤ) : + (box n ν).upper = fun i ↦ (ν i / n + 1 / n : ℝ) := rfl + +variable {n} in +@[simp] +theorem mem_box_iff {ν : ι → ℤ} {x : ι → ℝ} : + x ∈ box n ν ↔ ∀ i, ν i / n < x i ∧ x i ≤ ν i / n + 1 / n := by + simp_rw [Box.mem_def, box, Set.mem_Ioc] + +variable {n} in +theorem mem_box_iff' {ν : ι → ℤ} {x : ι → ℝ} : + x ∈ box n ν ↔ ∀ i, ν i < n * x i ∧ n * x i ≤ ν i + 1 := by + have h_npos : 0 < (n:ℝ) := Nat.cast_pos.mpr <| PNat.pos n + simp_rw [mem_box_iff, ← _root_.le_div_iff₀' h_npos, ← div_lt_iff₀' h_npos, add_div] + +/-- The tag of a `unitPartition.Box`. -/ +abbrev tag (ν : ι → ℤ) : ι → ℝ := fun i ↦ (ν i + 1) / n + +@[simp] +theorem tag_apply (ν : ι → ℤ) (i : ι) : tag n ν i = (ν i + 1) / n := rfl + +theorem tag_injective : Function.Injective (fun ν : ι → ℤ ↦ tag n ν) := by + refine fun _ _ h ↦ Function.funext_iff.mpr fun i ↦ ?_ + have := congr_arg (fun x ↦ x i) h + field_simp at this + exact this + +theorem tag_mem (ν : ι → ℤ) : + tag n ν ∈ box n ν := by + refine mem_box_iff.mpr fun _ ↦ ?_ + rw [tag, add_div] + exact ⟨by norm_num, le_rfl⟩ + +/-- For `x : ι → ℝ`, its index is the index of the unique `unitPartition.Box` to which +it belongs. -/ +def index (x : ι → ℝ) (i : ι) : ℤ := ⌈n * x i⌉ - 1 + +@[simp] +theorem index_apply {x : ι → ℝ} (i : ι) : + index n x i = ⌈n * x i⌉ - 1 := rfl + +variable {n} in +theorem mem_box_iff_index {x : ι → ℝ} {ν : ι → ℤ} : + x ∈ box n ν ↔ index n x = ν := by + simp_rw [mem_box_iff', Function.funext_iff, index_apply, sub_eq_iff_eq_add, Int.ceil_eq_iff, + Int.cast_add, Int.cast_one, add_sub_cancel_right] + +@[simp] +theorem index_tag (ν : ι → ℤ) : + index n (tag n ν) = ν := mem_box_iff_index.mp (tag_mem n ν) + +variable {n} in +theorem disjoint {ν ν' : ι → ℤ} : + ν ≠ ν' ↔ Disjoint (box n ν).toSet (box n ν').toSet := by + rw [not_iff_not.symm, ne_eq, not_not, Set.not_disjoint_iff] + refine ⟨fun h ↦ ?_, fun ⟨x, hx, hx'⟩ ↦ ?_⟩ + · exact ⟨tag n ν, tag_mem n ν, h ▸ tag_mem n ν⟩ + · rw [← mem_box_iff_index.mp hx, ← mem_box_iff_index.mp hx'] + +theorem box_injective : Function.Injective (fun ν : ι → ℤ ↦ box n ν) := by + intro _ _ h + contrapose! h + exact Box.ne_of_disjoint_coe (disjoint.mp h) + +variable [Fintype ι] + +theorem diam_boxIcc (ν : ι → ℤ) : + Metric.diam (Box.Icc (box n ν)) ≤ 1 / n := by + refine ENNReal.toReal_le_of_le_ofReal (by positivity) ?_ + rw [BoxIntegral.Box.Icc_eq_pi] + refine EMetric.diam_pi_le_of_le (fun i ↦ ?_) + rw [Real.ediam_Icc, box, add_sub_cancel_left, ENNReal.ofReal_div_of_pos (Nat.cast_pos.mpr n.pos), + ENNReal.ofReal_one] + +@[simp] +theorem volume_box (ν : ι → ℤ) : + volume (box n ν : Set (ι → ℝ)) = 1 / n ^ card ι := by + simp_rw [volume_pi, BoxIntegral.Box.coe_eq_pi, Measure.pi_pi, Real.volume_Ioc, box, + add_sub_cancel_left, Finset.prod_const, ENNReal.ofReal_div_of_pos (Nat.cast_pos.mpr n.pos), + ENNReal.ofReal_one, ENNReal.ofReal_natCast, Finset.card_univ, one_div, ENNReal.inv_pow] + +theorem setFinite_index {s : Set (ι → ℝ)} (hs₁ : NullMeasurableSet s) (hs₂ : volume s ≠ ⊤) : + Set.Finite {ν : ι → ℤ | ↑(box n ν) ⊆ s} := by + refine (Measure.finite_const_le_meas_of_disjoint_iUnion₀ volume (ε := 1 / n ^ card ι) + (by norm_num) (As := fun ν : ι → ℤ ↦ (box n ν) ∩ s) (fun ν ↦ ?_) (fun _ _ h ↦ ?_) ?_).subset + (fun _ hν ↦ ?_) + · refine NullMeasurableSet.inter ?_ hs₁ + exact (box n ν).measurableSet_coe.nullMeasurableSet + · exact ((Disjoint.inter_right _ (disjoint.mp h)).inter_left _ ).aedisjoint + · exact lt_top_iff_ne_top.mp <| measure_lt_top_of_subset (by aesop) hs₂ + · rw [Set.mem_setOf, Set.inter_eq_self_of_subset_left hν, volume_box] + +/-- For `B : BoxIntegral.Box`, the set of indices of `unitPartition.Box` that are subset of `B`. +It is a finite set. These boxes cover `B` if it has integral corners, see +`unitPartition.prepartition_isPartition`. -/ +def admissibleIndex (B : Box ι) : Finset (ι → ℤ) := by + refine (setFinite_index n B.measurableSet_coe.nullMeasurableSet ?_).toFinset + exact lt_top_iff_ne_top.mp (IsBounded.measure_lt_top B.isBounded) + +variable {n} in +theorem mem_admissibleIndex_iff {B : Box ι} {ν : ι → ℤ} : + ν ∈ admissibleIndex n B ↔ box n ν ≤ B := by + rw [admissibleIndex, Set.Finite.mem_toFinset, Set.mem_setOf_eq, Box.coe_subset_coe] + +open Classical in +/-- For `B : BoxIntegral.Box`, the `TaggedPrepartition` formed by the set of all +`unitPartition.Box` whose index is `B`-admissible. -/ +def prepartition (B : Box ι) : TaggedPrepartition B where + boxes := Finset.image (fun ν ↦ box n ν) (admissibleIndex n B) + le_of_mem' _ hI := by + obtain ⟨_, hν, rfl⟩ := Finset.mem_image.mp hI + exact mem_admissibleIndex_iff.mp hν + pairwiseDisjoint _ hI₁ _ hI₂ h := by + obtain ⟨_, _, rfl⟩ := Finset.mem_image.mp hI₁ + obtain ⟨_, _, rfl⟩ := Finset.mem_image.mp hI₂ + exact disjoint.mp fun x ↦ h (congrArg (box n) x) + tag I := by + by_cases hI : ∃ ν ∈ admissibleIndex n B, I = box n ν + · exact tag n hI.choose + · exact B.exists_mem.choose + tag_mem_Icc I := by + by_cases hI : ∃ ν ∈ admissibleIndex n B, I = box n ν + · simp_rw [dif_pos hI] + exact Box.coe_subset_Icc <| (mem_admissibleIndex_iff.mp hI.choose_spec.1) (tag_mem n _) + · simp_rw [dif_neg hI] + exact Box.coe_subset_Icc B.exists_mem.choose_spec + +variable {n} in +@[simp] +theorem mem_prepartition_iff {B I : Box ι} : + I ∈ prepartition n B ↔ ∃ ν ∈ admissibleIndex n B, box n ν = I := by + classical + rw [prepartition, TaggedPrepartition.mem_mk, Prepartition.mem_mk, Finset.mem_image] + +variable {n} in +theorem mem_prepartition_boxes_iff {B I : Box ι} : + I ∈ (prepartition n B).boxes ↔ ∃ ν ∈ admissibleIndex n B, box n ν = I := by + rw [Prepartition.mem_boxes, TaggedPrepartition.mem_toPrepartition, mem_prepartition_iff] + +theorem prepartition_tag {ν : ι → ℤ} {B : Box ι} (hν : ν ∈ admissibleIndex n B) : + (prepartition n B).tag (box n ν) = tag n ν := by + dsimp only [prepartition] + have h : ∃ ν' ∈ admissibleIndex n B, box n ν = box n ν' := ⟨ν, hν, rfl⟩ + rw [dif_pos h, (tag_injective n).eq_iff, ← (box_injective n).eq_iff] + exact h.choose_spec.2.symm + +theorem box_index_tag_eq_self {B I : Box ι} (hI : I ∈ (prepartition n B).boxes) : + box n (index n ((prepartition n B).tag I)) = I := by + obtain ⟨ν, hν, rfl⟩ := mem_prepartition_boxes_iff.mp hI + rw [prepartition_tag n hν, index_tag] + +theorem prepartition_isHenstock (B : Box ι) : + (prepartition n B).IsHenstock := by + intro _ hI + obtain ⟨ν, hν, rfl⟩ := mem_prepartition_iff.mp hI + rw [prepartition_tag n hν] + exact Box.coe_subset_Icc (tag_mem _ _) + +theorem prepartition_isSubordinate (B : Box ι) {r : ℝ} (hr : 0 < r) (hn : 1 / n ≤ r) : + (prepartition n B).IsSubordinate (fun _ ↦ ⟨r, hr⟩) := by + intro _ hI + obtain ⟨ν, hν, rfl⟩ := mem_prepartition_iff.mp hI + refine fun _ h ↦ le_trans (Metric.dist_le_diam_of_mem (Box.isBounded_Icc _) h ?_) ?_ + · rw [prepartition_tag n hν] + exact Box.coe_subset_Icc (tag_mem _ _) + · exact le_trans (diam_boxIcc n ν) hn + +theorem mem_admissibleIndex_of_mem_box_aux₁ (x : ℝ) (a : ℤ) : + a < x ↔ a ≤ (⌈n * x⌉ - 1) / (n : ℝ) := by + rw [le_div_iff₀' (by positivity), le_sub_iff_add_le, show (n : ℝ) * a + 1 = + (n * a + 1 : ℤ) by norm_cast, Int.cast_le, Int.add_one_le_ceil_iff, Int.cast_mul, + Int.cast_natCast, mul_lt_mul_left (by positivity)] + +theorem mem_admissibleIndex_of_mem_box_aux₂ (x : ℝ) (a : ℤ) : + x ≤ a ↔ (⌈n * x⌉ - 1) / (n : ℝ) + 1 / n ≤ a := by + rw [← add_div, sub_add_cancel, div_le_iff₀' (by positivity), show (n : ℝ) * a = (n * a : ℤ) + by norm_cast, Int.cast_le, Int.ceil_le, Int.cast_mul, Int.cast_natCast, mul_le_mul_left + (by positivity)] + +/-- If `B : BoxIntegral.Box` has integral corners and contains the point `x`, then the index of +`x` is admissible for `B`. -/ +theorem mem_admissibleIndex_of_mem_box {B : Box ι} (hB : hasIntegralCorners B) {x : ι → ℝ} + (hx : x ∈ B) : index n x ∈ admissibleIndex n B := by + obtain ⟨l, u, hl, hu⟩ := hB + simp_rw [mem_admissibleIndex_iff, Box.le_iff_bounds, box_lower, box_upper, Pi.le_def, + index_apply, hl, hu, ← forall_and] + push_cast + refine fun i ↦ ⟨?_, ?_⟩ + · exact (mem_admissibleIndex_of_mem_box_aux₁ n (x i) (l i)).mp ((hl i) ▸ (hx i).1) + · exact (mem_admissibleIndex_of_mem_box_aux₂ n (x i) (u i)).mp ((hu i) ▸ (hx i).2) + +theorem exists_admissibleIndex {B : Box ι} (hB : hasIntegralCorners B) {x : ι → ℝ} (hx : x ∈ B) : + ∃ ν ∈ admissibleIndex n B, box n ν = box n (index n x) := + ⟨index n x, mem_admissibleIndex_of_mem_box n hB hx, rfl⟩ + +/-- If `B : BoxIntegral.Box` has integral corners, then `prepartition n B` is a partition of +`B`. -/ +theorem prepartition_isPartition {B : Box ι} (hB : hasIntegralCorners B) : + (prepartition n B).IsPartition := by + refine fun x hx ↦ ⟨box n (index n x), ?_, mem_box_iff_index.mpr rfl⟩ + rw [TaggedPrepartition.mem_toPrepartition, mem_prepartition_iff] + exact exists_admissibleIndex n hB hx + +open Submodule Pointwise BigOperators + +open scoped Pointwise + +variable (c : ℝ) (s : Set (ι → ℝ)) (F : (ι → ℝ) → ℝ) + +local notation "L" => span ℤ (Set.range (Pi.basisFun ℝ ι)) + +theorem tag_mem_smul_span (ν : ι → ℤ) : + tag n ν ∈ (n : ℝ)⁻¹ • L := by + rw [ZSpan.smul _ (by aesop), Basis.mem_span_iff_repr_mem] + refine fun i ↦ ⟨ν i + 1, ?_⟩ + rw [Basis.repr_isUnitSMul, Pi.basisFun_repr, tag_apply, Units.smul_def, smul_eq_mul, + div_eq_inv_mul, ← mul_assoc, IsUnit.val_inv_mul, one_mul, map_add, map_one, eq_intCast] + +theorem tag_index_eq_self_of_mem_smul_span {x : ι → ℝ} (hx : x ∈ (n : ℝ)⁻¹ • L) : + tag n (index n x) = x := by + rw [ZSpan.smul _ (by aesop), Basis.mem_span_iff_repr_mem] at hx + ext i + rsuffices ⟨a, ha⟩ : ∃ a : ℤ, a = n * x i := by + specialize hx i + rwa [Basis.repr_isUnitSMul, Pi.basisFun_repr, Units.smul_def, Units.val_inv_eq_inv_val, + IsUnit.unit_spec, inv_inv] at hx + rw [tag_apply, index_apply, Int.cast_sub, Int.cast_one, sub_add_cancel, ← ha, Int.ceil_intCast, + ha, mul_div_assoc, mul_div_cancel₀ _ (by aesop)] + +theorem eq_of_index_eq_index_and_mem_smul_span {x y : ι → ℝ} (hx : x ∈ (n : ℝ)⁻¹ • L) + (hy : y ∈ (n : ℝ)⁻¹ • L) (h : index n x = index n y) : x = y := by + rw [← tag_index_eq_self_of_mem_smul_span n hx, ← tag_index_eq_self_of_mem_smul_span n hy, h] + +theorem integralSum_eq_tsum_div {B : Box ι} (hB : hasIntegralCorners B) (hs₀ : s ≤ B) : + BoxIntegral.integralSum (Set.indicator s F) + (BoxIntegral.BoxAdditiveMap.toSMul (Measure.toBoxAdditive volume)) + (prepartition n B) = (∑' x : ↑(s ∩ (n : ℝ)⁻¹ • L), F x) / n ^ card ι := by + classical + unfold BoxIntegral.integralSum + have : Fintype ↑(s ∩ (n : ℝ)⁻¹ • L) := by + apply Set.Finite.fintype + rw [← coe_pointwise_smul, ZSpan.smul _ (by aesop)] + exact ZSpan.setFinite_inter _ (IsBounded.subset B.isBounded hs₀) + rw [tsum_fintype, Finset.sum_set_coe, Finset.sum_div, eq_comm] + simp_rw [Set.indicator_apply, apply_ite, BoxAdditiveMap.toSMul_apply, Measure.toBoxAdditive_apply, + smul_eq_mul, mul_zero, Finset.sum_ite, Finset.sum_const_zero, add_zero] + refine Finset.sum_bij (fun x _ ↦ box n (index n x)) (fun _ hx ↦ Finset.mem_filter.mpr ?_) + (fun _ hx _ hy h ↦ ?_) (fun I hI ↦ ?_) (fun _ hx ↦ ?_) + · rw [Set.mem_toFinset] at hx + refine ⟨mem_prepartition_boxes_iff.mpr (exists_admissibleIndex n hB (hs₀ hx.1)), ?_⟩ + simp_rw [prepartition_tag n (mem_admissibleIndex_of_mem_box n hB (hs₀ hx.1)), + tag_index_eq_self_of_mem_smul_span n hx.2, hx.1] + · rw [Set.mem_toFinset] at hx hy + exact eq_of_index_eq_index_and_mem_smul_span n hx.2 hy.2 (box_injective n h) + · rw [Finset.mem_filter] at hI + refine ⟨(prepartition n B).tag I, Set.mem_toFinset.mpr ⟨hI.2, ?_⟩, box_index_tag_eq_self n hI.1⟩ + rw [← box_index_tag_eq_self n hI.1, prepartition_tag n + (mem_admissibleIndex_of_mem_box n hB (hs₀ hI.2))] + exact tag_mem_smul_span _ _ + · rw [Set.mem_toFinset] at hx + rw [volume_box, prepartition_tag n (mem_admissibleIndex_of_mem_box n hB (hs₀ hx.1)), + tag_index_eq_self_of_mem_smul_span n hx.2, ENNReal.toReal_div, + ENNReal.one_toReal, ENNReal.toReal_pow, ENNReal.toReal_nat, mul_comm_div, one_mul] + +open Filter + +/-- Let `s` be a bounded, measurable set of `ι → ℝ` whose frontier has zero volume and let `F` +be a continuous function. Then the limit as `n → ∞` of `∑ F x / n ^ card ι`, where the sum is +over the points in `s ∩ n⁻¹ • (ι → ℤ)`, tends to the integral of `F` over `s`. -/ +theorem tendsto_tsum_div_pow (hF : Continuous F) (hs₁ : Bornology.IsBounded s) + (hs₂ : MeasurableSet s) (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (∑' x : ↑(s ∩ (n : ℝ)⁻¹ • L), F x) / n ^ card ι) + atTop (nhds (∫ x in s, F x)) := by + obtain ⟨B, hB, hs₀⟩ := le_hasIntegralVertices_of_isBounded hs₁ + refine Metric.tendsto_atTop.mpr fun ε hε ↦ ?_ + have h₁ : ∃ C, ∀ x ∈ Box.Icc B, ‖Set.indicator s F x‖ ≤ C := by + obtain ⟨C₀, h₀⟩ := IsCompact.exists_bound_of_continuousOn (Box.isCompact_Icc B) hF.continuousOn + refine ⟨max 0 C₀, fun x hx ↦ ?_⟩ + by_cases hx' : x ∈ s + · rw [Set.indicator_of_mem hx'] + exact le_max_of_le_right (h₀ x hx) + · rw [Set.indicator_of_not_mem hx', norm_zero] + exact le_max_left 0 _ + have h₂ : ∀ᵐ x, ContinuousAt (s.indicator F) x := by + filter_upwards [compl_mem_ae_iff.mpr hs₃] with _ h + using ContinuousOn.continuousAt_indicator (hF.continuousOn) h + obtain ⟨r, hr₁, hr₂⟩ := (BoxIntegral.hasIntegral_iff.mp <| + AEContinuous.hasBoxIntegral (volume : Measure (ι → ℝ)) h₁ h₂ + BoxIntegral.IntegrationParams.Riemann) (ε / 2) (half_pos hε) + refine ⟨⌈(r 0 0 : ℝ)⁻¹⌉₊, fun n hn ↦ lt_of_le_of_lt ?_ (half_lt_self_iff.mpr hε)⟩ + lift n to ℕ+ using lt_of_lt_of_le (Nat.ceil_pos.mpr (inv_pos.mpr (by convert (r 0 0).prop))) hn + rw [← integralSum_eq_tsum_div _ s F hB hs₀, ← Measure.restrict_restrict_of_subset hs₀, + ← MeasureTheory.integral_indicator hs₂] + refine hr₂ 0 _ ⟨?_, fun _ ↦ ?_, fun h ↦ ?_, fun h ↦ ?_⟩ (prepartition_isPartition _ hB) + · rw [show r 0 = fun _ ↦ r 0 0 from Function.funext_iff.mpr (hr₁ 0 rfl)] + apply prepartition_isSubordinate n B + rw [one_div, inv_le_comm₀ (Nat.cast_pos.mpr <| PNat.pos n) (by convert (r 0 0).prop)] + exact le_trans (Nat.le_ceil _) (Nat.cast_le.mpr hn) + · exact prepartition_isHenstock n B + · simp only [IntegrationParams.Riemann, Bool.false_eq_true] at h + · simp only [IntegrationParams.Riemann, Bool.false_eq_true] at h + +/-- Let `s` be a bounded, measurable set of `ι → ℝ` whose frontier has zero volume. Then the limit +as `n → ∞` of `card (s ∩ n⁻¹ • (ι → ℤ)) / n ^ card ι` tends to the volume of `s`. This is a +special of `tendsto_card_div_pow` with `F = 1`. -/ +theorem tendsto_card_div_pow' (hs₁ : Bornology.IsBounded s) (hs₂ : MeasurableSet s) + (hs₃ : volume (frontier s) = 0) : + Tendsto (fun n : ℕ ↦ (Nat.card ↑(s ∩ (n : ℝ)⁻¹ • L) : ℝ) / n ^ card ι) + atTop (nhds (volume s).toReal) := by + convert tendsto_tsum_div_pow s (fun _ ↦ 1) continuous_const hs₁ hs₂ hs₃ + · rw [tsum_const, nsmul_eq_mul, mul_one, Nat.cast_inj] + · rw [MeasureTheory.setIntegral_const, smul_eq_mul, mul_one] + +private def tendsto_card_div_pow₁ {c : ℝ} (hc : c ≠ 0) : + ↑(s ∩ c⁻¹ • L) ≃ ↑(c • s ∩ L) := + Equiv.subtypeEquiv (Equiv.smulRight hc) (fun x ↦ by + simp_rw [Set.mem_inter_iff, Equiv.smulRight_apply, Set.smul_mem_smul_set_iff₀ hc, + ← Set.mem_inv_smul_set_iff₀ hc]) + +private theorem tendsto_card_div_pow₂ (hs₁ : Bornology.IsBounded s) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) {x y : ℝ} (hx : 0 < x) (hy : x ≤ y) : + Nat.card ↑(s ∩ x⁻¹ • L) ≤ Nat.card ↑(s ∩ y⁻¹ • L) := by + rw [Nat.card_congr (tendsto_card_div_pow₁ s (ne_of_gt hx)), + Nat.card_congr (tendsto_card_div_pow₁ s (ne_of_gt (lt_of_lt_of_le hx hy)))] + refine Nat.card_mono ?_ ?_ + · exact ZSpan.setFinite_inter _ (IsBounded.smul₀ hs₁ y) + · gcongr + exact hs₄ hx hy + +private theorem tendsto_card_div_pow₃ (hs₁ : Bornology.IsBounded s) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) : + ∀ᶠ x : ℝ in atTop, (Nat.card ↑(s ∩ (⌊x⌋₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι ≤ + (Nat.card ↑(s ∩ x⁻¹ • L) : ℝ) / x ^ card ι := by + filter_upwards [eventually_ge_atTop 1] with x hx + gcongr + exact tendsto_card_div_pow₂ s hs₁ hs₄ (Nat.cast_pos.mpr (Nat.floor_pos.mpr hx)) + (Nat.floor_le (by positivity)) + +private theorem tendsto_card_div_pow₄ (hs₁ : Bornology.IsBounded s) + (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) : + ∀ᶠ x : ℝ in atTop, (Nat.card ↑(s ∩ x⁻¹ • L) : ℝ) / x ^ card ι ≤ + (Nat.card ↑(s ∩ (⌈x⌉₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι := by + filter_upwards [eventually_gt_atTop 0] with x hx + gcongr + exact tendsto_card_div_pow₂ s hs₁ hs₄ hx (Nat.le_ceil _) + +private theorem tendsto_card_div_pow₅ : + (fun x ↦ (Nat.card ↑(s ∩ (⌊x⌋₊ : ℝ)⁻¹ • L) : ℝ) / ⌊x⌋₊ ^ card ι * (⌊x⌋₊ / x) ^ card ι) + =ᶠ[atTop] (fun x ↦ (Nat.card ↑(s ∩ (⌊x⌋₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι) := by + filter_upwards [eventually_ge_atTop 1] with x hx + have : 0 < ⌊x⌋₊ := Nat.floor_pos.mpr hx + field_simp [hx] + +private theorem tendsto_card_div_pow₆ : + (fun x ↦ (Nat.card ↑(s ∩ (⌈x⌉₊ : ℝ)⁻¹ • L) : ℝ) / ⌈x⌉₊ ^ card ι * (⌈x⌉₊ / x) ^ card ι) + =ᶠ[atTop] (fun x ↦ (Nat.card ↑(s ∩ (⌈x⌉₊ : ℝ)⁻¹ • L) : ℝ) / x ^ card ι) := by + filter_upwards [eventually_ge_atTop 1] with x hx + have : 0 < ⌊x⌋₊ := Nat.floor_pos.mpr hx + field_simp [hx] + +theorem tendsto_card_div_pow (hs₁ : Bornology.IsBounded s) (hs₂ : MeasurableSet s) + (hs₃ : volume (frontier s) = 0) (hs₄ : ∀ ⦃x y : ℝ⦄, 0 < x → x ≤ y → x • s ⊆ y • s) : + Tendsto (fun x : ℝ ↦ (Nat.card ↑(s ∩ x⁻¹ • L) : ℝ) / x ^ card ι) + atTop (nhds (volume s).toReal) := by + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' ?_ ?_ + (tendsto_card_div_pow₃ s hs₁ hs₄) (tendsto_card_div_pow₄ s hs₁ hs₄) + · rw [show (volume s).toReal = (volume s).toReal * 1 ^ card ι by ring] + refine Tendsto.congr' (tendsto_card_div_pow₅ s) (Tendsto.mul ?_ (Tendsto.pow ?_ _)) + · exact Tendsto.comp (tendsto_card_div_pow' s hs₁ hs₂ hs₃) tendsto_nat_floor_atTop + · exact tendsto_nat_floor_div_atTop + · rw [show (volume s).toReal = (volume s).toReal * 1 ^ card ι by ring] + refine Tendsto.congr' (tendsto_card_div_pow₆ s) (Tendsto.mul ?_ (Tendsto.pow ?_ _)) + · exact Tendsto.comp (tendsto_card_div_pow' s hs₁ hs₂ hs₃) tendsto_nat_ceil_atTop + · exact tendsto_nat_ceil_div_atTop + +end BoxIntegral.unitPartition diff --git a/Mathlib/Analysis/Complex/Basic.lean b/Mathlib/Analysis/Complex/Basic.lean index e4257988bc45b..e3597157acf0d 100644 --- a/Mathlib/Analysis/Complex/Basic.lean +++ b/Mathlib/Analysis/Complex/Basic.lean @@ -354,13 +354,20 @@ def ofRealLI : ℝ →ₗᵢ[ℝ] ℂ := theorem isometry_ofReal : Isometry ((↑) : ℝ → ℂ) := ofRealLI.isometry +theorem isUniformEmbedding_ofReal : IsUniformEmbedding ((↑) : ℝ → ℂ) := + ofRealLI.isometry.isUniformEmbedding + @[continuity, fun_prop] theorem continuous_ofReal : Continuous ((↑) : ℝ → ℂ) := ofRealLI.continuous +theorem _root_.Filter.tendsto_ofReal_iff {α : Type*} {l : Filter α} {f : α → ℝ} {x : ℝ} : + Tendsto (fun x ↦ (f x : ℂ)) l (𝓝 (x : ℂ)) ↔ Tendsto f l (𝓝 x) := + isUniformEmbedding_ofReal.toClosedEmbedding.tendsto_nhds_iff.symm + lemma _root_.Filter.Tendsto.ofReal {α : Type*} {l : Filter α} {f : α → ℝ} {x : ℝ} (hf : Tendsto f l (𝓝 x)) : Tendsto (fun x ↦ (f x : ℂ)) l (𝓝 (x : ℂ)) := - (continuous_ofReal.tendsto _).comp hf + tendsto_ofReal_iff.mpr hf /-- The only continuous ring homomorphism from `ℝ` to `ℂ` is the identity. -/ theorem ringHom_eq_ofReal_of_continuous {f : ℝ →+* ℂ} (h : Continuous f) : f = ofRealHom := by diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index 46513cef6be86..fbcd4627f44d1 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -89,6 +89,14 @@ def polarCoord : PartialHomeomorph (ℝ × ℝ) (ℝ × ℝ) where · exact (Complex.continuousAt_arg hz).continuousWithinAt · exact Complex.equivRealProdCLM.symm.continuous.continuousOn +theorem measurable_polarCoord : + Measurable (polarCoord) := + Measurable.prod_mk (by fun_prop) Complex.measurableEquivRealProd.symm.measurable.carg + +theorem continuous_polarCoord_symm : + Continuous (polarCoord.symm) := + Continuous.prod_mk (by fun_prop) (by fun_prop) + theorem hasFDerivAt_polarCoord_symm (p : ℝ × ℝ) : HasFDerivAt polarCoord.symm (LinearMap.toContinuousLinearMap (Matrix.toLin (Basis.finTwoProd ℝ) (Basis.finTwoProd ℝ) @@ -99,6 +107,14 @@ theorem hasFDerivAt_polarCoord_symm (p : ℝ × ℝ) : (hasFDerivAt_fst.mul ((hasDerivAt_sin p.2).comp_hasFDerivAt p hasFDerivAt_snd)) using 2 <;> simp [smul_smul, add_comm, neg_mul, smul_neg, neg_smul _ (ContinuousLinearMap.snd ℝ ℝ ℝ)] +theorem FDerivAt_polarCoord_symm_det (p : ℝ × ℝ) : + (LinearMap.toContinuousLinearMap (Matrix.toLin (Basis.finTwoProd ℝ) (Basis.finTwoProd ℝ) + !![cos p.2, -p.1 * sin p.2; sin p.2, p.1 * cos p.2])).det = p.1 := by + conv_rhs => rw [← one_mul p.1, ← cos_sq_add_sin_sq p.2] + simp only [neg_mul, LinearMap.det_toContinuousLinearMap, LinearMap.det_toLin, + Matrix.det_fin_two_of, sub_neg_eq_add] + ring + -- Porting note: this instance is needed but not automatically synthesised instance : Measure.IsAddHaarMeasure volume (G := ℝ × ℝ) := Measure.prod.instIsAddHaarMeasure _ _ @@ -121,30 +137,38 @@ theorem polarCoord_source_ae_eq_univ : polarCoord.source =ᵐ[volume] univ := by theorem integral_comp_polarCoord_symm {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] (f : ℝ × ℝ → E) : (∫ p in polarCoord.target, p.1 • f (polarCoord.symm p)) = ∫ p, f p := by - set B : ℝ × ℝ → ℝ × ℝ →L[ℝ] ℝ × ℝ := fun p => - LinearMap.toContinuousLinearMap (Matrix.toLin (Basis.finTwoProd ℝ) (Basis.finTwoProd ℝ) - !![cos p.2, -p.1 * sin p.2; sin p.2, p.1 * cos p.2]) - have A : ∀ p ∈ polarCoord.symm.source, HasFDerivAt polarCoord.symm (B p) p := fun p _ => - hasFDerivAt_polarCoord_symm p - have B_det : ∀ p, (B p).det = p.1 := by - intro p - conv_rhs => rw [← one_mul p.1, ← cos_sq_add_sin_sq p.2] - simp only [B, neg_mul, LinearMap.det_toContinuousLinearMap, LinearMap.det_toLin, - Matrix.det_fin_two_of, sub_neg_eq_add] - ring symm calc ∫ p, f p = ∫ p in polarCoord.source, f p := by rw [← setIntegral_univ] apply setIntegral_congr_set exact polarCoord_source_ae_eq_univ.symm - _ = ∫ p in polarCoord.target, abs (B p).det • f (polarCoord.symm p) := by - apply integral_target_eq_integral_abs_det_fderiv_smul volume A + _ = ∫ p in polarCoord.target, |p.1| • f (polarCoord.symm p) := by + rw [← PartialHomeomorph.symm_target, integral_target_eq_integral_abs_det_fderiv_smul volume + (fun p _ ↦ hasFDerivAt_polarCoord_symm p), PartialHomeomorph.symm_source] + simp_rw [FDerivAt_polarCoord_symm_det] _ = ∫ p in polarCoord.target, p.1 • f (polarCoord.symm p) := by apply setIntegral_congr_fun polarCoord.open_target.measurableSet fun x hx => ?_ - rw [B_det, abs_of_pos] + rw [abs_of_pos] exact hx.1 +theorem lintegral_comp_polarCoord_symm (f : ℝ × ℝ → ENNReal) : + ∫⁻ (p : ℝ × ℝ) in polarCoord.target, (p.1).toNNReal • f (polarCoord.symm p) = + ∫⁻ (p : ℝ × ℝ), f p := by + symm + calc + _ = ∫⁻ p in polarCoord.symm '' polarCoord.target, f p := by + rw [← setLIntegral_univ, setLIntegral_congr polarCoord_source_ae_eq_univ.symm, + polarCoord.symm_image_target_eq_source ] + _ = ∫⁻ (p : ℝ × ℝ) in polarCoord.target, |p.1|.toNNReal • f (polarCoord.symm p) := by + rw [lintegral_image_eq_lintegral_abs_det_fderiv_mul volume _ + (fun p _ ↦ (hasFDerivAt_polarCoord_symm p).hasFDerivWithinAt)] + simp_rw [FDerivAt_polarCoord_symm_det]; rfl + exacts [polarCoord.symm.injOn, measurableSet_Ioi.prod measurableSet_Ioo] + _ = ∫⁻ (p : ℝ × ℝ) in polarCoord.target, (p.1).toNNReal • f (polarCoord.symm p) := by + refine setLIntegral_congr_fun polarCoord.open_target.measurableSet ?_ + filter_upwards with _ hx using by rw [abs_of_pos (by convert hx.1)] + end Real noncomputable section Complex @@ -158,6 +182,14 @@ It is a homeomorphism between `ℂ - ℝ≤0` and `(0, +∞) × (-π, π)`. -/ protected noncomputable def polarCoord : PartialHomeomorph ℂ (ℝ × ℝ) := equivRealProdCLM.toHomeomorph.transPartialHomeomorph polarCoord +protected theorem measurable_polarCoord : + Measurable (Complex.polarCoord) := + measurable_polarCoord.comp Complex.measurableEquivRealProd.measurable + +protected theorem continuous_polarCoord_symm : + Continuous (Complex.polarCoord.symm) := + equivRealProdCLM.symm.continuous.comp continuous_polarCoord_symm + protected theorem polarCoord_apply (a : ℂ) : Complex.polarCoord a = (Complex.abs a, Complex.arg a) := by simp_rw [Complex.abs_def, Complex.normSq_apply, ← pow_two] @@ -176,6 +208,30 @@ protected theorem polarCoord_symm_apply (p : ℝ × ℝ) : theorem polarCoord_symm_abs (p : ℝ × ℝ) : Complex.abs (Complex.polarCoord.symm p) = |p.1| := by simp +open scoped ComplexOrder + +theorem polarCoord_symm_mem_polarCoord_source (x : ℝ × ℝ) : + Complex.polarCoord.symm x ∈ Complex.polarCoord.source ↔ + x.1 ≠ 0 ∧ (x.1 > 0 → ∀ k : ℤ, π + k * (2 * π) ≠ x.2) ∧ + (x.1 < 0 → ∀ k : ℤ, k * (2 * π) ≠ x.2) := by + simp_rw (config := {singlePass := true}) [← not_iff_not, Complex.polarCoord_symm_apply, + Complex.polarCoord_source, mem_slitPlane_iff_arg, not_and_or, ne_eq, not_not, mul_eq_zero, + not_and_or, Classical.not_imp, not_forall, not_not, ofReal_eq_zero, ofReal_cos, ofReal_sin, + cos_add_sin_I, exp_ne_zero, or_false] + obtain hx | hx | hx := lt_trichotomy x.1 0 + · simp_rw [hx, hx.ne, not_lt_of_gt hx, false_and, false_or, true_and, or_false] + have : (x.1 * cexp (x.2 * I)).arg = π ↔ (cexp (x.2 * I)).arg = 0 := by + simp_rw [arg_eq_pi_iff_lt_zero, arg_eq_zero_iff_zero_le, ofReal_mul_neg_iff, hx, + not_lt_of_gt hx, true_and, false_and, or_false, lt_iff_le_and_ne, ne_eq, eq_comm, + exp_ne_zero, not_false_eq_true, and_true] + simp_rw [this, arg_exp_mul_I, toIocMod_eq_iff, zero_add, zsmul_eq_mul, eq_comm, + and_iff_right_iff_imp] + exact fun _ ↦ ⟨neg_neg_iff_pos.mpr Real.pi_pos, by ring_nf; positivity⟩ + · simp_rw [hx, true_or, or_true] + · simp_rw [hx, hx.ne', not_lt_of_gt hx, false_and, or_false, true_and, false_or, arg_real_mul _ + hx, arg_exp_mul_I, toIocMod_eq_iff, zsmul_eq_mul, eq_comm, and_iff_right_iff_imp] + exact fun _ ↦ ⟨Left.neg_lt_self Real.pi_pos, by linarith⟩ + @[deprecated (since := "2024-07-15")] alias polardCoord_symm_abs := polarCoord_symm_abs protected theorem integral_comp_polarCoord_symm {E : Type*} [NormedAddCommGroup E] @@ -185,4 +241,10 @@ protected theorem integral_comp_polarCoord_symm {E : Type*} [NormedAddCommGroup measurableEquivRealProd.symm.measurableEmbedding, ← integral_comp_polarCoord_symm] rfl +protected theorem lintegral_comp_polarCoord_symm (f : ℂ → ENNReal) (hf : Measurable f) : + (∫⁻ p in polarCoord.target, (p.1).toNNReal • f (Complex.polarCoord.symm p)) = ∫⁻ p, f p := by + rw [← (Complex.volume_preserving_equiv_real_prod.symm).lintegral_comp hf, + ← lintegral_comp_polarCoord_symm] + rfl + end Complex diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean index abdc88a5f991e..bc70d5caeb295 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Real.lean @@ -82,6 +82,10 @@ lemma rpow_eq_zero (hx : 0 ≤ x) (hy : y ≠ 0) : x ^ y = 0 ↔ x = 0 := by lemma rpow_ne_zero (hx : 0 ≤ x) (hy : y ≠ 0) : x ^ y ≠ 0 ↔ x ≠ 0 := Real.rpow_eq_zero hx hy |>.not +@[simp] +theorem rpow_ne_zero_of_pos {x : ℝ} (hx : 0 < x) (y : ℝ) : x ^ y ≠ 0 := by + rw [rpow_def_of_pos hx]; apply exp_ne_zero _ + open Real theorem rpow_def_of_neg {x : ℝ} (hx : x < 0) (y : ℝ) : x ^ y = exp (log x * y) * cos (y * π) := by diff --git a/Mathlib/Data/Complex/Order.lean b/Mathlib/Data/Complex/Order.lean index 47d0dd85378f3..ab22a62dfd954 100644 --- a/Mathlib/Data/Complex/Order.lean +++ b/Mathlib/Data/Complex/Order.lean @@ -63,6 +63,32 @@ theorem nonneg_iff {z : ℂ} : 0 ≤ z ↔ 0 ≤ z.re ∧ 0 = z.im := theorem pos_iff {z : ℂ} : 0 < z ↔ 0 < z.re ∧ 0 = z.im := lt_def +theorem nonpos_iff {z : ℂ} : z ≤ 0 ↔ z.re ≤ 0 ∧ z.im = 0 := + le_def + +theorem neg_iff {z : ℂ} : z < 0 ↔ z.re < 0 ∧ z.im = 0 := + lt_def + +theorem ofReal_mul_pos_iff (x : ℝ) (z : ℂ) : + 0 < x * z ↔ (x < 0 ∧ z < 0) ∨ (0 < x ∧ 0 < z) := by + simp_rw [pos_iff, neg_iff, re_ofReal_mul, im_ofReal_mul] + obtain hx | hx | hx := lt_trichotomy x 0 + · simp_rw [mul_pos_iff, eq_comm, mul_eq_zero, hx, not_lt_of_gt hx, hx.ne, false_and, true_and, + false_or, or_false] + · simp_rw [hx, zero_mul, lt_self_iff_false, false_and, false_or] + · simp_rw [mul_pos_iff, eq_comm, mul_eq_zero, hx, not_lt_of_gt hx, hx.ne', false_and, true_and, + false_or, or_false] + +theorem ofReal_mul_neg_iff (x : ℝ) (z : ℂ) : + x * z < 0 ↔ (x < 0 ∧ 0 < z) ∨ (0 < x ∧ z < 0) := by + simp_rw [pos_iff, neg_iff, re_ofReal_mul, im_ofReal_mul] + obtain hx | hx | hx := lt_trichotomy x 0 + · simp_rw [mul_neg_iff, mul_eq_zero, hx, not_lt_of_gt hx, hx.ne, false_and, true_and, false_or, + or_false, eq_comm] + · simp_rw [hx, zero_mul, lt_self_iff_false, false_and, false_or] + · simp_rw [mul_neg_iff, mul_eq_zero, hx, not_lt_of_gt hx, hx.ne', false_and, true_and, false_or, + or_false] + @[simp, norm_cast] theorem real_le_real {x y : ℝ} : (x : ℂ) ≤ (y : ℂ) ↔ x ≤ y := by simp [le_def, ofReal] diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index bccf679b634a5..50bffb2b3038c 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -192,6 +192,9 @@ theorem ofReal_lt_ofReal_iff_of_nonneg {p q : ℝ} (hp : 0 ≤ p) : @[simp] theorem ofReal_pos {p : ℝ} : 0 < ENNReal.ofReal p ↔ 0 < p := by simp [ENNReal.ofReal] +theorem ofReal_ne_zero_iff {r : ℝ} : ENNReal.ofReal r ≠ 0 ↔ 0 < r := by + rw [← zero_lt_iff, ENNReal.ofReal_pos] + @[bound] private alias ⟨_, Bound.ofReal_pos_of_pos⟩ := ofReal_pos @[simp] diff --git a/Mathlib/LinearAlgebra/Basis/Basic.lean b/Mathlib/LinearAlgebra/Basis/Basic.lean index 46ed8282bcd9a..ed7b353cf862b 100644 --- a/Mathlib/LinearAlgebra/Basis/Basic.lean +++ b/Mathlib/LinearAlgebra/Basis/Basic.lean @@ -362,6 +362,10 @@ theorem isUnitSMul_apply {v : Basis ι R M} {w : ι → R} (hw : ∀ i, IsUnit ( v.isUnitSMul hw i = w i • v i := unitsSMul_apply i +theorem repr_isUnitSMul {v : Basis ι R₂ M} {w : ι → R₂} (hw : ∀ i, IsUnit (w i)) (x : M) (i : ι) : + (v.isUnitSMul hw).repr x i = (hw i).unit⁻¹ • v.repr x i := + repr_unitsSMul _ _ _ _ + section Fin /-- Let `b` be a basis for a submodule `N` of `M`. If `y : M` is linear independent of `N` diff --git a/Mathlib/LinearAlgebra/Determinant.lean b/Mathlib/LinearAlgebra/Determinant.lean index dc9936ab62f2a..7d8b1a2d7531a 100644 --- a/Mathlib/LinearAlgebra/Determinant.lean +++ b/Mathlib/LinearAlgebra/Determinant.lean @@ -548,6 +548,15 @@ theorem Basis.det_comp_basis [Module A M'] (b : Basis ι A M) (b' : Basis ι A M congr 1; ext i j rw [Basis.toMatrix_apply, LinearMap.toMatrix_apply, Function.comp_apply] +theorem Basis.det_basis (b : Basis ι A M) (b' : Basis ι A M) : + b'.det b = LinearMap.det (b'.equiv b (Equiv.refl ι)).toLinearMap := + b.det_comp_basis b' (LinearMap.id) + +theorem Basis.det_inv (b : Basis ι A M) (b' : Basis ι A M) : + (b.isUnit_det b').unit⁻¹ = b'.det b := by + rw [← Units.mul_eq_one_iff_inv_eq, IsUnit.unit_spec, Basis.det_basis, Basis.det_basis] + exact LinearEquiv.det_mul_det_symm _ + theorem Basis.det_reindex {ι' : Type*} [Fintype ι'] [DecidableEq ι'] (b : Basis ι R M) (v : ι' → M) (e : ι ≃ ι') : (b.reindex e).det v = b.det (v ∘ e) := by rw [Basis.det_apply, Basis.toMatrix_reindex', det_reindexAlgEquiv, Basis.det_apply] @@ -574,6 +583,11 @@ theorem Pi.basisFun_det : (Pi.basisFun R ι).det = Matrix.detRowAlternating := b ext M rw [Basis.det_apply, Basis.coePiBasisFun.toMatrix_eq_transpose, det_transpose] +theorem Pi.basisFun_det_apply (v : ι → ι → R) : + (Pi.basisFun R ι).det v = (Matrix.of v).det := by + rw [Pi.basisFun_det] + rfl + /-- If we fix a background basis `e`, then for any other basis `v`, we can characterise the coordinates provided by `v` in terms of determinants relative to `e`. -/ theorem Basis.det_smul_mk_coord_eq_det_update {v : ι → M} (hli : LinearIndependent R v) diff --git a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean index a09f3bc5eb9bd..b2c7caa1a6f54 100644 --- a/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/BorelSpace/Basic.lean @@ -294,6 +294,10 @@ theorem IsCompact.measure_closure [R1Space γ] {K : Set γ} (hK : IsCompact K) ( theorem measurableSet_closure : MeasurableSet (closure s) := isClosed_closure.measurableSet +@[measurability] +theorem measurableSet_frontier : MeasurableSet (frontier s) := + measurableSet_closure.diff measurableSet_interior + theorem measurable_of_isOpen {f : δ → γ} (hf : ∀ s, IsOpen s → MeasurableSet (f ⁻¹' s)) : Measurable f := by rw [‹BorelSpace γ›.measurable_eq] diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 40341ab6df349..4f89b384208ea 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -435,6 +435,12 @@ theorem pi_hyperplane (i : ι) [NoAtoms (μ i)] (x : α i) : theorem ae_eval_ne (i : ι) [NoAtoms (μ i)] (x : α i) : ∀ᵐ y : ∀ i, α i ∂Measure.pi μ, y i ≠ x := compl_mem_ae_iff.2 (pi_hyperplane μ i x) +theorem restrict_pi_pi (s : (i : ι) → Set (α i)) : + (Measure.pi μ).restrict (Set.univ.pi fun i ↦ s i) = .pi (fun i ↦ (μ i).restrict (s i)) := by + refine (pi_eq fun _ h ↦ ?_).symm + simp_rw [restrict_apply (MeasurableSet.univ_pi h), restrict_apply (h _), + ← Set.pi_inter_distrib, pi_pi] + variable {μ} theorem tendsto_eval_ae_ae {i : ι} : Tendsto (eval i) (ae (Measure.pi μ)) (ae (μ i)) := fun _ hs => @@ -753,6 +759,34 @@ theorem volume_measurePreserving_piCongrLeft (α : ι → Type*) (f : ι' ≃ ι MeasurePreserving (MeasurableEquiv.piCongrLeft α f) volume volume := measurePreserving_piCongrLeft (fun _ ↦ volume) f +theorem measurePreserving_arrowProdEquivProdArrow (α β γ : Type*) [MeasurableSpace α] + [MeasurableSpace β] [Fintype γ] (μ : γ → Measure α) (ν : γ → Measure β) [∀ i, SigmaFinite (μ i)] + [∀ i, SigmaFinite (ν i)] : + MeasurePreserving (MeasurableEquiv.arrowProdEquivProdArrow α β γ) + (.pi fun i ↦ (μ i).prod (ν i)) + ((Measure.pi fun i ↦ μ i).prod (Measure.pi fun i ↦ ν i)) where + measurable := (MeasurableEquiv.arrowProdEquivProdArrow α β γ).measurable + map_eq := by + refine (FiniteSpanningSetsIn.ext ?_ (isPiSystem_pi.prod isPiSystem_pi) + ((FiniteSpanningSetsIn.pi fun i ↦ (μ i).toFiniteSpanningSetsIn).prod + (FiniteSpanningSetsIn.pi (fun i ↦ (ν i).toFiniteSpanningSetsIn))) ?_).symm + · refine (generateFrom_eq_prod generateFrom_pi generateFrom_pi ?_ ?_).symm + exact (FiniteSpanningSetsIn.pi (fun i ↦ (μ i).toFiniteSpanningSetsIn)).isCountablySpanning + exact (FiniteSpanningSetsIn.pi (fun i ↦ (ν i).toFiniteSpanningSetsIn)).isCountablySpanning + · rintro _ ⟨s, ⟨s, _, rfl⟩, ⟨_, ⟨t, _, rfl⟩, rfl⟩⟩ + rw [MeasurableEquiv.map_apply, MeasurableEquiv.arrowProdEquivProdArrow, + MeasurableEquiv.coe_mk] + rw [show Equiv.arrowProdEquivProdArrow α β γ ⁻¹' (univ.pi s ×ˢ univ.pi t) = + (univ.pi fun i ↦ s i ×ˢ t i) by + ext; simp [Equiv.arrowProdEquivProdArrow, Equiv.coe_fn_mk, Set.mem_pi, forall_and]] + simp_rw [pi_pi, prod_prod, pi_pi, Finset.prod_mul_distrib] + +theorem volume_measurePreserving_arrowProdEquivProdArrow (α β γ : Type*) [MeasureSpace α] + [MeasureSpace β] [Fintype γ] [SigmaFinite (volume : Measure α)] + [SigmaFinite (volume : Measure β)] : + MeasurePreserving (MeasurableEquiv.arrowProdEquivProdArrow α β γ) := + measurePreserving_arrowProdEquivProdArrow α β γ (fun _ ↦ volume) (fun _ ↦ volume) + theorem measurePreserving_sumPiEquivProdPi_symm {π : ι ⊕ ι' → Type*} {m : ∀ i, MeasurableSpace (π i)} (μ : ∀ i, Measure (π i)) [∀ i, SigmaFinite (μ i)] : MeasurePreserving (MeasurableEquiv.sumPiEquivProdPi π).symm @@ -882,7 +916,7 @@ theorem volume_preserving_piFinsetUnion {ι : Type*} [DecidableEq ι] (α : ι measurePreserving_piFinsetUnion h (fun _ ↦ volume) theorem measurePreserving_pi {ι : Type*} [Fintype ι] {α : ι → Type v} {β : ι → Type*} - [∀ i, MeasureSpace (α i)] [∀ i, MeasurableSpace (β i)] + [∀ i, MeasurableSpace (α i)] [∀ i, MeasurableSpace (β i)] (μ : (i : ι) → Measure (α i)) (ν : (i : ι) → Measure (β i)) {f : (i : ι) → (α i) → (β i)} [∀ i, SigmaFinite (ν i)] (hf : ∀ i, MeasurePreserving (f i) (μ i) (ν i)) : @@ -903,6 +937,28 @@ theorem volume_preserving_pi {α' β' : ι → Type*} [∀ i, MeasureSpace (α' MeasurePreserving (fun (a : (i : ι) → α' i) (i : ι) ↦ (f i) (a i)) := measurePreserving_pi _ _ hf +/-- Docstring. -/ +theorem measurePreserving_arrowCongr' {α₁ β₁ α₂ β₂ : Type*} [Fintype α₁] [Fintype α₂] + [MeasurableSpace β₁] [MeasurableSpace β₂] (μ : α₁ → Measure β₁) (ν : α₂ → Measure β₂) + [∀ i, SigmaFinite (ν i)] (hα : α₁ ≃ α₂) (hβ : β₁ ≃ᵐ β₂) + (hm : ∀ i, MeasurePreserving hβ (μ i) (ν (hα i))) : + MeasurePreserving (MeasurableEquiv.arrowCongr' hα hβ) (Measure.pi fun i ↦ μ i) + (Measure.pi fun i ↦ ν i) := by + classical + convert (measurePreserving_piCongrLeft (fun i : α₂ ↦ ν i) hα).comp + (measurePreserving_pi μ (fun i : α₁ ↦ ν (hα i)) hm) + simp only [MeasurableEquiv.arrowCongr', Equiv.arrowCongr', Equiv.arrowCongr, EquivLike.coe_coe, + MeasurableEquiv.coe_mk, Equiv.coe_fn_mk, MeasurableEquiv.piCongrLeft, Equiv.piCongrLeft, + Equiv.symm_symm_apply, Equiv.piCongrLeft'_symm, Equiv.symm_symm] + rfl + +/-- Docstring. -/ +theorem volume_preserving_arrowCongr' {α₁ β₁ α₂ β₂ : Type*} [Fintype α₁] [Fintype α₂] + [MeasureSpace β₁] [MeasureSpace β₂] [SigmaFinite (volume : Measure β₂)] + (hα : α₁ ≃ α₂) (hβ : β₁ ≃ᵐ β₂) (hm : MeasurePreserving hβ) : + MeasurePreserving (MeasurableEquiv.arrowCongr' hα hβ) := + measurePreserving_arrowCongr' (fun _ ↦ volume) (fun _ ↦ volume) hα hβ (fun _ ↦ hm) + end MeasurePreserving end MeasureTheory diff --git a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean index a41e36a07b43e..84bd2c1f4e915 100644 --- a/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean +++ b/Mathlib/MeasureTheory/Constructions/Prod/Basic.lean @@ -723,6 +723,11 @@ theorem restrict_prod_eq_prod_univ (s : Set α) : have : ν = ν.restrict Set.univ := Measure.restrict_univ.symm rw [this, Measure.prod_restrict, ← this] +theorem restrict_prod_eq_univ_prod (t : Set β) : + μ.prod (ν.restrict t) = (μ.prod ν).restrict (Set.univ ×ˢ t) := by + have : μ = μ.restrict Set.univ := Measure.restrict_univ.symm + rw [this, Measure.prod_restrict, ← this] + theorem prod_dirac (y : β) : μ.prod (dirac y) = map (fun x => (x, y)) μ := by classical rw [← sum_sfiniteSeq μ, prod_sum_left, map_sum measurable_prod_mk_right.aemeasurable] @@ -827,6 +832,33 @@ protected theorem prod [SFinite μa] [SFinite μc] {f : α → β} {g : γ → have : Measurable (uncurry fun _ : α => g) := hg.1.comp measurable_snd hf.skew_product this <| ae_of_all _ fun _ => hg.map_eq + +variable (μa μb μc) in +/-- The measurable equiv induced by the equiv `(α × β) × γ ≃ α × (β × γ)` is measure preserving. -/ +theorem _root_.MeasureTheory.measurePreserving_prodAssoc [SigmaFinite μa] [SigmaFinite μb] + [SigmaFinite μc] : + MeasurePreserving (MeasurableEquiv.prodAssoc : (α × β) × γ ≃ᵐ α × β × γ) + ((μa.prod μb).prod μc) (μa.prod (μb.prod μc)) where + measurable := MeasurableEquiv.prodAssoc.measurable + map_eq := by + refine (FiniteSpanningSetsIn.ext ?_ + (isPiSystem_measurableSet.prod (isPiSystem_measurableSet.prod isPiSystem_measurableSet)) + (μa.toFiniteSpanningSetsIn.prod (μb.toFiniteSpanningSetsIn.prod μc.toFiniteSpanningSetsIn)) + ?_).symm + · refine (generateFrom_eq_prod generateFrom_measurableSet + (generateFrom_eq_prod ?_ ?_ ?_ ?_) ?_ (IsCountablySpanning.prod ?_ ?_)).symm + any_goals exact generateFrom_measurableSet + all_goals exact isCountablySpanning_measurableSet + · rintro _ ⟨s, _, _, ⟨t, _, ⟨u, _, rfl⟩⟩, rfl⟩ + rw [MeasurableEquiv.map_apply, MeasurableEquiv.prodAssoc, MeasurableEquiv.coe_mk, + Equiv.prod_assoc_preimage, prod_prod, prod_prod, prod_prod, prod_prod, mul_assoc] + +theorem _root_.MeasureTheory.volume_preserving_prodAssoc {α₁ β₁ γ₁ : Type*} [MeasureSpace α₁] + [MeasureSpace β₁] [MeasureSpace γ₁] [SigmaFinite (volume : Measure α₁)] + [SigmaFinite (volume : Measure β₁)] [SigmaFinite (volume : Measure γ₁)] : + MeasurePreserving (MeasurableEquiv.prodAssoc : (α₁ × β₁) × γ₁ ≃ᵐ α₁ × β₁ × γ₁) := + MeasureTheory.measurePreserving_prodAssoc volume volume volume + end MeasurePreserving namespace QuasiMeasurePreserving diff --git a/Mathlib/MeasureTheory/Integral/Marginal.lean b/Mathlib/MeasureTheory/Integral/Marginal.lean index 76fcaa682df92..991f1e8fed928 100644 --- a/Mathlib/MeasureTheory/Integral/Marginal.lean +++ b/Mathlib/MeasureTheory/Integral/Marginal.lean @@ -94,6 +94,10 @@ theorem _root_.Measurable.lmarginal [∀ i, SigmaFinite (μ i)] (hf : Measurable · simpa [hi, updateFinset] using measurable_pi_iff.1 measurable_snd _ · simpa [hi, updateFinset] using measurable_pi_iff.1 measurable_fst _ +theorem _root_.Measurable.lmarginal_update [∀ i, SigmaFinite (μ i)] (hf : Measurable f) : + Measurable fun xᵢ ↦ (∫⋯∫⁻_s, f ∂μ) (Function.update x i xᵢ) := + (Measurable.lmarginal _ hf).comp (measurable_update x) + @[simp] theorem lmarginal_empty (f : (∀ i, π i) → ℝ≥0∞) : ∫⋯∫⁻_∅, f ∂μ = f := by ext1 x simp_rw [lmarginal, Measure.pi_of_empty fun i : (∅ : Finset δ) => μ i] @@ -131,6 +135,18 @@ variable {μ} in theorem lmarginal_mono {f g : (∀ i, π i) → ℝ≥0∞} (hfg : f ≤ g) : ∫⋯∫⁻_s, f ∂μ ≤ ∫⋯∫⁻_s, g ∂μ := fun _ => lintegral_mono fun _ => hfg _ +variable {μ} in +theorem lmarginal_const_smul (hf : Measurable f) (r : ENNReal) : + (∫⋯∫⁻_s, r • f ∂μ) x = r * (∫⋯∫⁻_s, f ∂μ) x := by + simp_rw [lmarginal, Pi.smul_apply, smul_eq_mul] + rw [lintegral_const_mul _ (by convert hf.comp measurable_updateFinset)] + +variable {μ} in +theorem lmarginal_const_smul' (r : ENNReal) (hr : r ≠ ⊤): + (∫⋯∫⁻_s, r • f ∂μ) x = r * (∫⋯∫⁻_s, f ∂μ) x := by + simp_rw [lmarginal, Pi.smul_apply, smul_eq_mul] + rw [lintegral_const_mul' _ _ hr] + variable [∀ i, SigmaFinite (μ i)] theorem lmarginal_union (f : (∀ i, π i) → ℝ≥0∞) (hf : Measurable f) diff --git a/Mathlib/MeasureTheory/Integral/Pi.lean b/Mathlib/MeasureTheory/Integral/Pi.lean index 3445cc3f930f1..30d99f281b0e8 100644 --- a/Mathlib/MeasureTheory/Integral/Pi.lean +++ b/Mathlib/MeasureTheory/Integral/Pi.lean @@ -3,24 +3,34 @@ Copyright (c) 2023 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ +import Mathlib.Analysis.SpecialFunctions.PolarCoord import Mathlib.MeasureTheory.Constructions.Pi import Mathlib.MeasureTheory.Constructions.Prod.Integral /-! # Integration with respect to a finite product of measures +## Fubini's theorem + On a finite product of measure spaces, we show that a product of integrable functions each depending on a single coordinate is integrable, in `MeasureTheory.integrable_fintype_prod`, and that its integral is the product of the individual integrals, in `MeasureTheory.integral_fintype_prod_eq_prod`. + +## Polar coordinates change of variables + + + -/ open Fintype MeasureTheory MeasureTheory.Measure -variable {𝕜 : Type*} [RCLike 𝕜] - namespace MeasureTheory +section Fubini + +variable {𝕜 : Type*} [RCLike 𝕜] + /-- On a finite product space in `n` variables, for a natural number `n`, a product of integrable functions depending on each coordinate is integrable. -/ theorem Integrable.fin_nat_prod {n : ℕ} {E : Fin n → Type*} @@ -97,4 +107,60 @@ theorem integral_fintype_prod_eq_pow {E : Type*} (ι : Type*) [Fintype ι] (f : ∫ x : ι → E, ∏ i, f (x i) = (∫ x, f x) ^ (card ι) := by rw [integral_fintype_prod_eq_prod, Finset.prod_const, card] +end Fubini + end MeasureTheory + +section polarCoord + +variable {ι : Type*} [DecidableEq ι] (f : (ι → ℂ) → ENNReal) + +private theorem Complex.lintegral_pi_comp_polarCoord_symm_aux (hf : Measurable f) (s : Finset ι) + (a : ι → ℝ × ℝ) : + (∫⋯∫⁻_s, f ∂fun _ ↦ (volume : Measure ℂ)) (fun i ↦ Complex.polarCoord.symm (a i)) = + (∫⋯∫⁻_s, fun p ↦ + ((∏ i ∈ s, (p i).1.toNNReal) * f (fun i ↦ Complex.polarCoord.symm (p i))) + ∂fun _ ↦ ((volume : Measure (ℝ × ℝ)).restrict polarCoord.target)) a := by + induction s using Finset.induction generalizing f a with + | empty => simp + | @insert i₀ s hi₀ h_ind => + have h : ∀ t : Finset ι, Measurable fun p : ι → ℝ × ℝ ↦ + (∏ i ∈ t, (p i).1.toNNReal) * f fun i ↦ Complex.polarCoord.symm (p i) := by + intro _ + refine Measurable.mul ?_ ?_ + · exact measurable_coe_nnreal_ennreal_iff.mpr <| + Finset.measurable_prod _ fun _ _ ↦ by fun_prop + · exact hf.comp <| measurable_pi_lambda _ fun _ ↦ + Complex.continuous_polarCoord_symm.measurable.comp (measurable_pi_apply _) + calc + _ = ∫⁻ x in polarCoord.target, x.1.toNNReal • + (∫⋯∫⁻_s, f ∂fun _ ↦ volume) + fun j ↦ Complex.polarCoord.symm (Function.update a i₀ x j) := by + rw [MeasureTheory.lmarginal_insert _ hf hi₀, ← Complex.lintegral_comp_polarCoord_symm _ + (hf.lmarginal_update (fun _ ↦ (volume : Measure ℂ)))] + congr! + simp_rw [Function.update_apply] + split_ifs <;> rfl + _ = ∫⁻ (x : ℝ × ℝ) in polarCoord.target, + (∫⋯∫⁻_s, + (fun p ↦ ↑(∏ i ∈ insert i₀ s, (p i).1.toNNReal) * + (f fun i ↦ Complex.polarCoord.symm (p i))) ∘ fun p ↦ Function.update p i₀ x + ∂fun _ ↦ volume.restrict polarCoord.target) a := by + simp_rw [h_ind _ hf, lmarginal_update_of_not_mem (h s) hi₀, Function.comp_def, + ENNReal.smul_def, smul_eq_mul, ← lmarginal_const_smul' _ ENNReal.coe_ne_top, + Pi.smul_def, Finset.prod_insert hi₀, Function.update_same, smul_eq_mul, + ENNReal.coe_mul, mul_assoc] + _ = (∫⋯∫⁻_insert i₀ s, fun p ↦ (∏ i ∈ insert i₀ s, (p i).1.toNNReal) * + f (fun i ↦ Complex.polarCoord.symm (p i)) + ∂fun _ ↦ volume.restrict polarCoord.target) a := by + simp_rw [← lmarginal_update_of_not_mem (h _) hi₀] + rw [MeasureTheory.lmarginal_insert _ (h _) hi₀] + +theorem Complex.lintegral_pi_comp_polarCoord_symm [Fintype ι] (hf : Measurable f) : + ∫⁻ p, f p = ∫⁻ p in (Set.univ.pi fun _ : ι ↦ polarCoord.target), + (∏ i, (p i).1.toNNReal) * f (fun i ↦ Complex.polarCoord.symm (p i)) := by + rw [volume_pi, lintegral_eq_lmarginal_univ (fun _ ↦ Complex.polarCoord.symm 0), + Complex.lintegral_pi_comp_polarCoord_symm_aux _ hf, lmarginal_univ, ← restrict_pi_pi] + rfl + +end polarCoord diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean index 67ab6dc3c5625..8d6e2a402cc65 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Embedding.lean @@ -468,6 +468,32 @@ lemma piCongrLeft_apply_apply {ι ι' : Type*} (e : ι ≃ ι') {β : ι' → Ty piCongrLeft (fun i' ↦ β i') e x (e i) = x i := by rw [piCongrLeft, coe_mk, Equiv.piCongrLeft_apply_apply] +/-- The isomorphism `(γ → α × β) ≃ (γ → α) × (γ → β)` as a measurable equivalence. -/ +def arrowProdEquivProdArrow (α β γ : Type*) [MeasurableSpace α] + [MeasurableSpace β] : + (γ → α × β) ≃ᵐ (γ → α) × (γ → β) where + __ := Equiv.arrowProdEquivProdArrow α β γ + measurable_toFun _ h := by + simp_rw [Equiv.arrowProdEquivProdArrow, coe_fn_mk] + exact MeasurableSet.preimage h (Measurable.prod_mk + (measurable_pi_lambda (fun a c ↦ (a c).1) fun a ↦ (measurable_pi_apply a).fst) + (measurable_pi_lambda (fun a c ↦ (a c).2) fun a ↦ (measurable_pi_apply a).snd )) + measurable_invFun _ h := by + simp_rw [Equiv.arrowProdEquivProdArrow, coe_fn_symm_mk] + exact MeasurableSet.preimage h (by measurability) + +/-- The measurable equivalence `(α₁ → β₁) ≃ᵐ (α₂ → β₂)` induced by `α₁ ≃ α₂` and `β₁ ≃ᵐ β₂`. -/ +def arrowCongr' {α₁ β₁ α₂ β₂ : Type*} [MeasurableSpace β₁] [MeasurableSpace β₂] + (hα : α₁ ≃ α₂) (hβ : β₁ ≃ᵐ β₂) : + (α₁ → β₁) ≃ᵐ (α₂ → β₂) where + __ := Equiv.arrowCongr' hα hβ + measurable_toFun _ h := by + exact MeasurableSet.preimage h <| + measurable_pi_iff.mpr fun _ ↦ hβ.measurable.comp' (measurable_pi_apply _) + measurable_invFun _ h := by + exact MeasurableSet.preimage h <| + measurable_pi_iff.mpr fun _ ↦ hβ.symm.measurable.comp' (measurable_pi_apply _) + /-- Pi-types are measurably equivalent to iterated products. -/ @[simps! (config := .asFn)] def piMeasurableEquivTProd [DecidableEq δ'] {l : List δ'} (hnd : l.Nodup) (h : ∀ i, i ∈ l) : diff --git a/Mathlib/NumberTheory/LSeries/Basic.lean b/Mathlib/NumberTheory/LSeries/Basic.lean index 7221b06b0360f..c476dbab384cd 100644 --- a/Mathlib/NumberTheory/LSeries/Basic.lean +++ b/Mathlib/NumberTheory/LSeries/Basic.lean @@ -91,6 +91,11 @@ lemma term_congr {f g : ℕ → ℂ} (h : ∀ {n}, n ≠ 0 → f n = g n) (s : term f s n = term g s n := by rcases eq_or_ne n 0 with hn | hn <;> simp [hn, h] +theorem term_eq_coe (f : ℕ → ℝ) (s : ℝ) (n : ℕ) : + term (fun n ↦ f n) s n = ↑(if n = 0 then 0 else f n / (n : ℝ) ^ s) := by + rw [LSeries.term_def, apply_ite ofReal, ofReal_zero, ofReal_div, ofReal_cpow (Nat.cast_nonneg _), + ofReal_natCast] + lemma norm_term_eq (f : ℕ → ℂ) (s : ℂ) (n : ℕ) : ‖term f s n‖ = if n = 0 then 0 else ‖f n‖ / n ^ s.re := by rcases eq_or_ne n 0 with rfl | hn diff --git a/Mathlib/NumberTheory/LSeries/Residue.lean b/Mathlib/NumberTheory/LSeries/Residue.lean new file mode 100644 index 0000000000000..a56dd53225e7f --- /dev/null +++ b/Mathlib/NumberTheory/LSeries/Residue.lean @@ -0,0 +1,330 @@ +/- +Copyright (c) 2024 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.NumberTheory.LSeries.RiemannZeta + +/-! +# Docstring + +-/ + +open Filter Topology + +namespace NumberTheory.LSeries.residueFormula + +noncomputable section + +variable {a : ℕ → ℕ} {l : ℝ} + +variable (a) in +/-- Docstring. -/ +abbrev A (n : ℕ) : ℕ := ∑ i ∈ Finset.range (n + 1), a i + +variable (a) in +theorem monotone_A : Monotone (A a) := by + intro x y h + rw [A, A, ← Finset.sum_range_add_sum_Ico _ ( Nat.add_le_add_right h 1)] + exact Nat.le_add_right _ _ + +variable (a) in +/-- Docstring. -/ +def u (n : ℕ) : ℕ := sInf {k : ℕ | n ≤ A a k} + +theorem A_u_lt {n : ℕ} (hn : 0 < u a n) : A a ((u a n) - 1) < n := by + by_contra! h + have := csInf_le' (by exact h : (u a n) - 1 ∈ {k : ℕ | n ≤ A a k}) + exact (lt_irrefl _) <| (Nat.le_sub_one_iff_lt hn).mp this + +theorem tendsto_mul_sum_rpow (T : Finset ℕ) (v : ℕ → ℕ) : + Tendsto (fun s ↦ (s - 1) * ∑ n ∈ T, (v n : ℝ) ^ (- s)) (𝓝[>] 1) (𝓝 0) := by + have h₀ : Tendsto (fun s : ℝ ↦ (s - 1) * (0 : ℝ) ^ (- s)) (𝓝[>] 1) (𝓝 0) := by + refine Tendsto.congr' (eventuallyEq_nhdsWithin_of_eqOn fun s hs ↦ ?_) tendsto_const_nhds + rw [Real.zero_rpow (neg_ne_zero.mpr (lt_trans zero_lt_one hs).ne'), mul_zero] + have : ∀ n ∈ T, Tendsto (fun s ↦ (s - 1) * (v n : ℝ) ^ (- s)) (𝓝[>] 1) (𝓝 0) := by + intro n _ + by_cases hv : v n = 0 + · simp_rw [hv, Nat.cast_zero, h₀] + · rw [show 0 = 0 * (v n : ℝ) ^ (- 1 : ℝ) by rw [zero_mul]] + refine tendsto_nhdsWithin_of_tendsto_nhds (Tendsto.mul ?_ (Continuous.tendsto ?_ 1)) + · convert (continuous_sub_right (1 : ℝ)).tendsto 1 + rw [sub_self] + · exact continuous_const.rpow continuous_neg fun _ ↦ Or.inl (Nat.cast_ne_zero.mpr hv) + convert tendsto_finset_sum _ this + · rw [Finset.mul_sum] + · rw [Finset.sum_const_zero] + +theorem tendsto_rpow_mul_tsum_rpow {c : ℝ} (hc : c ≠ 0) (T : Finset ℕ) : + Tendsto (fun s ↦ c ^ s * (s - 1) * + ∑' n : ↑((T : Set ℕ)ᶜ), (n : ℝ) ^ (- s)) (𝓝[>] 1) (𝓝 c) := by + simp_rw [mul_assoc, show 𝓝 c = 𝓝 (c * (1 - 0)) by rw [sub_zero, mul_one]] + refine Tendsto.mul (tendsto_nhdsWithin_of_tendsto_nhds ?_) ?_ + · refine Continuous.tendsto' ?_ 1 c (by rw [Real.rpow_one]) + exact continuous_const.rpow continuous_id fun _ ↦ Or.inl hc + · refine (riemannZeta_residue_one'.sub (tendsto_mul_sum_rpow T (fun n ↦ n))).congr' ?_ + filter_upwards [eventually_mem_nhdsWithin] with s hs + simp_rw [sub_eq_iff_eq_add', ← mul_add, sum_add_tsum_compl (Real.summable_nat_rpow.mpr + (neg_lt_neg_iff.mpr hs)), Real.rpow_neg (Nat.cast_nonneg _), one_div] + +variable (hl : 0 < l) + (hA₁ : Tendsto (fun n ↦ (∑ i ∈ Finset.range (n + 1), a i : ℝ) / n) atTop (𝓝 l)) + +include hl hA₁ + +theorem tends_atTop_A : Tendsto (A a) atTop atTop := by + have : Tendsto (fun n ↦ (A a n : ℝ)) atTop atTop := by + have : Tendsto (fun n : ℕ ↦ l * (n : ℝ)) atTop atTop := by + refine Tendsto.const_mul_atTop hl tendsto_natCast_atTop_atTop + refine Asymptotics.IsEquivalent.tendsto_atTop ?_ this + rw [Asymptotics.isEquivalent_comm, Asymptotics.isEquivalent_iff_tendsto_one] + convert Tendsto.mul hA₁ (tendsto_const_nhds (x := l⁻¹)) + · simp + ring + · rw [mul_inv_cancel₀ hl.ne'] + · filter_upwards [eventually_ne_atTop 0] with n hn + refine mul_ne_zero hl.ne' (Nat.cast_ne_zero.mpr hn) + exact tendsto_natCast_atTop_iff.mp this + +theorem le_A_u (n : ℕ) : n ≤ A a (u a n) := by + have : {k : ℕ | n ≤ A a k}.Nonempty := by + have := tendsto_atTop_atTop.mp (tends_atTop_A hl hA₁) n + exact ⟨this.choose, this.choose_spec this.choose le_rfl⟩ + have := csInf_mem this + exact this + +theorem card_u_eq {n : ℕ} (hn : 0 < n) : Nat.card {k | u a k = n} = a n := by + have : {k | u a k = n} = Finset.Ioc (A a (n - 1)) (A a n) := by + ext x + rw [Set.mem_setOf_eq, Finset.coe_Ioc, Set.mem_Ioc] + refine ⟨?_, ?_⟩ + · intro h + rw [← h] + refine ⟨A_u_lt (h ▸ hn), le_A_u hl hA₁ x⟩ + · intro h + refine IsLeast.csInf_eq ⟨?_, ?_⟩ + exact h.2 + intro y hy + simp at hy + have := lt_of_lt_of_le h.1 hy + contrapose! this + rw [Nat.lt_iff_le_pred hn] at this + exact monotone_A a this + simp_rw [this, Nat.card_eq_card_toFinset, Finset.coe_Ioc, Set.toFinset_Ioc, Nat.card_Ioc, A] + rw [Finset.sum_range_succ, Nat.sub_one_add_one_eq_of_pos hn, Nat.add_sub_cancel_left] + +theorem monotone_u : Monotone (u a) := by + intro x y h + exact le_csInf ⟨u a y, le_A_u hl hA₁ y⟩ fun _ h' ↦ csInf_le (OrderBot.bddBelow _) (h.trans h') + +theorem tendsto_atTop_u : Tendsto (u a) atTop atTop := by + refine Monotone.tendsto_atTop_atTop (monotone_u hl hA₁) ?_ + by_contra! h + obtain ⟨B, hB⟩ := h + have : ∀ n, n ≤ A a B := by + intro n + have t₀ := le_A_u hl hA₁ n + have t₁ := monotone_A a (hB n) + have t₃ := monotone_A a (by exact Nat.le_add_right (u a n) 1 : u a n ≤ u a n + 1) + exact t₀.trans (t₃.trans t₁) + specialize this (A a B + 1) + simp only [add_le_iff_nonpos_right, nonpos_iff_eq_zero, one_ne_zero] at this + +theorem tendsto_atTop_u_div : Tendsto (fun n : ℕ ↦ (n : ℝ) / (u a n)) atTop (𝓝 l) := by + have h₁ : Tendsto (fun n ↦ (A a (u a n) : ℝ)/ (u a n)) atTop (𝓝 l) := by + convert hA₁.comp (tendsto_atTop_u hl hA₁) + simp + have h₂ : Tendsto (fun n ↦ ((A a (u a n - 1) : ℝ) / (u a n - 1 : ℕ)) * ((u a n - 1) / u a n)) + atTop (𝓝 l) := by + have : Tendsto (fun n ↦ n - 1) atTop atTop := by exact tendsto_sub_atTop_nat 1 + have := hA₁.comp this + have := this.comp (tendsto_atTop_u hl hA₁) + simp [Function.comp_def] at this + rw [show 𝓝 l = 𝓝 (l * 1) by ring_nf] + simp_rw [← Nat.cast_sum] at this + refine Tendsto.mul this ?_ + have : Tendsto (fun n : ℕ ↦ (n - 1 : ℝ) / n) atTop (𝓝 1) := by + have : Tendsto (fun n : ℕ ↦ (n : ℝ) / (n + 1)) atTop (𝓝 1) := tendsto_natCast_div_add_atTop 1 + have := this.comp (tendsto_sub_atTop_nat 1) + simp [Function.comp_def] at this + refine Tendsto.congr' ?_ this + filter_upwards [eventually_ge_atTop 1] with n hn + rw [Nat.cast_sub hn, Nat.cast_one, sub_add_cancel] + have := this.comp (tendsto_atTop_u hl hA₁) + exact this + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' h₂ h₁ ?_ ?_ + · have := tendsto_atTop_u hl hA₁ + rw [tendsto_atTop] at this + filter_upwards [this 2] with n hn + rw [Nat.cast_sub, Nat.cast_one, ← mul_div_assoc, div_mul_cancel₀] + · refine div_le_div_of_nonneg_right ?_ ?_ + · rw [Nat.cast_le] + exact (A_u_lt (lt_of_lt_of_le zero_lt_two hn)).le + · exact Nat.cast_nonneg _ + · refine sub_ne_zero_of_ne ?_ + refine LT.lt.ne' ?_ + rw [Nat.one_lt_cast] + exact lt_of_lt_of_le one_lt_two hn + · exact le_trans one_le_two hn + · filter_upwards with n + refine div_le_div_of_nonneg_right ?_ ?_ + · rw [Nat.cast_le] + exact le_A_u hl hA₁ n + · exact Nat.cast_nonneg _ + +theorem lt_u_rpow_lt {ε : ℝ} (hε₁ : 0 < ε) (hε₂ : ε ≤ l) : + ∀ᶠ n : ℕ in atTop, ∀ s : ℝ, 0 < s → (l - ε) ^ s * (n : ℝ) ^ (- s) < u a n ^ (- s) ∧ + u a n ^ (- s) < (l + ε) ^ s * (n : ℝ) ^ (- s) := by + rw [← sub_nonneg] at hε₂ -- To help positivity + filter_upwards [eventually_gt_atTop 0, Metric.tendsto_nhds.mp (tendsto_atTop_u_div hl hA₁) ε hε₁] + with _ _ h + simp_rw [Real.rpow_neg (Nat.cast_nonneg _), ← Real.inv_rpow (Nat.cast_nonneg _)] + intro s hs + rw [← Real.mul_rpow, ← Real.mul_rpow, Real.rpow_lt_rpow_iff, Real.rpow_lt_rpow_iff, + mul_inv_lt_iff₀, lt_mul_inv_iff₀, ← neg_add_lt_iff_lt_add, sub_eq_add_neg, + ← lt_neg_add_iff_add_lt (a := l), neg_add_eq_sub, ← abs_lt, mul_comm] + exact h + all_goals positivity + +theorem summable_u_rpow {s : ℝ} (hs : 1 < s) : + Summable (fun n ↦ (u a n : ℝ) ^ (- s)) := by + have : Summable (fun n : ℕ ↦ (l + l) ^ s * (n : ℝ) ^ (- s)) := by + refine Summable.mul_left _ ?_ + rw [Real.summable_nat_rpow] + exact neg_lt_neg_iff.mpr hs + refine summable_of_isBigO this ?_ + rw [Nat.cofinite_eq_atTop] + have := lt_u_rpow_lt (ε := l) hl hA₁ hl le_rfl + refine Eventually.isBigO ?_ + filter_upwards [this] with n hn + rw [Real.norm_eq_abs, abs_of_nonneg] + exact (hn s (lt_trans zero_lt_one hs)).2.le + refine Real.rpow_nonneg ?_ _ + exact Nat.cast_nonneg _ + +theorem exist_finset_lt_tsum_u_lt {ε : ℝ} (hε₁ : 0 < ε) (hε₂ : ε ≤ l) : + ∃ T : Finset ℕ, ∀ s, 1 < s → + (s - 1) * ∑ n ∈ T, (u a n : ℝ) ^ (- s) + + (l - ε) ^ s * (s - 1) * ∑' n : ↑((T : Set ℕ)ᶜ), (n : ℝ) ^ (- s) < + (s - 1) * ∑' n, (u a n : ℝ) ^ (-s) ∧ + (s - 1) * ∑' n, (u a n : ℝ) ^ (-s) < + (s - 1) * ∑ n ∈ T, (u a n : ℝ) ^ (- s) + + (l + ε) ^ s * (s - 1) * ∑' n : ↑((T : Set ℕ)ᶜ), (n : ℝ) ^ (- s) := by + obtain ⟨N, hN⟩ := eventually_atTop.mp <| lt_u_rpow_lt hl hA₁ hε₁ hε₂ + refine ⟨Finset.range N, fun s hs ↦ ?_⟩ + simp_rw [← sum_add_tsum_compl (s := Finset.range N) (summable_u_rpow hl hA₁ hs), mul_add, + add_lt_add_iff_left, mul_assoc, mul_left_comm _ (s- 1), mul_lt_mul_left (sub_pos.mpr hs), + ← tsum_mul_left] + have h₁ : ∀ (S : Set ℕ) (c : ℝ), Summable fun n : S ↦ c * (n : ℝ) ^ (-s) := fun S c ↦ by + have : Summable fun n : ℕ ↦ c * (n : ℝ) ^ (- s) := by + refine Summable.mul_left _ ?_ + rw [Real.summable_nat_rpow] + rwa [neg_lt_neg_iff] + exact (summable_subtype_and_compl.mpr this).1 + have h₂ : ∀ (S : Set ℕ), Summable fun n : S ↦ (u a n : ℝ) ^ (-s) := + fun S ↦ (summable_subtype_and_compl.mpr (summable_u_rpow hl hA₁ hs)).1 + refine ⟨tsum_lt_tsum (i := ⟨N+1, by simp⟩) ?_ ?_ (h₁ _ ((l - ε) ^ s)) (h₂ _), + tsum_lt_tsum (i := ⟨N+1, by simp⟩) ?_ ?_ (h₂ _) (h₁ _ ((l + ε) ^ s))⟩ + · rintro ⟨i, hi⟩ + simp only [Finset.coe_range, Set.compl_Iio, Set.mem_Ici] at hi + exact (hN i hi s (zero_lt_one.trans hs)).1.le + · exact (hN (N + 1) (Nat.le_add_right N 1) s (zero_lt_one.trans hs)).1 + · rintro ⟨i, hi⟩ + simp only [Finset.coe_range, Set.compl_Iio, Set.mem_Ici] at hi + exact (hN i hi s (zero_lt_one.trans hs)).2.le + · exact (hN (N + 1) (Nat.le_add_right N 1) s (zero_lt_one.trans hs)).2 + +theorem tendsto_mul_u_rpow : + Tendsto (fun s ↦ (s - 1) * ∑' n, (u a n : ℝ) ^ (- s)) (𝓝[>] 1) (𝓝 l) := by + rw [Metric.tendsto_nhdsWithin_nhds] + intro ε' hε' + let ε := min l ε' + have h₀ : 0 < ε := by + aesop + have h₁ : 0 < ε / 3 := by positivity + have h₂ : ε / 3 < l := by + refine lt_of_lt_of_le ?_ (min_le_left l ε') + refine div_lt_self ?_ (by norm_num) + exact h₀ + have h₃ : 0 < l - ε / 3 := by + exact sub_pos.mpr h₂ + have h₄ : 0 < l + ε / 3 := by + positivity + obtain ⟨T, hT⟩ := exist_finset_lt_tsum_u_lt hl hA₁ h₁ h₂.le + obtain ⟨δ₁, hδ₁, hδ₁'⟩ := Metric.tendsto_nhdsWithin_nhds.mp + (tendsto_mul_sum_rpow T (u a)) (ε / 3) h₁ + obtain ⟨δ₂, hδ₂, hδ₂'⟩ := Metric.tendsto_nhdsWithin_nhds.mp + (tendsto_rpow_mul_tsum_rpow h₃.ne' T) (ε / 3) h₁ + obtain ⟨δ₃, hδ₃, hδ₃'⟩ := Metric.tendsto_nhdsWithin_nhds.mp + (tendsto_rpow_mul_tsum_rpow h₄.ne' T) (ε / 3) h₁ + let δ := min δ₁ (min δ₂ δ₃) + refine ⟨δ, ?_, ?_⟩ + · simp_all only [gt_iff_lt, lt_min_iff, and_self, div_pos_iff_of_pos_left, Nat.ofNat_pos, sub_pos, + Set.mem_Ioi, dist_zero_right, norm_mul, Real.norm_eq_abs, dist_sub_eq_dist_add_right, ε, δ] + · intro s hs hsδ + specialize hδ₁' hs (lt_of_lt_of_le hsδ (by simp [δ])) + specialize hδ₂' hs (lt_of_lt_of_le hsδ (by simp [δ])) + specialize hδ₃' hs (lt_of_lt_of_le hsδ (by simp [δ])) + simp_rw [Real.dist_eq, abs_lt] at hδ₂' hδ₃' ⊢ + rw [Real.dist_0_eq_abs, abs_lt] at hδ₁' + refine ⟨?_, ?_⟩ + · refine lt_of_le_of_lt ?_ (sub_lt_sub_right (hT s hs).1 l) + have := add_lt_add hδ₁'.1 hδ₂'.1 + rw [← add_sub_assoc, ← sub_add, ← sub_lt_iff_lt_add] at this + refine le_trans ?_ this.le + rw [sub_eq_add_neg, ← neg_div, add_thirds, neg_le_neg_iff] + exact min_le_right l ε' + · refine lt_of_lt_of_le (sub_lt_sub_right (hT s hs).2 l) ?_ + have := add_lt_add hδ₁'.2 hδ₃'.2 + rw [← add_sub_assoc, ← sub_sub, sub_lt_iff_lt_add] at this + refine le_trans this.le ?_ + rw [add_thirds] + exact min_le_right l ε' + +include hl hA₁ in +theorem LSeries_eq_of_summable {s : ℝ} (hs₁ : s ≠ 0) + (hs₂ : Summable (fun n ↦ (u a n : ℝ) ^ (- s))) : + LSeries (fun n ↦ a n) s = ∑' (n : ℕ), (u a n : ℝ) ^ (- s) := by + have : ∀ (n : ℕ), {k | u a k = n}.Finite := by + intro n + have := tendsto_atTop_u hl hA₁ + rw [← Nat.cofinite_eq_atTop, tendsto_def] at this + have := this {n}ᶜ (by simp only [mem_cofinite, compl_compl, Set.finite_singleton]) + rwa [Set.preimage_compl, mem_cofinite, compl_compl] at this + simp_rw [← tsum_card_nsmul_eq_tsum this (fun n ↦ (n : ℝ) ^ (- s)) hs₂, nsmul_eq_mul, LSeries, + ← Complex.ofReal_natCast, LSeries.term_eq_coe, ← Complex.ofReal_tsum] + congr with n + obtain hn | hn := Nat.eq_zero_or_pos n + · rw [hn, Nat.cast_zero, if_pos rfl, Real.zero_rpow (neg_ne_zero.mpr hs₁), mul_zero] + · rw [card_u_eq hl hA₁ hn, if_neg hn.ne', Real.rpow_neg (Nat.cast_nonneg _), ← div_eq_mul_inv] + +-- Use previous result +include hl hA₁ in +theorem _root_.NumberTheory.LSeries.tendsto_mul_of_sum_div_tendsto : + Tendsto (fun s : ℝ ↦ (s - 1) * LSeries (fun n ↦ a n) s) (𝓝[>] 1) (𝓝 l) := by + have : ∀ (n : ℕ), {k | u a k = n}.Finite := by + intro n + have := tendsto_atTop_u hl hA₁ + rw [← Nat.cofinite_eq_atTop, tendsto_def] at this + have := this {n}ᶜ (by simp only [mem_cofinite, compl_compl, Set.finite_singleton]) + rwa [Set.preimage_compl, mem_cofinite, compl_compl] at this + have t₀ := fun s (hs : s ∈ Set.Ioi (1 : ℝ)) ↦ + tsum_card_nsmul_eq_tsum this (fun n : ℕ ↦ (n : ℝ) ^ (- s)) (summable_u_rpow hl hA₁ hs) + simp_rw [nsmul_eq_mul] at t₀ + have t₁ := tendsto_mul_u_rpow hl hA₁ + simp_rw [LSeries, ← Complex.ofReal_natCast, LSeries.term_eq_coe, ← Complex.ofReal_tsum, + ← Complex.ofReal_one, ← Complex.ofReal_sub, ← Complex.ofReal_mul] + rw [Filter.tendsto_ofReal_iff] + refine Tendsto.congr' ?_ t₁ + filter_upwards [eventually_mem_nhdsWithin] with s hs + simp_rw [← t₀ s hs] + congr with n + obtain hn | hn := Nat.eq_zero_or_pos n + · rw [hn, Nat.cast_zero, if_pos rfl, Real.zero_rpow, mul_zero] + rw [neg_ne_zero] + exact (zero_lt_one.trans hs).ne' + · rw [card_u_eq hl hA₁ hn, if_neg hn.ne', Real.rpow_neg (Nat.cast_nonneg _), ← div_eq_mul_inv] + +end + +end NumberTheory.LSeries.residueFormula diff --git a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean index bcc35496bf6f6..e6b13d6b9d027 100644 --- a/Mathlib/NumberTheory/LSeries/RiemannZeta.lean +++ b/Mathlib/NumberTheory/LSeries/RiemannZeta.lean @@ -202,6 +202,18 @@ theorem zeta_nat_eq_tsum_of_gt_one {k : ℕ} (hk : 1 < k) : lemma riemannZeta_residue_one : Tendsto (fun s ↦ (s - 1) * riemannZeta s) (𝓝[≠] 1) (𝓝 1) := by exact hurwitzZetaEven_residue_one 0 +/-- The residue of `ζ(s)` at `s = 1` is equal to 1 expressed using `tsum`. -/ +theorem riemannZeta_residue_one' : + Tendsto (fun s : ℝ ↦ (s - 1) * ∑' (n : ℕ), 1 / (n : ℝ) ^ s) (𝓝[>] 1) (𝓝 1) := by + rw [← tendsto_ofReal_iff, ofReal_one] + have : Tendsto (fun s : ℝ ↦ (s : ℂ)) (𝓝[>] 1) (𝓝[≠] 1) := + continuous_ofReal.continuousWithinAt.tendsto_nhdsWithin (fun _ _ ↦ by aesop) + refine Tendsto.congr' ?_ (riemannZeta_residue_one.comp this) + filter_upwards [eventually_mem_nhdsWithin] with _ _ + simp_rw [Function.comp_apply, zeta_eq_tsum_one_div_nat_cpow (by rwa [ofReal_re]), + ofReal_mul, ofReal_tsum, ofReal_sub, ofReal_one, one_div, ofReal_inv, + ofReal_cpow ( Nat.cast_nonneg _), ofReal_natCast] + /- naming scheme was changed from `riemannCompletedZeta` to `completedRiemannZeta`; add aliases for the old names -/ section aliases diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 73c8984828756..5c501966dda21 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -4,6 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ import Mathlib.Algebra.Module.ZLattice.Basic +import Mathlib.Analysis.InnerProductSpace.ProdL2 +import Mathlib.MeasureTheory.Measure.Haar.Unique import Mathlib.NumberTheory.NumberField.FractionalIdeal import Mathlib.NumberTheory.NumberField.Units.Basic @@ -182,6 +184,26 @@ open NumberField.InfinitePlace Module Finset abbrev mixedSpace := ({w : InfinitePlace K // IsReal w} → ℝ) × ({w : InfinitePlace K // IsComplex w} → ℂ) +section Measure + +open MeasureTheory.Measure MeasureTheory + +variable [NumberField K] + +open Classical in +instance : IsAddHaarMeasure (volume : Measure (mixedSpace K)) := + prod.instIsAddHaarMeasure volume volume + +open Classical in +instance : NoAtoms (volume : Measure (mixedSpace K)) := by + obtain ⟨w⟩ := (inferInstance : Nonempty (InfinitePlace K)) + by_cases hw : IsReal w + · exact @prod.instNoAtoms_fst _ _ _ _ volume volume _ (pi_noAtoms ⟨w, hw⟩) + · exact @prod.instNoAtoms_snd _ _ _ _ volume volume _ + (pi_noAtoms ⟨w, not_isReal_iff_isComplex.mp hw⟩) + +end Measure + /-- The mixed embedding of a number field `K` into the mixed space of `K`. -/ noncomputable def _root_.NumberField.mixedEmbedding : K →+* (mixedSpace K) := RingHom.prod (Pi.ringHom fun w => embedding_of_isReal w.prop) @@ -343,6 +365,11 @@ theorem exists_normAtPlace_ne_zero_iff {x : mixedSpace K} : (∃ w, normAtPlace w x ≠ 0) ↔ x ≠ 0 := by rw [ne_eq, ← forall_normAtPlace_eq_zero_iff, not_forall] +theorem continuous_normAtPlace (w : InfinitePlace K) : + Continuous (normAtPlace w) := by + simp_rw [normAtPlace, MonoidWithZeroHom.coe_mk, ZeroHom.coe_mk] + split_ifs <;> fun_prop + variable [NumberField K] theorem nnnorm_eq_sup_normAtPlace (x : mixedSpace K) : @@ -421,6 +448,12 @@ theorem norm_eq_zero_iff' {x : mixedSpace K} (hx : x ∈ Set.range (mixedEmbeddi rw [norm_eq_norm, Rat.cast_abs, abs_eq_zero, Rat.cast_eq_zero, Algebra.norm_eq_zero_iff, map_eq_zero] +variable (K) in +protected theorem continuous_norm : Continuous (mixedEmbedding.norm : (mixedSpace K) → ℝ) := by + refine continuous_finset_prod Finset.univ fun _ _ ↦ ?_ + simp_rw [normAtPlace, MonoidWithZeroHom.coe_mk, ZeroHom.coe_mk, dite_pow] + split_ifs <;> fun_prop + end norm noncomputable section stdBasis @@ -603,7 +636,7 @@ theorem latticeBasis_apply (i : ChooseBasisIndex ℤ (𝓞 K)) : theorem mem_span_latticeBasis (x : (mixedSpace K)) : x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ - x ∈ ((mixedEmbedding K).comp (algebraMap (𝓞 K) K)).range := by + x ∈ mixedEmbedding.integerLattice K := by rw [show Set.range (latticeBasis K) = (mixedEmbedding K).toIntAlgHom.toLinearMap '' (Set.range (integralBasis K)) by rw [← Set.range_comp]; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] @@ -626,6 +659,13 @@ instance : IsZLattice ℝ (mixedEmbedding.integerLattice K) := by simp_rw [← span_latticeBasis] exact ZSpan.isZLattice (latticeBasis K) +open Classical in +theorem fundamentalDomain_integerLattice : + MeasureTheory.IsAddFundamentalDomain (mixedEmbedding.integerLattice K) + (ZSpan.fundamentalDomain (latticeBasis K)) := by + rw [← span_latticeBasis] + exact ZSpan.isAddFundamentalDomain (latticeBasis K) _ + theorem mem_rat_span_latticeBasis (x : K) : mixedEmbedding K x ∈ Submodule.span ℚ (Set.range (latticeBasis K)) := by rw [← Basis.sum_repr (integralBasis K) x, map_sum] @@ -652,6 +692,14 @@ theorem latticeBasis_repr_apply (x : K) (i : ChooseBasisIndex ℤ (𝓞 K)) : variable (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) +/-- The image of the fractional ideal `I` in the mixed space. -/ +abbrev idealLattice : Submodule ℤ (mixedSpace K) := LinearMap.range <| + (mixedEmbedding K).toIntAlgHom.toLinearMap ∘ₗ ((I : Submodule (𝓞 K) K).subtype.restrictScalars ℤ) + +theorem mem_idealLattice (x : mixedSpace K) : + x ∈ idealLattice K I ↔ ∃ y, y ∈ (I : Set K) ∧ mixedEmbedding K y = x := by + simp [idealLattice] + /-- The generalized index of the lattice generated by `I` in the lattice generated by `𝓞 K` is equal to the norm of the ideal `I`. The result is stated in terms of base change determinant and is the translation of `NumberField.det_basisOfFractionalIdeal_eq_absNorm` in @@ -706,6 +754,334 @@ theorem mem_span_fractionalIdealLatticeBasis (x : (mixedSpace K)) : ext; erw [mem_span_basisOfFractionalIdeal]] rfl +theorem span_idealLatticeBasis : + (Submodule.span ℤ (Set.range (fractionalIdealLatticeBasis K I))) = + (mixedEmbedding.idealLattice K I) := by + ext x + simp [mem_span_fractionalIdealLatticeBasis] + +instance : DiscreteTopology (mixedEmbedding.idealLattice K I) := by + classical + rw [← span_idealLatticeBasis] + infer_instance + +open Classical in +instance : IsZLattice ℝ (mixedEmbedding.idealLattice K I) := by + simp_rw [← span_idealLatticeBasis] + exact ZSpan.isZLattice (fractionalIdealLatticeBasis K I) + +open Classical in +theorem fundamentalDomain_idealLattice : + MeasureTheory.IsAddFundamentalDomain (mixedEmbedding.idealLattice K I) + (ZSpan.fundamentalDomain (fractionalIdealLatticeBasis K I)) := by + rw [← span_idealLatticeBasis] + exact ZSpan.isAddFundamentalDomain (fractionalIdealLatticeBasis K I) _ + end integerLattice +noncomputable section plusPart + +variable {K} + +open Classical in +/-- Let `s` be a set of real places, define the continuous linear equiv of the mixed space that +changes sign at places in `s` and leaves the rest unchanged. -/ +def negAt (s : Set {w : InfinitePlace K // IsReal w}) : + (mixedSpace K) ≃L[ℝ] (mixedSpace K) := + (ContinuousLinearEquiv.piCongrRight + fun w ↦ if w ∈ s then ContinuousLinearEquiv.neg ℝ else ContinuousLinearEquiv.refl ℝ ℝ).prod + (ContinuousLinearEquiv.refl ℝ _) + +@[simp] +theorem negAt_apply_of_isReal_and_mem {s : Set {w // IsReal w}} (x : mixedSpace K) + {w : {w // IsReal w}} (hw : w ∈ s) : + (negAt s x).1 w = - x.1 w := by + simp_rw [negAt, ContinuousLinearEquiv.prod_apply, ContinuousLinearEquiv.piCongrRight_apply, + if_pos hw, ContinuousLinearEquiv.neg_apply] + +@[simp] +theorem negAt_apply_of_isReal_and_not_mem {s : Set {w // IsReal w}} (x : mixedSpace K) + {w : {w // IsReal w}} (hw : w ∉ s) : + (negAt s x).1 w = x.1 w := by + simp_rw [negAt, ContinuousLinearEquiv.prod_apply, ContinuousLinearEquiv.piCongrRight_apply, + if_neg hw, ContinuousLinearEquiv.refl_apply] + +@[simp] +theorem negAt_apply_abs_of_isReal {s : Set {w // IsReal w}} (x : mixedSpace K) + (w : {w // IsReal w}) : + |(negAt s x).1 w| = |x.1 w| := by + by_cases hw : w ∈ s + · rw [negAt_apply_of_isReal_and_mem _ hw, abs_neg] + · rw [negAt_apply_of_isReal_and_not_mem _ hw] + +@[simp] +theorem negAt_apply_snd (s : Set {w // IsReal w}) (x : mixedSpace K) : + (negAt s x).2 = x.2 := rfl + +open MeasureTheory Classical in +/-- `negAt` preserves the volume . -/ +theorem volume_preserving_negAt [NumberField K] (s : Set {w : InfinitePlace K // IsReal w}) : + MeasurePreserving (negAt s) := by + refine MeasurePreserving.prod (volume_preserving_pi fun w ↦ ?_) (MeasurePreserving.id _) + by_cases hw : w ∈ s + · simp_rw [if_pos hw] + exact Measure.measurePreserving_neg _ + · simp_rw [if_neg hw] + exact MeasurePreserving.id _ + +/-- `negAt` preserves `normAtPlace`. -/ +@[simp] +theorem normAtPlace_negAt (s : Set {w // IsReal w}) (x : mixedSpace K) (w : InfinitePlace K) : + normAtPlace w (negAt s x) = normAtPlace w x := by + obtain hw | hw := isReal_or_isComplex w + · simp_rw [normAtPlace_apply_isReal hw, Real.norm_eq_abs, negAt_apply_abs_of_isReal] + · simp_rw [normAtPlace_apply_isComplex hw, congr_fun (negAt_apply_snd s x) ⟨w, hw⟩] + +/-- `negAt` preserves the `norm`. -/ +@[simp] +theorem norm_negAt [NumberField K] (s : Set {w // IsReal w}) (x : mixedSpace K) : + mixedEmbedding.norm (negAt s x) = mixedEmbedding.norm x := + norm_eq_of_normAtPlace_eq (fun w ↦ normAtPlace_negAt _ _ w) + +open ContinuousLinearEquiv in +/-- `negAt` is equal to its inverse. -/ +theorem negAt_symm (s : Set {w : InfinitePlace K // IsReal w}) : + (negAt s).symm = negAt s := by + ext x w + · by_cases hw : w ∈ s + · simp_rw [negAt_apply_of_isReal_and_mem _ hw, negAt, prod_symm, + ContinuousLinearEquiv.prod_apply, piCongrRight_symm_apply, if_pos hw, symm_neg, neg_apply] + · simp_rw [negAt_apply_of_isReal_and_not_mem _ hw, negAt, prod_symm, + ContinuousLinearEquiv.prod_apply, piCongrRight_symm_apply, if_neg hw, refl_symm, refl_apply] + · rfl + +variable (A : Set (mixedSpace K)) + +/-- `negAt s A` is also equal to the preimage of `A` by `negAt s`. This fact can be useful to +simplify some proofs. -/ +theorem negAt_preimage (s : Set {w // IsReal w} ) : + negAt s ⁻¹' A = negAt s '' A := by + rw [ContinuousLinearEquiv.image_eq_preimage, negAt_symm] + +/-- The `Plus` part of a subset `A` of the `mixedSpace` is the set of points in `A` that are +positive at all real places. -/ +abbrev plusPart : Set (mixedSpace K) := A ∩ {x | ∀ w, 0 < x.1 w} + +/-- Let `s` be a set of real places. `negAtPlusPart A s` is the image of `plusPart A` by `negAt s`, +thus it is equal to the set of elements of `A` that are negative at real places in `s` and +positive at real places not in `s`. -/ +abbrev negAtPlusPart (s : Set {w : InfinitePlace K // IsReal w}) : Set (mixedSpace K) := + negAt s '' (plusPart A) + +theorem negAtPlusPart_neg_of_mem {s : Set {w // IsReal w}} {x : mixedSpace K} + (hx : x ∈ negAtPlusPart A s) {w : {w // IsReal w}} (hw : w ∈ s) : + x.1 w < 0 := by + obtain ⟨y, hy, rfl⟩ := hx + rw [negAt_apply_of_isReal_and_mem _ hw, neg_lt_zero] + exact hy.2 w + +theorem negAtPlusPart_pos_of_not_mem {s : Set {w // IsReal w}} {x : mixedSpace K} + (hx : x ∈ negAtPlusPart A s) {w : {w // IsReal w}} (hw : w ∉ s) : + 0 < x.1 w := by + obtain ⟨y, hy, rfl⟩ := hx + rw [negAt_apply_of_isReal_and_not_mem _ hw] + exact hy.2 w + +open Classical in +theorem mem_negAtPlusPart_of_mem (hA : ∀ x, x ∈ A ↔ (fun w ↦ |x.1 w|, x.2) ∈ A) + (s : Set {w // IsReal w}) {x : mixedSpace K} (hx₁ : x ∈ A) (hx₂ : ∀ w, x.1 w ≠ 0) : + x ∈ negAtPlusPart A s ↔ (∀ w, w ∈ s → x.1 w < 0) ∧ (∀ w, w ∉ s → x.1 w > 0) := by + refine ⟨fun hx ↦ ⟨fun _ hw ↦ negAtPlusPart_neg_of_mem A hx hw, + fun _ hw ↦ negAtPlusPart_pos_of_not_mem A hx hw⟩, + fun ⟨h₁, h₂⟩ ↦ ⟨(fun w ↦ |x.1 w|, x.2), ⟨(hA x).mp hx₁, fun w ↦ abs_pos.mpr (hx₂ w)⟩, ?_⟩⟩ + ext w + · by_cases hw : w ∈ s + · simp_rw [negAt_apply_of_isReal_and_mem _ hw, abs_of_neg (h₁ w hw), neg_neg] + · simp_rw [negAt_apply_of_isReal_and_not_mem _ hw, abs_of_pos (h₂ w hw)] + · rfl + +/-- The sets `negAtPlusPart` are pairwise disjoint. -/ +theorem disjoint_negAtPlusPart : Pairwise (Disjoint on (negAtPlusPart A)) := by + classical + intro s t hst + refine Set.disjoint_left.mpr fun _ hx hx' ↦ ?_ + obtain ⟨w, hw | hw⟩ : ∃ w, (w ∈ s ∧ w ∉ t) ∨ (w ∈ t ∧ w ∉ s) := by + exact Set.symmDiff_nonempty.mpr hst + · exact lt_irrefl _ <| + (negAtPlusPart_neg_of_mem A hx hw.1).trans (negAtPlusPart_pos_of_not_mem A hx' hw.2) + · exact lt_irrefl _ <| + (negAtPlusPart_neg_of_mem A hx' hw.1).trans (negAtPlusPart_pos_of_not_mem A hx hw.2) + +/-- For `x : mixedSpace K`, the set `signSet x` is the set of real places `w` s.t. `x w ≤ 0`. -/ +def signSet (x : mixedSpace K) : Set {w : InfinitePlace K // IsReal w} := {w | x.1 w ≤ 0} + +@[simp] +theorem negAt_signSet_apply_of_isReal (x : mixedSpace K) (w : {w // IsReal w}) : + (negAt (signSet x) x).1 w = |x.1 w| := by + by_cases hw : x.1 w ≤ 0 + · rw [negAt_apply_of_isReal_and_mem _ hw, abs_of_nonpos hw] + · rw [negAt_apply_of_isReal_and_not_mem _ hw, abs_of_pos (lt_of_not_ge hw)] + +@[simp] +theorem negAt_signSet_apply_of_isComplex (x : mixedSpace K) (w : {w // IsComplex w}) : + (negAt (signSet x) x).2 w = x.2 w := rfl + +/-- Assume that the images of `plusPart A` under `negAt` are all subset of `A`. Then, the union of +all the `negPlusPart` and the set of elements of `A` that are zero at at least one real place is +equal to `A`. -/ +theorem iUnion_negAtPlusPart_union (hA : ∀ x, x ∈ A ↔ (fun w ↦ |x.1 w|, x.2) ∈ A) : + (⋃ s, negAtPlusPart A s) ∪ (A ∩ (⋃ w, {x | x.1 w = 0})) = A := by + ext x + rw [Set.mem_union, Set.mem_inter_iff, Set.mem_iUnion, Set.mem_iUnion] + refine ⟨?_, fun h ↦ ?_⟩ + · rintro (⟨s, ⟨x, ⟨hx, _⟩, rfl⟩⟩ | h) + · simp_rw (config := {singlePass := true}) [hA, negAt_apply_abs_of_isReal, negAt_apply_snd] + rwa [← hA] + · exact h.left + · obtain hx | hx := exists_or_forall_not (fun w ↦ x.1 w = 0) + · exact Or.inr ⟨h, hx⟩ + · refine Or.inl ⟨signSet x, + (mem_negAtPlusPart_of_mem A hA _ h hx).mpr ⟨fun w hw ↦ ?_, fun w hw ↦ ?_⟩⟩ + · exact lt_of_le_of_ne hw (hx w) + · exact lt_of_le_of_ne (lt_of_not_ge hw).le (Ne.symm (hx w)) + +open MeasureTheory + +variable [NumberField K] + +open Classical in +/-- The set of points in the mixedSpace that are equal to `0` at a fixed (real) place has +volume zero. -/ +theorem volume_eq_zero (w : {w // IsReal w}) : + volume ({x : mixedSpace K | x.1 w = 0}) = 0 := by + let A : AffineSubspace ℝ (mixedSpace K) := + Submodule.toAffineSubspace (Submodule.mk ⟨⟨{x | x.1 w = 0}, by aesop⟩, rfl⟩ (by aesop)) + convert Measure.addHaar_affineSubspace volume A fun h ↦ ?_ + have : 1 ∈ A := h ▸ Set.mem_univ _ + simp [A] at this + +open Classical in +theorem iUnion_negAtPart_ae (hA : ∀ x, x ∈ A ↔ (fun w ↦ |x.1 w|, x.2) ∈ A) : + ⋃ s, negAtPlusPart A s =ᵐ[volume] A := by + nth_rewrite 2 [← iUnion_negAtPlusPart_union A hA] + refine (MeasureTheory.union_ae_eq_left_of_ae_eq_empty (ae_eq_empty.mpr ?_)).symm + exact measure_mono_null Set.inter_subset_right + (measure_iUnion_null_iff.mpr fun _ ↦ volume_eq_zero _) + +variable {A} in +theorem measurableSet_plusPart (hm : MeasurableSet A) : + MeasurableSet (plusPart A) := by + convert_to MeasurableSet (A ∩ (⋂ w, {x | 0 < x.1 w})) + · ext; simp + · refine hm.inter (MeasurableSet.iInter fun _ ↦ ?_) + exact measurableSet_lt measurable_const ((measurable_pi_apply _).comp' measurable_fst) + +theorem measurableSet_negAtPlusPart (s : Set {w : InfinitePlace K // IsReal w}) + (hm : MeasurableSet A) : + MeasurableSet (negAtPlusPart A s) := by + rw [negAtPlusPart, ← negAt_preimage] + exact (measurableSet_plusPart hm).preimage (negAt s).continuous.measurable + +open Classical in +theorem volume_negAtPlusPart (hm : MeasurableSet A) (s : Set {w // IsReal w}) : + volume (negAtPlusPart A s) = volume (plusPart A) := by + rw [negAtPlusPart, ← negAt_symm, ContinuousLinearEquiv.image_symm_eq_preimage, + (volume_preserving_negAt s).measure_preimage (measurableSet_plusPart hm).nullMeasurableSet] + +open Classical in +/-- If a subset `A` of the `mixedSpace` is symmetric at real places, then its volume is +`2^r₁` times the volume of its `plusPart` where `r₁` is the number of real places. -/ +theorem volume_eq_two_pow_mul_volume_plusPart (hA : ∀ x, x ∈ A ↔ (fun w ↦ |x.1 w|, x.2) ∈ A) + (hm : MeasurableSet A) : + volume A = 2 ^ NrRealPlaces K * volume (plusPart A) := by + simp_rw [← measure_congr (iUnion_negAtPart_ae A hA), measure_iUnion (disjoint_negAtPlusPart A) + (fun _ ↦ measurableSet_negAtPlusPart A _ hm), volume_negAtPlusPart _ hm, tsum_fintype, + Finset.sum_const, Finset.card_univ, NrRealPlaces, nsmul_eq_mul, Fintype.card_set, Nat.cast_pow, + Nat.cast_ofNat] + +end plusPart + +noncomputable section + +namespace euclidean + +open MeasureTheory NumberField Submodule + +/-- The mixed space `ℝ^r₁ × ℂ^r₂`, with `(r₁, r₂)` the signature of `K`, as an Euclidean space. -/ +protected abbrev mixedSpace := + (WithLp 2 ((EuclideanSpace ℝ {w : InfinitePlace K // IsReal w}) × + (EuclideanSpace ℂ {w : InfinitePlace K // IsComplex w}))) + +instance : Ring (euclidean.mixedSpace K) := by + have : Ring (EuclideanSpace ℝ {w : InfinitePlace K // IsReal w}) := Pi.ring + have : Ring (EuclideanSpace ℂ {w : InfinitePlace K // IsComplex w}) := Pi.ring + exact Prod.instRing + +instance : MeasurableSpace (euclidean.mixedSpace K) := borel _ + +instance : BorelSpace (euclidean.mixedSpace K) := ⟨rfl⟩ + +variable [NumberField K] + +open Classical in +instance : T2Space (euclidean.mixedSpace K) := Prod.t2Space + +open Classical in +/-- The continuous linear equivalence between the euclidean mixed space and the mixed space. -/ +def toMixed : (euclidean.mixedSpace K) ≃L[ℝ] (mixedSpace K) := + (WithLp.linearEquiv _ _ _).toContinuousLinearEquiv + +instance : Nontrivial (euclidean.mixedSpace K) := (toMixed K).toEquiv.nontrivial + +protected theorem finrank : + finrank ℝ (euclidean.mixedSpace K) = finrank ℚ K := by + rw [LinearEquiv.finrank_eq (toMixed K).toLinearEquiv, mixedEmbedding.finrank] + +open Classical in +/-- An orthonormal basis of the euclidean mixed space. -/ +def stdOrthonormalBasis : OrthonormalBasis (index K) ℝ (euclidean.mixedSpace K) := + OrthonormalBasis.prod (EuclideanSpace.basisFun _ ℝ) + ((Pi.orthonormalBasis fun _ ↦ Complex.orthonormalBasisOneI).reindex (Equiv.sigmaEquivProd _ _)) + +open Classical in +theorem stdOrthonormalBasis_map_eq : + (euclidean.stdOrthonormalBasis K).toBasis.map (toMixed K).toLinearEquiv = + mixedEmbedding.stdBasis K := by + ext _ _ <;> rfl + +open Classical in +theorem volumePreserving_toMixed : + MeasurePreserving (toMixed K) where + measurable := (toMixed K).continuous.measurable + map_eq := by + rw [← (OrthonormalBasis.addHaar_eq_volume (euclidean.stdOrthonormalBasis K)), Basis.map_addHaar, + stdOrthonormalBasis_map_eq, Basis.addHaar_eq_iff, Basis.coe_parallelepiped, + ← measure_congr (ZSpan.fundamentalDomain_ae_parallelepiped (stdBasis K) volume), + volume_fundamentalDomain_stdBasis K] + +open Classical in +theorem volumePreserving_toMixed_symm : + MeasurePreserving (toMixed K).symm := by + have : MeasurePreserving (toMixed K).toHomeomorph.toMeasurableEquiv := volumePreserving_toMixed K + exact this.symm + +open Classical in +/-- The image of ring of integers `𝓞 K` in the euclidean mixed space. -/ +protected def integerLattice : Submodule ℤ (euclidean.mixedSpace K) := + ZLattice.comap ℝ (mixedEmbedding.integerLattice K) (toMixed K).toLinearMap + +instance : DiscreteTopology (euclidean.integerLattice K) := by + classical + rw [euclidean.integerLattice] + infer_instance + +open Classical in +instance : IsZLattice ℝ (euclidean.integerLattice K) := by + simp_rw [euclidean.integerLattice] + infer_instance + +end euclidean + +end + end NumberField.mixedEmbedding diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean index 36a06aabe738b..1b946441109c7 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/ConvexBody.lean @@ -77,15 +77,6 @@ open scoped Classical variable [NumberField K] -instance : IsAddHaarMeasure (volume : Measure (mixedSpace K)) := prod.instIsAddHaarMeasure _ _ - -instance : NoAtoms (volume : Measure (mixedSpace K)) := by - obtain ⟨w⟩ := (inferInstance : Nonempty (InfinitePlace K)) - by_cases hw : IsReal w - · exact @prod.instNoAtoms_fst _ _ _ _ volume volume _ (pi_noAtoms ⟨w, hw⟩) - · exact @prod.instNoAtoms_snd _ _ _ _ volume volume _ - (pi_noAtoms ⟨w, not_isReal_iff_isComplex.mp hw⟩) - /-- The fudge factor that appears in the formula for the volume of `convexBodyLT`. -/ noncomputable abbrev convexBodyLTFactor : ℝ≥0 := (2 : ℝ≥0) ^ NrRealPlaces K * NNReal.pi ^ NrComplexPlaces K @@ -315,7 +306,7 @@ theorem convexBodySumFun_eq_zero_iff (x : mixedSpace K) : conv => enter [1, w, hw] rw [mul_left_mem_nonZeroDivisors_eq_zero_iff - (mem_nonZeroDivisors_iff_ne_zero.mpr <| Nat.cast_ne_zero.mpr mult_ne_zero)] + (mem_nonZeroDivisors_iff_ne_zero.mpr mult_coe_ne_zero)] simp_rw [Finset.mem_univ, true_implies] theorem norm_le_convexBodySumFun (x : mixedSpace K) : ‖x‖ ≤ convexBodySumFun x := by diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean index 419b0f6e623c0..e9c75171152e3 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/FundamentalCone.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ import Mathlib.RingTheory.Ideal.IsPrincipal +import Mathlib.RingTheory.ClassGroup import Mathlib.NumberTheory.NumberField.Units.DirichletTheorem /-! @@ -164,6 +165,23 @@ theorem logMap_eq_of_normAtPlace_eq (h : ∀ w, normAtPlace w x = normAtPlace w ext simp_rw [logMap_apply, h, norm_eq_of_normAtPlace_eq h] +variable (K) + +theorem measurable_logMap : + Measurable (logMap : (mixedSpace K) → _) := + measurable_pi_iff.mpr fun _ ↦ + measurable_const.mul <| (continuous_normAtPlace _).measurable.log.sub + <| (mixedEmbedding.continuous_norm K).measurable.log.mul measurable_const + +theorem continuousOn_logMap : + ContinuousOn (logMap : (mixedSpace K) → _) {x | mixedEmbedding.norm x ≠ 0} := by + refine continuousOn_pi.mpr fun w ↦ continuousOn_const.mul (ContinuousOn.sub ?_ ?_) + · exact Real.continuousOn_log.comp'' (continuous_normAtPlace _).continuousOn + fun _ hx ↦ mixedEmbedding.norm_ne_zero_iff.mp hx _ + · exact ContinuousOn.mul + (Real.continuousOn_log.comp'' (mixedEmbedding.continuous_norm K).continuousOn + fun _ hx ↦ hx) continuousOn_const + end logMap noncomputable section @@ -180,6 +198,17 @@ def fundamentalCone : Set (mixedSpace K) := logMap⁻¹' (ZSpan.fundamentalDomain ((basisUnitLattice K).ofZLatticeBasis ℝ _)) \ {x | mixedEmbedding.norm x = 0} +theorem measurableSet_fundamentalCone : + MeasurableSet (fundamentalCone K) := by + classical + refine MeasurableSet.diff ?_ ?_ + · unfold logMap + refine MeasurableSet.preimage (ZSpan.fundamentalDomain_measurableSet _) <| + measurable_pi_iff.mpr fun w ↦ measurable_const.mul ?_ + exact (continuous_normAtPlace _).measurable.log.sub <| + (mixedEmbedding.continuous_norm _).measurable.log.mul measurable_const + · exact measurableSet_eq_fun (mixedEmbedding.continuous_norm K).measurable measurable_const + namespace fundamentalCone variable {K} {x y : mixedSpace K} {c : ℝ} @@ -286,10 +315,10 @@ theorem mixedEmbedding_preimageOfIntegralPoint (a : integralPoint K) : mixedEmbedding K (preimageOfIntegralPoint a : 𝓞 K) = (a : mixedSpace K) := by rw [preimageOfIntegralPoint, (mem_integralPoint.mp a.prop).2.choose_spec] -theorem preimageOfIntegralPoint_mixedEmbedding {x : (𝓞 K)⁰} +theorem preimageOfIntegralPoint_mixedEmbedding {x : (𝓞 K)} (hx : mixedEmbedding K (x : 𝓞 K) ∈ integralPoint K) : preimageOfIntegralPoint (⟨mixedEmbedding K (x : 𝓞 K), hx⟩) = x := by - simp_rw [Subtype.ext_iff, RingOfIntegers.ext_iff, ← (mixedEmbedding_injective K).eq_iff, + simp_rw [RingOfIntegers.ext_iff, ← (mixedEmbedding_injective K).eq_iff, mixedEmbedding_preimageOfIntegralPoint] /-- If `x : mixedSpace K` is nonzero and the image of an algebraic integer, then there exists a @@ -427,29 +456,30 @@ variable (K) in /-- For an integer `n`, The equivalence between the `integralPoint K` of norm `n` and the product of the set of nonzero principal ideals of `K` of norm `n` and the torsion of `K`. -/ def integralPointEquivNorm (n : ℕ) : - {a : integralPoint K // intNorm a = n} ≃ + {a : integralPoint K // mixedEmbedding.norm (a : mixedSpace K) = n} ≃ {I : (Ideal (𝓞 K))⁰ // IsPrincipal (I : Ideal (𝓞 K)) ∧ absNorm (I : Ideal (𝓞 K)) = n} × (torsion K) := - calc {a // intNorm a = n} - ≃ {I : {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1} × torsion K // + calc + _ ≃ {I : {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1} × torsion K // absNorm (I.1 : Ideal (𝓞 K)) = n} := - (Equiv.subtypeEquiv (integralPointEquiv K) fun _ ↦ by simp [intNorm, absNorm_span_singleton]) + Equiv.subtypeEquiv (integralPointEquiv K) fun _ ↦ by simp_rw [← intNorm_coe, intNorm, + Nat.cast_inj, integralPointEquiv_apply_fst, absNorm_span_singleton] _ ≃ {I : {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1} // absNorm (I.1 : Ideal (𝓞 K)) = n} × - torsion K := - Equiv.prodSubtypeFstEquivSubtypeProd (p := fun I : {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1} ↦ - absNorm (I : Ideal (𝓞 K)) = n) + torsion K := Equiv.prodSubtypeFstEquivSubtypeProd + (p := fun I : {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1} ↦ absNorm (I : Ideal (𝓞 K)) = n) _ ≃ {I : (Ideal (𝓞 K))⁰ // IsPrincipal (I : Ideal (𝓞 K)) ∧ - absNorm (I : Ideal (𝓞 K)) = n} × (torsion K) := - Equiv.prodCongrLeft fun _ ↦ (Equiv.subtypeSubtypeEquivSubtypeInter + absNorm (I : Ideal (𝓞 K)) = n} × (torsion K) := Equiv.prodCongrLeft fun _ ↦ + (Equiv.subtypeSubtypeEquivSubtypeInter (fun I : (Ideal (𝓞 K))⁰ ↦ IsPrincipal I.1) (fun I ↦ absNorm I.1 = n)) @[simp] -theorem integralPointEquivNorm_apply_fst {n : ℕ} {a : integralPoint K} (ha : intNorm a = n) : - ((integralPointEquivNorm K n ⟨a, ha⟩).1 : Ideal (𝓞 K)) = - span {(preimageOfIntegralPoint a : 𝓞 K)} := by - simp_rw [integralPointEquivNorm, Equiv.prodSubtypeFstEquivSubtypeProd, Equiv.instTrans_trans, - Equiv.prodCongrLeft, Equiv.trans_apply, Equiv.subtypeEquiv_apply, Equiv.coe_fn_mk, - Equiv.subtypeSubtypeEquivSubtypeInter_apply_coe, integralPointEquiv_apply_fst] +theorem integralPointEquivNorm_apply_fst {n : ℕ} + (a : {a : integralPoint K // mixedEmbedding.norm (a : mixedSpace K) = n}) : + ((integralPointEquivNorm K n a).1 : Ideal (𝓞 K)) = + span {(preimageOfIntegralPoint a.val : 𝓞 K)} := by + simp_rw [integralPointEquivNorm, Equiv.prodSubtypeFstEquivSubtypeProd, Equiv.instTrans_trans, + Equiv.prodCongrLeft, Equiv.trans_apply, Equiv.subtypeEquiv_apply, Equiv.coe_fn_mk, + Equiv.subtypeSubtypeEquivSubtypeInter_apply_coe, integralPointEquiv_apply_fst] variable (K) @@ -458,42 +488,141 @@ of the torsion of `K` is equal to the number of `integralPoint K` of norm `n`. - theorem card_isPrincipal_norm_eq_mul_torsion (n : ℕ) : Nat.card {I : (Ideal (𝓞 K))⁰ | IsPrincipal (I : Ideal (𝓞 K)) ∧ absNorm (I : Ideal (𝓞 K)) = n} * torsionOrder K = - Nat.card {a : integralPoint K | intNorm a = n} := by + Nat.card {a : integralPoint K | mixedEmbedding.norm (a : mixedSpace K) = n} := by rw [torsionOrder, PNat.mk_coe, ← Nat.card_eq_fintype_card, ← Nat.card_prod] exact Nat.card_congr (integralPointEquivNorm K n).symm -/-- For `s : ℝ`, the number of principal nonzero ideals in `𝓞 K` of norm `≤ s` multiplied by the -order of the torsion of `K` is equal to the number of `integralPoint K` of norm `≤ s`. -/ -theorem card_isPrincipal_norm_le_mul_torsion (s : ℝ) : - Nat.card {I : (Ideal (𝓞 K))⁰ | IsPrincipal (I : Ideal (𝓞 K)) ∧ +variable (J : (Ideal (𝓞 K))⁰) + +/-- The set of images by `mixedEmbedding` of the elements of the integral ideal `J` contained in +the fundamental cone. -/ +def idealPoint : Set (mixedSpace K) := + fundamentalCone K ∩ (mixedEmbedding.idealLattice K (FractionalIdeal.mk0 K J)) + +variable {K J} in +theorem mem_idealPoint : + x ∈ idealPoint K J ↔ x ∈ fundamentalCone K ∧ ∃ a : (𝓞 K), (a : 𝓞 K) ∈ (J : Set (𝓞 K)) ∧ + mixedEmbedding K (a : 𝓞 K) = x := by + simp_rw [idealPoint, Set.mem_inter_iff, idealLattice, SetLike.mem_coe, FractionalIdeal.coe_mk0, + LinearMap.mem_range, LinearMap.coe_comp, LinearMap.coe_restrictScalars, coe_subtype, + Function.comp_apply, AlgHom.toLinearMap_apply, RingHom.toIntAlgHom_coe, Subtype.exists, + FractionalIdeal.mem_coe, FractionalIdeal.mem_coeIdeal, exists_prop', nonempty_prop, + exists_exists_and_eq_and] + +/-- The map that sends an `idealPoint` to an `integralPoint`. This map exists because `J` is an +integral ideal. -/ +def idealPointMap : idealPoint K J → integralPoint K := + fun ⟨a, ha⟩ ↦ ⟨a, + mem_integralPoint.mpr ⟨(mem_idealPoint.mp ha).1, (mem_idealPoint.mp ha).2.choose, + (mem_idealPoint.mp ha).2.choose_spec.2⟩⟩ + +@[simp] +theorem idealPointMap_apply (a : idealPoint K J) : + (idealPointMap K J a : mixedSpace K) = a := rfl + +theorem preimage_of_IdealPointMap (a : idealPoint K J) : + (preimageOfIntegralPoint (idealPointMap K J a) : 𝓞 K) ∈ (J : Set (𝓞 K)) := by + obtain ⟨_, ⟨x, hx₁, hx₂⟩⟩ := mem_idealPoint.mp a.prop + simp_rw [idealPointMap, ← hx₂, preimageOfIntegralPoint_mixedEmbedding] + exact hx₁ + +/-- The map `idealPointMap` is actually an equiv between the `idealPoint K J` and the +`integralPoint K` whose preimage lies in `J`. -/ +def idealPointEquiv : idealPoint K J ≃ + {a : integralPoint K | (preimageOfIntegralPoint a : 𝓞 K) ∈ (J : Set (𝓞 K))} := + Equiv.ofBijective (fun a ↦ ⟨idealPointMap K J a, preimage_of_IdealPointMap K J a⟩) + ⟨fun _ _ h ↦ (by + simp_rw [Subtype.ext_iff_val, idealPointMap_apply] at h + rwa [Subtype.ext_iff_val]), + fun ⟨a, ha₂⟩ ↦ ⟨⟨a.val, mem_idealPoint.mpr ⟨a.prop.1, + ⟨preimageOfIntegralPoint a, ha₂, mixedEmbedding_preimageOfIntegralPoint a⟩⟩⟩, rfl⟩⟩ + +variable {K J} + +theorem idealPointEquiv_apply (a : idealPoint K J) : + (idealPointEquiv K J a : mixedSpace K) = a := rfl + +theorem idealPointEquiv_symm_apply + (a : {a : integralPoint K // (preimageOfIntegralPoint a : 𝓞 K) ∈ (J : Set (𝓞 K)) }) : + ((idealPointEquiv K J).symm a : mixedSpace K) = a := by + rw [← (idealPointEquiv_apply ((idealPointEquiv K J).symm a)), Equiv.apply_symm_apply] + +theorem intNorm_idealPointEquiv_apply (a : idealPoint K J) : + intNorm (idealPointEquiv K J a).val = mixedEmbedding.norm (a : mixedSpace K) := by + rw [intNorm_coe, idealPointEquiv_apply] + +variable (K J) + +/-- For an integer `n`, The equivalence between the `idealPoint K` of norm `n` and the product +of the set of nonzero principal ideals of `K` divisible by `J` of norm `n` and the +torsion of `K`. -/ +def idealPointEquivNorm (n : ℕ) : + {a : idealPoint K J // mixedEmbedding.norm (a : mixedSpace K) = n} ≃ + {I : (Ideal (𝓞 K))⁰ // (J : Ideal (𝓞 K)) ∣ I ∧ IsPrincipal (I : Ideal (𝓞 K)) ∧ + absNorm (I : Ideal (𝓞 K)) = n} × (torsion K) := by + calc + _ ≃ {a : {a : integralPoint K // (preimageOfIntegralPoint a).1 ∈ J.1} // + mixedEmbedding.norm a.1.1 = n} := by + convert (Equiv.subtypeEquivOfSubtype (idealPointEquiv K J).symm).symm using 3 + rw [idealPointEquiv_symm_apply] + _ ≃ {a : integralPoint K // (preimageOfIntegralPoint a).1 ∈ J.1 ∧ + mixedEmbedding.norm a.1 = n} := Equiv.subtypeSubtypeEquivSubtypeInter + (fun a : integralPoint K ↦ (preimageOfIntegralPoint a).1 ∈ J.1) + (fun a ↦ mixedEmbedding.norm a.1 = n) + _ ≃ {a : {a :integralPoint K // mixedEmbedding.norm a.1 = n} // + (preimageOfIntegralPoint a.1).1 ∈ J.1} := ((Equiv.subtypeSubtypeEquivSubtypeInter + (fun a : integralPoint K ↦ mixedEmbedding.norm a.1 = n) + (fun a ↦ (preimageOfIntegralPoint a).1 ∈ J.1)).trans + (Equiv.subtypeEquivRight (fun _ ↦ by simp [and_comm]))).symm + _ ≃ {I : {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1 ∧ absNorm I.1 = n} × (torsion K) // + J.1 ∣ I.1.1} := by + convert Equiv.subtypeEquivOfSubtype (p := fun I ↦ J.1 ∣ I.1) (integralPointEquivNorm K n) + rw [integralPointEquivNorm_apply_fst, dvd_span_singleton] + _ ≃ {I : {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1 ∧ absNorm I.1 = n} // J.1 ∣ I.1} × + (torsion K) := Equiv.prodSubtypeFstEquivSubtypeProd + (p := fun I : {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1 ∧ absNorm I.1 = n} ↦ J.1 ∣ I.1) + _ ≃ {I : (Ideal (𝓞 K))⁰ // (IsPrincipal I.1 ∧ absNorm I.1 = n) ∧ J.1 ∣ I.1} × (torsion K) := + Equiv.prodCongrLeft fun _ ↦ (Equiv.subtypeSubtypeEquivSubtypeInter + (fun I : (Ideal (𝓞 K))⁰ ↦ IsPrincipal I.1 ∧ absNorm I.1 = n) + (fun I ↦ J.1 ∣ I.1)) + _ ≃ {I : (Ideal (𝓞 K))⁰ // J.1 ∣ I.1 ∧ IsPrincipal I.1 ∧ absNorm I.1 = n} × + (Units.torsion K) := + Equiv.prodCongrLeft fun _ ↦ Equiv.subtypeEquivRight fun _ ↦ by rw [and_comm] + +/-- For `s : ℝ`, the number of principal nonzero ideals in `𝓞 K` divisible par `J` of norm `≤ s` +multiplied by the order of the torsion of `K` is equal to the number of `idealPoint K J` +of norm `≤ s`. -/ +theorem card_isPrincipal_dvd_norm_le (s : ℝ) : + Nat.card {I : (Ideal (𝓞 K))⁰ // (J : Ideal (𝓞 K)) ∣ I ∧ IsPrincipal (I : Ideal (𝓞 K)) ∧ absNorm (I : Ideal (𝓞 K)) ≤ s} * torsionOrder K = - Nat.card {a : integralPoint K | intNorm a ≤ s} := by + Nat.card {a : idealPoint K J // mixedEmbedding.norm (a : mixedSpace K) ≤ s} := by obtain hs | hs := le_or_gt 0 s - · simp_rw [← Nat.le_floor_iff hs] + · simp_rw [← intNorm_idealPointEquiv_apply, ← Nat.le_floor_iff hs] rw [torsionOrder, PNat.mk_coe, ← Nat.card_eq_fintype_card, ← Nat.card_prod] - refine Nat.card_congr <| @Equiv.ofFiberEquiv _ (γ := Finset.Iic _) _ - (fun I ↦ ⟨absNorm (I.1 : Ideal (𝓞 K)), Finset.mem_Iic.mpr I.1.2.2⟩) - (fun a ↦ ⟨intNorm a.1, Finset.mem_Iic.mpr a.2⟩) fun ⟨i, hi⟩ ↦ ?_ + refine Nat.card_congr <| @Equiv.ofFiberEquiv _ (γ := Finset.Iic ⌊s⌋₊) _ + (fun I ↦ ⟨absNorm I.1.val.1, Finset.mem_Iic.mpr I.1.prop.2.2⟩) + (fun a ↦ ⟨intNorm (idealPointEquiv K J a.1).1, Finset.mem_Iic.mpr a.prop⟩) fun ⟨i, hi⟩ ↦ ?_ simp_rw [Subtype.mk.injEq] calc - _ ≃ {I : {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1 ∧ absNorm I.1 ≤ _} // absNorm I.1.1 = i} - × torsion K := Equiv.prodSubtypeFstEquivSubtypeProd - _ ≃ {I : (Ideal (𝓞 K))⁰ // (IsPrincipal I.1 ∧ absNorm I.1 ≤ _) ∧ absNorm I.1 = i} - × torsion K := - Equiv.prodCongrLeft fun _ ↦ (Equiv.subtypeSubtypeEquivSubtypeInter - (p := fun I : (Ideal (𝓞 K))⁰ ↦ IsPrincipal I.1 ∧ absNorm I.1 ≤ _) - (q := fun I ↦ absNorm I.1 = i)) - _ ≃ {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1 ∧ absNorm I.1 = i ∧ absNorm I.1 ≤ _} - × torsion K := - Equiv.prodCongrLeft fun _ ↦ (Equiv.subtypeEquivRight fun _ ↦ by aesop) - _ ≃ {I : (Ideal (𝓞 K))⁰ // IsPrincipal I.1 ∧ absNorm I.1 = i} × torsion K := - Equiv.prodCongrLeft fun _ ↦ (Equiv.subtypeEquivRight fun _ ↦ by - rw [and_iff_left_of_imp (a := absNorm _ = _) fun h ↦ Finset.mem_Iic.mp (h ▸ hi)]) - _ ≃ {a : integralPoint K // intNorm a = i} := (integralPointEquivNorm K i).symm - _ ≃ {a : {a : integralPoint K // intNorm a ≤ _} // intNorm a.1 = i} := - (Equiv.subtypeSubtypeEquivSubtype fun h ↦ Finset.mem_Iic.mp (h ▸ hi)).symm - · simp_rw [lt_iff_not_le.mp (lt_of_lt_of_le hs (Nat.cast_nonneg _)), and_false, Set.setOf_false, - Nat.card_eq_fintype_card, Fintype.card_ofIsEmpty, zero_mul] + _ ≃ {I : {I : (Ideal (𝓞 K))⁰ // _ ∧ _ ∧ _} // absNorm I.1.1 = i} × torsion K := + Equiv.prodSubtypeFstEquivSubtypeProd + _ ≃ {I : (Ideal (𝓞 K))⁰ // (_ ∧ _ ∧ absNorm I.1 ≤ ⌊s⌋₊) ∧ absNorm I.1 = i} + × torsion K := Equiv.prodCongrLeft fun _ ↦ (Equiv.subtypeSubtypeEquivSubtypeInter + (p := fun I : (Ideal (𝓞 K))⁰ ↦ J.1 ∣ I.1 ∧ IsPrincipal I.1 ∧ absNorm I.1 ≤ ⌊s⌋₊) + (q := fun I ↦ absNorm I.1 = i)) + _ ≃ {I : (Ideal (𝓞 K))⁰ // J.1 ∣ I.1 ∧ IsPrincipal I.1 ∧ absNorm I.1 = i} + × torsion K := Equiv.prodCongrLeft fun _ ↦ Equiv.subtypeEquivRight fun _ ↦ by aesop + _ ≃ {a : idealPoint K J // mixedEmbedding.norm (a : mixedSpace K) = i} := + (idealPointEquivNorm K J i).symm + _ ≃ {a : idealPoint K J // intNorm (idealPointEquiv K J a).1 = i} := by + simp_rw [← intNorm_idealPointEquiv_apply, Nat.cast_inj] + rfl + _ ≃ {b : {a : idealPoint K J // intNorm (idealPointEquiv K J a).1 ≤ ⌊s⌋₊} // + intNorm (idealPointEquiv K J b).1 = i} := + (Equiv.subtypeSubtypeEquivSubtype fun h ↦ Finset.mem_Iic.mp (h ▸ hi)).symm + · simp_rw [lt_iff_not_le.mp (lt_of_lt_of_le hs (Nat.cast_nonneg _)), lt_iff_not_le.mp + (lt_of_lt_of_le hs (mixedEmbedding.norm_nonneg _)), and_false, Nat.card_of_isEmpty, + zero_mul] end fundamentalCone diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean new file mode 100644 index 0000000000000..a4cde6fece926 --- /dev/null +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean @@ -0,0 +1,1612 @@ +/- +Copyright (c) 2024 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.FundamentalCone +import Mathlib.NumberTheory.NumberField.Units.Regulator + +/-! +# Fundamental Cone + +In this file, we study the subset `NormLessThanOne` of the `fundamentalCone` defined as the +subset of elements `x` such that `mixedEmbedding.norm x ≤ 1`. + +Mainly, we want to prove that its frontier has volume zero and compute its volume. For this, we +follow mainly the strategy given in [D. Marcus, *Number Fields*][marcus1977number]. + +## Strategy of proof + +* `polarCoordMixedSpace` and `lintegral_eq_lintegral_polarCoordMixedSpace_symm` + +-/ + +variable (K : Type*) [Field K] + +open Bornology NumberField.InfinitePlace NumberField.mixedEmbedding NumberField.Units + NumberField.Units.dirichletUnitTheorem + +open scoped Real + +namespace NumberField.mixedEmbedding + +noncomputable section realSpace + +/-- DOCSTRING -/ +abbrev realSpace := InfinitePlace K → ℝ + +variable {K} + +/-- DOCSTRING -/ +def realToMixed : (realSpace K) →L[ℝ] (mixedSpace K) := ContinuousLinearMap.prod + (ContinuousLinearMap.pi fun w ↦ ContinuousLinearMap.proj w.val) + (ContinuousLinearMap.pi fun w ↦ Complex.ofRealCLM.comp (ContinuousLinearMap.proj w.val)) + +@[simp] +theorem realToMixed_apply_of_isReal (x :realSpace K) {w : InfinitePlace K} + (hw : IsReal w) : (realToMixed x).1 ⟨w, hw⟩ = x w := rfl + +@[simp] +theorem realToMixed_apply_of_isComplex (x : realSpace K) {w : InfinitePlace K} + (hw : IsComplex w) : (realToMixed x).2 ⟨w, hw⟩ = x w := rfl + +@[simp] +theorem normAtPlace_realToMixed (w : InfinitePlace K) (x : realSpace K) : + normAtPlace w (realToMixed x) = ‖x w‖ := by + obtain hw | hw := isReal_or_isComplex w + · simp [normAtPlace_apply_isReal hw, realToMixed] + · simp [normAtPlace_apply_isComplex hw, realToMixed] + +@[simp] +theorem norm_realToMixed [NumberField K] (x : realSpace K) : + mixedEmbedding.norm (realToMixed x) = ∏ w, ‖x w‖ ^ w.mult := + Finset.prod_congr rfl fun w _ ↦ by simp + +theorem pos_norm_realToMixed [NumberField K] {x : realSpace K} (hx : ∀ w, x w ≠ 0) : + 0 < mixedEmbedding.norm (realToMixed x) := by + rw [norm_realToMixed] + exact Finset.prod_pos fun w _ ↦ pow_pos (abs_pos.mpr (hx w)) _ + +theorem logMap_realToMixed [NumberField K] {x : realSpace K} + (hx : mixedEmbedding.norm (realToMixed x) = 1) : + logMap (realToMixed x) = fun w ↦ (mult w.val) * Real.log (x w.val) := by + ext + rw [logMap_apply_of_norm_one hx, normAtPlace_realToMixed, Real.norm_eq_abs, Real.log_abs] + +open Classical in +/-- DOCSTRING -/ +def mixedToReal (x : mixedSpace K) : realSpace K := + fun w ↦ if hw : IsReal w then x.1 ⟨w, hw⟩ else ‖x.2 ⟨w, not_isReal_iff_isComplex.mp hw⟩‖ + +@[simp] +theorem mixedToReal_apply_of_isReal (x : mixedSpace K) (w : {w // IsReal w}) : + mixedToReal x w.val = x.1 w := by + rw [mixedToReal, dif_pos] + +theorem mixedToReal_apply_of_isComplex (x : mixedSpace K) (w : {w // IsComplex w}) : + mixedToReal x w.val = ‖x.2 w‖ := by + rw [mixedToReal, dif_neg (not_isReal_iff_isComplex.mpr w.prop)] + +theorem mixedToReal_nonneg (x : mixedSpace K) (w : {w // IsComplex w}) : + 0 ≤ mixedToReal x w.val := by + rw [mixedToReal_apply_of_isComplex] + exact norm_nonneg _ + +theorem mixedToReal_smul (x : mixedSpace K) {r : ℝ} (hr : 0 ≤ r) : + mixedToReal (r • x) = r • mixedToReal x := by + ext w + obtain hw | hw := isReal_or_isComplex w + · simp_rw [Pi.smul_apply, mixedToReal_apply_of_isReal _ ⟨w, hw⟩, Prod.smul_fst, Pi.smul_apply] + · simp_rw [Pi.smul_apply, mixedToReal_apply_of_isComplex _ ⟨w, hw⟩, Prod.smul_snd, Pi.smul_apply, + _root_.norm_smul, Real.norm_eq_abs, abs_of_nonneg hr, smul_eq_mul] + +open Classical in +theorem realToMixedToReal (x : realSpace K) : + mixedToReal (realToMixed x) = fun w ↦ if IsReal w then x w else ‖x w‖ := by + ext w + obtain hw | hw := isReal_or_isComplex w + · rw [mixedToReal_apply_of_isReal _ ⟨w, hw⟩, realToMixed_apply_of_isReal _ hw, if_pos hw] + · rw [mixedToReal_apply_of_isComplex _ ⟨w, hw⟩, if_neg (not_isReal_iff_isComplex.mpr hw), + realToMixed_apply_of_isComplex, Complex.norm_real, Real.norm_eq_abs] + +@[simp] +theorem realToMixedToReal_eq_self_of_nonneg {x : realSpace K} (hx : ∀ w, IsComplex w → 0 ≤ x w) : + mixedToReal (realToMixed x) = x := by + rw [realToMixedToReal] + ext w + obtain hw | hw := isReal_or_isComplex w + · rw [if_pos hw] + · rw [if_neg (not_isReal_iff_isComplex.mpr hw), Real.norm_eq_abs, abs_of_nonneg (hx w hw)] + +theorem mixedToRealToMixed (x : mixedSpace K) : + realToMixed (mixedToReal x) = (fun w ↦ x.1 w, fun w ↦ (‖x.2 w‖ : ℂ)) := by + ext w + · rw [realToMixed_apply_of_isReal, mixedToReal_apply_of_isReal] + · rw [realToMixed_apply_of_isComplex, mixedToReal_apply_of_isComplex] + +theorem norm_mixedToReal (x : mixedSpace K) (w : InfinitePlace K) : + ‖mixedToReal x w‖ = normAtPlace w x := by + obtain hw | hw := isReal_or_isComplex w + · rw [mixedToReal_apply_of_isReal _ ⟨w, hw⟩, normAtPlace_apply_isReal] + · rw [mixedToReal_apply_of_isComplex _ ⟨w, hw⟩, normAtPlace_apply_isComplex, norm_norm] + +theorem norm_mixedToRealToMixed [NumberField K] (x : mixedSpace K) : + mixedEmbedding.norm (realToMixed (mixedToReal x)) = mixedEmbedding.norm x := by + simp_rw [norm_realToMixed, norm_mixedToReal, mixedEmbedding.norm_apply] + +@[simp] +theorem logMap_mixedToRealToMixed_of_norm_one [NumberField K] {x : mixedSpace K} + (hx : mixedEmbedding.norm x = 1) : logMap (realToMixed (mixedToReal x)) = logMap x := by + ext + rw [logMap_apply_of_norm_one hx, logMap_apply_of_norm_one (by rwa [norm_mixedToRealToMixed]), + normAtPlace_realToMixed, ← norm_mixedToReal] + +open Classical in +theorem norm_realToMixed_prod_units_rpow [NumberField K] {ι : Type*} [Fintype ι] (c : ι → ℝ) + (u : ι → (𝓞 K)ˣ) : + mixedEmbedding.norm (realToMixed fun w : InfinitePlace K ↦ ∏ i, w (u i) ^ c i) = 1 := + calc + _ = |∏ w : InfinitePlace K, ∏ i, (w (u i) ^ c i) ^ w.mult| := by + simp_rw [norm_realToMixed, Real.norm_eq_abs, Finset.abs_prod, abs_pow, Finset.prod_pow] + _ = |∏ w : InfinitePlace K, ∏ i, (w (u i) ^ w.mult) ^ c i| := by + simp_rw [← Real.rpow_natCast, ← Real.rpow_mul (apply_nonneg _ _), mul_comm] + _ = |∏ i, (∏ w : InfinitePlace K, (w (u i) ^ w.mult)) ^ c i| := by + rw [Finset.prod_comm] + simp_rw [Real.finset_prod_rpow _ _ (fun _ _ ↦ pow_nonneg (apply_nonneg _ _) _)] + _ = 1 := by + simp_rw [prod_eq_abs_norm, Units.norm, Rat.cast_one, Real.one_rpow, + Finset.prod_const_one, abs_one] + +end realSpace + +noncomputable section polarCoord + +open MeasureTheory MeasureTheory.Measure MeasurableEquiv + +open Classical in +/-- DOCSTRING -/ +def realProdComplexProdMeasurableEquiv : + ({w : InfinitePlace K // IsReal w} → ℝ) × ({w : InfinitePlace K // IsComplex w} → ℝ × ℝ) ≃ᵐ + (realSpace K) × ({w : InfinitePlace K // IsComplex w} → ℝ) := + MeasurableEquiv.trans (prodCongr (refl _) + (arrowProdEquivProdArrow ℝ ℝ _)) <| + MeasurableEquiv.trans prodAssoc.symm <| + MeasurableEquiv.trans + (prodCongr (prodCongr (refl _) + (arrowCongr' (Equiv.subtypeEquivRight (fun _ ↦ not_isReal_iff_isComplex.symm)) (refl _))) + (refl _)) + (prodCongr (piEquivPiSubtypeProd (fun _ ↦ ℝ) _).symm (refl _)) + +open Classical in +/-- DOCSTRING -/ +def realProdComplexProdEquiv : + ({w : InfinitePlace K // IsReal w} → ℝ) × + ({w : InfinitePlace K // IsComplex w} → ℝ × ℝ) ≃ₜ + (realSpace K) × ({w : InfinitePlace K // IsComplex w} → ℝ) where + __ := realProdComplexProdMeasurableEquiv K + continuous_toFun := by + change Continuous fun x ↦ ⟨fun w ↦ if hw : w.IsReal then x.1 ⟨w, hw⟩ else + (x.2 ⟨w, not_isReal_iff_isComplex.mp hw⟩).1, fun w ↦ (x.2 w).2⟩ + refine continuous_prod_mk.mpr ⟨continuous_pi_iff.mpr fun w ↦ ?_, by fun_prop⟩ + by_cases hw : IsReal w + · simp_rw [dif_pos hw]; fun_prop + · simp_rw [dif_neg hw]; fun_prop + continuous_invFun := by + change Continuous fun x ↦ (fun w ↦ x.1 w.val, fun w ↦ ⟨x.1 w.val, x.2 w⟩) + fun_prop + +open Classical in +theorem volume_preserving_realProdComplexProdEquiv [NumberField K] : + MeasurePreserving (realProdComplexProdEquiv K) := by + change MeasurePreserving (realProdComplexProdMeasurableEquiv K) volume volume + exact MeasurePreserving.trans ((MeasurePreserving.id volume).prod + (volume_measurePreserving_arrowProdEquivProdArrow ℝ ℝ {w : InfinitePlace K // IsComplex w})) <| + MeasurePreserving.trans (volume_preserving_prodAssoc.symm) <| + MeasurePreserving.trans + (((MeasurePreserving.id volume).prod (volume_preserving_arrowCongr' _ + (MeasurableEquiv.refl ℝ) + (MeasurePreserving.id volume))).prod (MeasurePreserving.id volume)) + <| ((volume_preserving_piEquivPiSubtypeProd (fun _ : InfinitePlace K ↦ ℝ) + (fun w : InfinitePlace K ↦ IsReal w)).symm).prod (MeasurePreserving.id volume) + +open Classical in +@[simp] +theorem realProdComplexProdEquiv_apply (x : ({w : InfinitePlace K // IsReal w} → ℝ) × + ({w : InfinitePlace K // IsComplex w} → ℝ × ℝ)) : + realProdComplexProdEquiv K x = ⟨fun w ↦ if hw : w.IsReal then x.1 ⟨w, hw⟩ else + (x.2 ⟨w, not_isReal_iff_isComplex.mp hw⟩).1, fun w ↦ (x.2 w).2⟩ := rfl + +@[simp] +theorem realProdComplexProdEquiv_symm_apply (x : (realSpace K) × + ({w : InfinitePlace K // IsComplex w} → ℝ)) : + (realProdComplexProdEquiv K).symm x = (fun w ↦ x.1 w.val, fun w ↦ ⟨x.1 w.val, x.2 w⟩) := rfl + +variable [NumberField K] + +/-- DOCSTRING -/ +def polarCoordMixedSpace : PartialHomeomorph + (mixedSpace K) ((realSpace K) × ({w : InfinitePlace K // IsComplex w} → ℝ)) := + ((PartialHomeomorph.refl _).prod + (PartialHomeomorph.pi fun _ ↦ Complex.polarCoord)).transHomeomorph (realProdComplexProdEquiv K) + +theorem polarCoordMixedSpace_source : + (polarCoordMixedSpace K).source = Set.univ ×ˢ Set.univ.pi fun _ ↦ Complex.slitPlane := by + simp [polarCoordMixedSpace, Complex.polarCoord_source] + +open Classical in +theorem polarCoordMixedSpace_target : (polarCoordMixedSpace K).target = + (Set.univ.pi fun w ↦ + if IsReal w then Set.univ else Set.Ioi 0) ×ˢ (Set.univ.pi fun _ ↦ Set.Ioo (-π) π):= by + rw [polarCoordMixedSpace, PartialHomeomorph.transHomeomorph_target] + ext + simp_rw [Set.mem_preimage, realProdComplexProdEquiv_symm_apply, PartialHomeomorph.prod_target, + Set.mem_prod, PartialHomeomorph.refl_target, PartialHomeomorph.pi_target, + Complex.polarCoord_target] + aesop + +theorem polarCoordMixedSpace_symm_apply (x : (realSpace K) × ({w // IsComplex w} → ℝ)) : + (polarCoordMixedSpace K).symm x = ⟨fun w ↦ x.1 w.val, + fun w ↦ Complex.polarCoord.symm (x.1 w, x.2 w)⟩ := rfl + +theorem measurableSet_polarCoordMixedSpace_target : + MeasurableSet (polarCoordMixedSpace K).target := + (polarCoordMixedSpace K).open_target.measurableSet + +theorem polarCoordMixedSpace_apply (x : mixedSpace K) : + polarCoordMixedSpace K x = + (realProdComplexProdEquiv K) (x.1, fun w ↦ Complex.polarCoord (x.2 w)) := by + rw [polarCoordMixedSpace] + simp_rw [PartialHomeomorph.transHomeomorph_apply, PartialHomeomorph.prod_apply, + PartialHomeomorph.refl_apply, id_eq, Function.comp_apply] + rfl + +theorem continuous_polarCoordMixedSpace_symm : + Continuous (polarCoordMixedSpace K).symm := by + change Continuous (fun x ↦ (polarCoordMixedSpace K).symm x) + simp_rw [polarCoordMixedSpace_symm_apply] + rw [continuous_prod_mk] + exact ⟨by fun_prop, + continuous_pi_iff.mpr fun i ↦ Complex.continuous_polarCoord_symm.comp' (by fun_prop)⟩ + +theorem realProdComplexProdEquiv_preimage_target : + (realProdComplexProdEquiv K) ⁻¹' (polarCoordMixedSpace K).target = + Set.univ ×ˢ Set.univ.pi fun _ ↦ polarCoord.target := by + ext + simp_rw [polarCoordMixedSpace_target, Set.mem_preimage, realProdComplexProdEquiv_apply, + polarCoord_target, Set.mem_prod, Set.mem_pi, Set.mem_univ, true_implies, true_and, + Set.mem_ite_univ_left, not_isReal_iff_isComplex, Set.mem_prod] + refine ⟨fun ⟨h₁, h₂⟩ i ↦ ⟨?_, h₂ i⟩, fun h ↦ ⟨fun i hi ↦ ?_, fun i ↦ (h i).2⟩⟩ + · specialize h₁ i i.prop + rwa [dif_neg (not_isReal_iff_isComplex.mpr i.prop)] at h₁ + · rw [dif_neg (not_isReal_iff_isComplex.mpr hi)] + exact (h ⟨i, hi⟩).1 + +open Classical in +theorem lintegral_eq_lintegral_polarCoordMixedSpace_symm (f : (mixedSpace K) → ENNReal) + (hf : Measurable f) : + ∫⁻ x, f x = + ∫⁻ x in (polarCoordMixedSpace K).target, + (∏ w : {w // IsComplex w}, (x.1 w.val).toNNReal) * + f ((polarCoordMixedSpace K).symm x) := by + have h : Measurable fun x ↦ (∏ w : { w // IsComplex w}, (x.1 w.val).toNNReal) * + f ((polarCoordMixedSpace K).symm x) := by + refine Measurable.mul ?_ ?_ + · exact measurable_coe_nnreal_ennreal_iff.mpr <| Finset.measurable_prod _ fun _ _ ↦ by fun_prop + · exact hf.comp' (continuous_polarCoordMixedSpace_symm K).measurable + rw [← (volume_preserving_realProdComplexProdEquiv K).setLIntegral_comp_preimage + (measurableSet_polarCoordMixedSpace_target K) h, volume_eq_prod, volume_eq_prod, + lintegral_prod _ hf.aemeasurable] + simp_rw [Complex.lintegral_pi_comp_polarCoord_symm _ (hf.comp' measurable_prod_mk_left)] + rw [realProdComplexProdEquiv_preimage_target, ← Measure.restrict_prod_eq_univ_prod, + lintegral_prod _ (h.comp' (realProdComplexProdEquiv K).measurable).aemeasurable] + simp_rw [realProdComplexProdEquiv_apply, polarCoordMixedSpace_symm_apply, + dif_pos (Subtype.prop _), dif_neg (not_isReal_iff_isComplex.mpr (Subtype.prop _))] + +end polarCoord + +noncomputable section mapToUnitsPow + +open Module Finset + +variable [NumberField K] + +open Classical in +/-- DOCSTRING -/ +-- This cannot be a `PartiaHomeomorph` because the target is not an open set +def mapToUnitsPow₀_aux : + PartialEquiv ({w : InfinitePlace K // w ≠ w₀} → ℝ) (realSpace K) where + toFun := fun c w ↦ if hw : w = w₀ then + Real.exp (- ((w₀ : InfinitePlace K).mult : ℝ)⁻¹ * ∑ w : {w // w ≠ w₀}, c w) + else Real.exp ((w.mult : ℝ)⁻¹ * c ⟨w, hw⟩) + invFun := fun x w ↦ w.val.mult * Real.log (x w.val) + source := Set.univ + target := {x | ∀ w, 0 < x w} ∩ {x | ∑ w, w.mult * Real.log (x w) = 0} + map_source' _ _ := by + dsimp only + refine ⟨Set.mem_setOf.mpr fun w ↦ by split_ifs <;> exact Real.exp_pos _, ?_⟩ + simp_rw [Set.mem_setOf_eq, ← Finset.univ.sum_erase_add _ (mem_univ w₀), dif_pos] + rw [sum_subtype _ (by aesop : ∀ w, w ∈ univ.erase w₀ ↔ w ≠ w₀)] + conv_lhs => enter [1,2,w]; rw [dif_neg w.prop] + simp_rw [Real.log_exp, neg_mul, mul_neg, mul_inv_cancel_left₀ mult_coe_ne_zero, + add_neg_eq_zero] + infer_instance + map_target' _ _ := trivial + left_inv' := by + intro x _ + dsimp only + ext w + rw [dif_neg w.prop, Real.log_exp, mul_inv_cancel_left₀ mult_coe_ne_zero] + right_inv' := by + intro x hx + ext w + dsimp only + by_cases hw : w = w₀ + · rw [hw, dif_pos rfl, ← sum_subtype _ + (by aesop : ∀ w, w ∈ univ.erase w₀ ↔ w ≠ w₀) (fun w ↦ w.mult * Real.log (x w))] + rw [sum_erase_eq_sub (mem_univ w₀), hx.2, _root_.zero_sub, neg_mul, mul_neg, + neg_neg, inv_mul_cancel_left₀ mult_coe_ne_zero, Real.exp_log (hx.1 w₀)] + · rw [dif_neg hw, inv_mul_cancel_left₀ mult_coe_ne_zero, Real.exp_log (hx.1 w)] + +theorem mapToUnitsPow₀_aux_symm_apply (x : realSpace K) : + (mapToUnitsPow₀_aux K).symm x = fun w ↦ w.val.mult * Real.log (x w.val) := rfl + +theorem continuous_mapToUnitsPow₀_aux : + Continuous (mapToUnitsPow₀_aux K) := by + unfold mapToUnitsPow₀_aux + refine continuous_pi_iff.mpr fun w ↦ ?_ + by_cases hw : w = w₀ + · simp_rw [dif_pos hw] + fun_prop + · simp_rw [dif_neg hw] + fun_prop + +variable {K} + +/-- DOCSTRING -/ +def equivFinRank : Fin (rank K) ≃ {w : InfinitePlace K // w ≠ w₀} := by + classical + refine Fintype.equivOfCardEq ?_ + rw [Fintype.card_subtype_compl, Fintype.card_ofSubsingleton, Fintype.card_fin, rank] + +variable (K) + +open Classical in +/-- DOCSTRING -/ +-- This cannot be a `PartiaHomeomorph` because the target is not an open set +def mapToUnitsPow₀ : + PartialEquiv ({w : InfinitePlace K // w ≠ w₀} → ℝ) (realSpace K) := + (((basisUnitLattice K).ofZLatticeBasis ℝ).reindex + equivFinRank).equivFun.symm.toEquiv.toPartialEquiv.trans (mapToUnitsPow₀_aux K) + +theorem mapToUnitsPow₀_source : + (mapToUnitsPow₀ K).source = Set.univ := by simp [mapToUnitsPow₀, mapToUnitsPow₀_aux] + +theorem mapToUnitsPow₀_target : + (mapToUnitsPow₀ K).target = {x | (∀ w, 0 < x w) ∧ mixedEmbedding.norm (realToMixed x) = 1} := by + rw [mapToUnitsPow₀, mapToUnitsPow₀_aux] + dsimp + ext x + simp_rw [Set.inter_univ, Set.mem_inter_iff, Set.mem_setOf, and_congr_right_iff] + intro hx + rw [← Real.exp_injective.eq_iff, Real.exp_zero, Real.exp_sum, norm_realToMixed] + refine Eq.congr (Finset.prod_congr rfl fun _ _ ↦ ?_) rfl + rw [← Real.log_rpow (hx _), Real.exp_log (Real.rpow_pos_of_pos (hx _) _), Real.norm_eq_abs, + abs_of_pos (hx _), Real.rpow_natCast] + +variable {K} + +theorem mixedToReal_mem_target {x : mixedSpace K} (hx₁ : ∀ w, 0 < x.1 w) + (hx₂ : mixedEmbedding.norm x = 1): + mixedToReal x ∈ (mapToUnitsPow₀ K).target := by + rw [mapToUnitsPow₀_target] + refine ⟨fun w ↦ ?_, by rwa [norm_mixedToRealToMixed]⟩ + obtain hw | hw := isReal_or_isComplex w + · rw [mixedToReal_apply_of_isReal _ ⟨w, hw⟩] + exact hx₁ _ + · refine lt_of_le_of_ne' (mixedToReal_nonneg _ ⟨w, hw⟩) ?_ + contrapose! hx₂ + rw [mixedToReal_apply_of_isComplex _ ⟨w, hw⟩] at hx₂ + rw [mixedEmbedding.norm_eq_zero_iff.mpr ⟨w, by rwa [normAtPlace_apply_isComplex hw x]⟩] + exact zero_ne_one + +theorem norm_mapToUnitsPow₀ (c : {w : InfinitePlace K // w ≠ w₀} → ℝ) : + mixedEmbedding.norm (realToMixed (mapToUnitsPow₀ K c)) = 1 := by + suffices mapToUnitsPow₀ K c ∈ (mapToUnitsPow₀ K).target by + rw [mapToUnitsPow₀_target] at this + exact this.2 + refine PartialEquiv.map_source (mapToUnitsPow₀ K) ?_ + rw [mapToUnitsPow₀_source] + exact trivial + +theorem mapToUnitsPow₀_pos (c : {w : InfinitePlace K // w ≠ w₀} → ℝ) (w : InfinitePlace K) : + 0 < mapToUnitsPow₀ K c w := by + suffices mapToUnitsPow₀ K c ∈ (mapToUnitsPow₀ K).target by + rw [mapToUnitsPow₀_target] at this + exact this.1 w + refine PartialEquiv.map_source (mapToUnitsPow₀ K) ?_ + rw [mapToUnitsPow₀_source] + exact trivial + +open Classical in +theorem mapToUnitsPow₀_symm_prod_fundSystem_rpow (c : {w : InfinitePlace K // w ≠ w₀} → ℝ) : + (mapToUnitsPow₀ K).symm (fun w ↦ ∏ i, w (fundSystem K (equivFinRank.symm i)) ^ c i) = c := by + ext + simp_rw [mapToUnitsPow₀, PartialEquiv.coe_trans_symm, Equiv.toPartialEquiv_symm_apply, + LinearEquiv.coe_toEquiv_symm, LinearEquiv.symm_symm, Function.comp_apply, + mapToUnitsPow₀_aux_symm_apply, EquivLike.coe_coe, Basis.equivFun_apply, Basis.repr_reindex, + Real.log_prod _ _ (fun _ _ ↦ (Real.rpow_pos_of_pos (Units.pos_at_place _ _) _).ne'), + Real.log_rpow (Units.pos_at_place _ _), mul_sum, mul_left_comm, ← logEmbedding_component, + logEmbedding_fundSystem, ← sum_fn, _root_.map_sum, ← smul_eq_mul, ← Pi.smul_def, + _root_.map_smul, Finsupp.mapDomain_equiv_apply, sum_apply', Finsupp.coe_smul, Pi.smul_apply, + Basis.ofZLatticeBasis_repr_apply, Basis.repr_self, smul_eq_mul, Finsupp.single_apply, + Int.cast_ite, mul_ite, Int.cast_zero, mul_zero, EmbeddingLike.apply_eq_iff_eq, sum_ite_eq', + mem_univ, if_true, Int.cast_one, mul_one] + +open Classical in +theorem mapToUnitsPow₀_apply (c : {w : InfinitePlace K // w ≠ w₀} → ℝ) : + mapToUnitsPow₀ K c = fun w ↦ ∏ i, w (fundSystem K (equivFinRank.symm i)) ^ c i := by + refine (PartialEquiv.eq_symm_apply _ (by trivial) ?_).mp + (mapToUnitsPow₀_symm_prod_fundSystem_rpow c).symm + rw [mapToUnitsPow₀_target] + refine ⟨?_, norm_realToMixed_prod_units_rpow c _⟩ + exact fun _ ↦ Finset.prod_pos fun _ _ ↦ Real.rpow_pos_of_pos (Units.pos_at_place _ _) _ + +open Classical in +theorem mapToUnitsPow₀_symm_apply_of_norm_one {x : InfinitePlace K → ℝ} + (hx : mixedEmbedding.norm (realToMixed x) = 1) : + (mapToUnitsPow₀ K).symm x = (((basisUnitLattice K).ofZLatticeBasis ℝ).reindex + equivFinRank).equivFun (logMap (realToMixed x)) := by + simp_rw [mapToUnitsPow₀, PartialEquiv.coe_trans_symm, Equiv.toPartialEquiv_symm_apply, + LinearEquiv.coe_toEquiv_symm, LinearEquiv.symm_symm, EquivLike.coe_coe, Function.comp_apply] + congr with x + rw [logMap_apply_of_norm_one hx, mapToUnitsPow₀_aux, PartialEquiv.coe_symm_mk, + normAtPlace_realToMixed, Real.norm_eq_abs, Real.log_abs] + +variable (K) in +open Classical in +theorem continuous_mapToUnitsPow₀ : + Continuous (mapToUnitsPow₀ K) := (continuous_mapToUnitsPow₀_aux K).comp <| + LinearEquiv.continuous_symm _ (continuous_equivFun_basis _) + +open Classical in +/-- DOCSTRING -/ +abbrev mapToUnitsPow_single (c : realSpace K) : InfinitePlace K → (realSpace K) := + fun i ↦ if hi : i = w₀ then fun _ ↦ |c w₀| else + fun w ↦ (w (fundSystem K (equivFinRank.symm ⟨i, hi⟩))) ^ (c i) + +open Classical in +theorem mapToUnitsPow₀_eq_prod_single (c : realSpace K) (w : InfinitePlace K) : + mapToUnitsPow₀ K (fun w ↦ c w.val) w = + ∏ i ∈ univ.erase w₀, mapToUnitsPow_single c i w := by + rw [mapToUnitsPow₀_apply, Finset.prod_subtype (Finset.univ.erase w₀) + (fun w ↦ (by aesop : w ∈ univ.erase w₀ ↔ w ≠ w₀))] + exact Finset.prod_congr rfl fun w _ ↦ by rw [mapToUnitsPow_single, dif_neg w.prop] + +theorem prod_mapToUnitsPow_single (c : realSpace K) : + ∏ i, mapToUnitsPow_single c i = |c w₀| • mapToUnitsPow₀ K (fun w ↦ c w.val) := by + classical + ext + rw [Pi.smul_apply, mapToUnitsPow₀_eq_prod_single, ← Finset.univ.mul_prod_erase _ + (Finset.mem_univ w₀), mapToUnitsPow_single, dif_pos rfl, smul_eq_mul, Pi.mul_apply, prod_apply] + +variable (K) + +open Classical in +/-- DOCSTRING -/ +@[simps source target] +def mapToUnitsPow : PartialHomeomorph (realSpace K) (realSpace K) where + toFun c := ∏ i, mapToUnitsPow_single c i + invFun x w := + if hw : w = w₀ then mixedEmbedding.norm (realToMixed x) ^ (finrank ℚ K : ℝ)⁻¹ else + (((basisUnitLattice K).ofZLatticeBasis ℝ).reindex + equivFinRank).equivFun (logMap (realToMixed x)) ⟨w, hw⟩ + source := {x | 0 < x w₀} + target := {x | ∀ w, 0 < x w} + map_source' _ h _ := by + simp_rw [prod_mapToUnitsPow_single, Pi.smul_apply, smul_eq_mul] + exact mul_pos (abs_pos.mpr h.ne') (mapToUnitsPow₀_pos _ _) + map_target' x hx := by + refine Set.mem_setOf.mpr ?_ + dsimp only + rw [dif_pos rfl] + exact Real.rpow_pos_of_pos (pos_norm_realToMixed (fun w ↦ (hx w).ne')) _ + left_inv' _ h := by + dsimp only + ext w + by_cases hw : w = w₀ + · rw [hw, dif_pos rfl, prod_mapToUnitsPow_single, map_smul, mixedEmbedding.norm_smul, + norm_mapToUnitsPow₀, mul_one, ← Real.rpow_natCast, ← Real.rpow_mul (abs_nonneg _), + mul_inv_cancel₀ (Nat.cast_ne_zero.mpr (finrank_pos).ne'), Real.rpow_one, abs_of_nonneg + (abs_nonneg _), abs_of_pos (by convert h)] + · rw [dif_neg hw, prod_mapToUnitsPow_single, map_smul, logMap_real_smul + (by rw [norm_mapToUnitsPow₀]; exact one_ne_zero) (abs_ne_zero.mpr (ne_of_gt h)), + ← mapToUnitsPow₀_symm_apply_of_norm_one, PartialEquiv.left_inv _ + (by rw [mapToUnitsPow₀_source]; trivial)] + rw [mapToUnitsPow₀_apply] + exact norm_realToMixed_prod_units_rpow _ _ + right_inv' x hx := by + have h₀ : mixedEmbedding.norm + (realToMixed (mixedEmbedding.norm (realToMixed x) ^ (- (finrank ℚ K : ℝ)⁻¹) • x)) = 1 := by + rw [map_smul, norm_smul, ← abs_pow, ← Real.rpow_natCast, ← Real.rpow_mul, neg_mul, + inv_mul_cancel₀, Real.rpow_neg_one, abs_of_nonneg, inv_mul_cancel₀] + · rw [mixedEmbedding.norm_ne_zero_iff] + intro w + rw [normAtPlace_realToMixed, Real.norm_eq_abs, abs_ne_zero] + exact (hx w).ne' + refine inv_nonneg.mpr (mixedEmbedding.norm_nonneg _) + rw [Nat.cast_ne_zero] + exact finrank_pos.ne' + exact mixedEmbedding.norm_nonneg _ + have hx' : ∀ w, x w ≠ 0 := fun w ↦ (hx w).ne' + dsimp only + rw [prod_mapToUnitsPow_single, dif_pos rfl] + conv_lhs => + enter [2, 2, w] + rw [dif_neg w.prop] + ext w + rw [Pi.smul_apply, ← logMap_real_smul] + rw [← _root_.map_smul] + rw [← mapToUnitsPow₀_symm_apply_of_norm_one h₀] + rw [PartialEquiv.right_inv, Pi.smul_apply, smul_eq_mul, smul_eq_mul] + rw [abs_of_nonneg, Real.rpow_neg, mul_inv_cancel_left₀] + · refine Real.rpow_ne_zero_of_pos ?_ _ + exact pos_norm_realToMixed hx' + · exact mixedEmbedding.norm_nonneg (realToMixed x) + · refine Real.rpow_nonneg ?_ _ + exact mixedEmbedding.norm_nonneg (realToMixed x) + · rw [mapToUnitsPow₀_target] + refine ⟨fun w ↦ ?_, h₀⟩ + exact mul_pos (Real.rpow_pos_of_pos (pos_norm_realToMixed hx') _) (hx w) + · exact ne_of_gt (pos_norm_realToMixed hx') + · refine Real.rpow_ne_zero_of_pos ?_ _ + exact pos_norm_realToMixed hx' + open_source := isOpen_lt continuous_const (continuous_apply w₀) + open_target := by + convert_to IsOpen (⋂ w, {x : InfinitePlace K → ℝ | 0 < x w}) + · ext; simp + · exact isOpen_iInter_of_finite fun w ↦ isOpen_lt continuous_const (continuous_apply w) + continuousOn_toFun := by + simp_rw [prod_mapToUnitsPow_single] + exact ContinuousOn.smul (by fun_prop) <| + (continuous_mapToUnitsPow₀ K).comp_continuousOn' (by fun_prop) + continuousOn_invFun := by + simp only + rw [continuousOn_pi] + intro w + by_cases hw : w = w₀ + · simp_rw [hw, dite_true] + refine Continuous.continuousOn ?_ + refine Continuous.rpow_const ?_ ?_ + · refine Continuous.comp' ?_ ?_ + exact mixedEmbedding.continuous_norm K + exact ContinuousLinearMap.continuous realToMixed + · intro _ + right + rw [inv_nonneg] + exact Nat.cast_nonneg' (finrank ℚ K) + · simp_rw [dif_neg hw] + refine Continuous.comp_continuousOn' (continuous_apply _) <| + (continuous_equivFun_basis _).comp_continuousOn' ?_ + refine ContinuousOn.comp'' (continuousOn_logMap K) ?_ ?_ + refine Continuous.continuousOn ?_ + exact ContinuousLinearMap.continuous realToMixed + intro x hx + refine ne_of_gt ?_ + exact pos_norm_realToMixed (fun w ↦ (hx w).ne') + +theorem measurable_mapToUnitsPow_symm : + Measurable (mapToUnitsPow K).symm := by + classical + dsimp [mapToUnitsPow, PartialHomeomorph.mk_coe_symm, PartialEquiv.coe_symm_mk] + refine measurable_pi_iff.mpr fun _ ↦ ?_ + split_ifs + · refine Measurable.pow_const ?_ _ + exact Measurable.comp' (mixedEmbedding.continuous_norm K).measurable realToMixed.measurable + · exact Measurable.eval <| + (continuous_equivFun_basis ((Basis.ofZLatticeBasis ℝ (unitLattice K) + (basisUnitLattice K)).reindex equivFinRank)).measurable.comp' + ((measurable_logMap _).comp' realToMixed.measurable) + +variable {K} + +theorem mapToUnitsPow_apply (c : realSpace K) : + mapToUnitsPow K c = ∏ i, mapToUnitsPow_single c i := rfl + +/-- Docstring. -/ +theorem mapToUnitsPow_apply' (c : realSpace K) : + mapToUnitsPow K c = |c w₀| • mapToUnitsPow₀ K (fun w ↦ c w.val) := by + rw [mapToUnitsPow_apply, prod_mapToUnitsPow_single] + +variable (K) in +theorem continuous_mapToUnitsPow : + Continuous (mapToUnitsPow K) := by + change Continuous (fun c ↦ mapToUnitsPow K c) + simp_rw [mapToUnitsPow_apply'] + exact Continuous.smul (by fun_prop) ((continuous_mapToUnitsPow₀ K).comp' (by fun_prop)) + +theorem mapToUnitsPow_nonneg (c : realSpace K) (w : InfinitePlace K) : + 0 ≤ mapToUnitsPow K c w := by + rw [mapToUnitsPow_apply'] + exact mul_nonneg (abs_nonneg _) (mapToUnitsPow₀_pos _ _).le + +theorem mapToUnitsPow_zero_iff {c : realSpace K} : + mapToUnitsPow K c = 0 ↔ c w₀ = 0 := by + rw [mapToUnitsPow_apply', smul_eq_zero, abs_eq_zero, or_iff_left] + obtain ⟨w⟩ := (inferInstance : Nonempty (InfinitePlace K)) + refine Function.ne_iff.mpr ⟨w, ?_⟩ + convert (mapToUnitsPow₀_pos (fun i ↦ c i) w).ne' + +/-- Docstring. -/ +theorem mapToUnitsPow_zero_iff' {c : InfinitePlace K → ℝ} {w : InfinitePlace K} : + mapToUnitsPow K c w = 0 ↔ c w₀ = 0 := by + rw [mapToUnitsPow_apply', Pi.smul_apply, smul_eq_mul, mul_eq_zero, abs_eq_zero, + or_iff_left (ne_of_gt (mapToUnitsPow₀_pos _ _))] + +theorem mapToUnitsPow_pos {c : InfinitePlace K → ℝ} (hc : c w₀ ≠ 0) (w : InfinitePlace K) : + 0 < mapToUnitsPow K c w := + lt_of_le_of_ne' (mapToUnitsPow_nonneg c w) (mapToUnitsPow_zero_iff'.not.mpr hc) + +open ContinuousLinearMap + +/-- DOCSTRING -/ +abbrev mapToUnitsPow_fDeriv_single (c : realSpace K) (i w : InfinitePlace K) : + (realSpace K) →L[ℝ] ℝ := by + classical + exact if hi : i = w₀ then proj w₀ else + (w (fundSystem K (equivFinRank.symm ⟨i, hi⟩)) ^ c i * + (w (fundSystem K (equivFinRank.symm ⟨i, hi⟩))).log) • proj i + +theorem hasFDeriv_mapToUnitsPow_single (c : realSpace K) (i w : InfinitePlace K) + (hc : 0 ≤ c w₀) : + HasFDerivWithinAt (fun x ↦ mapToUnitsPow_single x i w) + (mapToUnitsPow_fDeriv_single c i w) {x | 0 ≤ x w₀} c := by + unfold mapToUnitsPow_single mapToUnitsPow_fDeriv_single + split_ifs + · refine HasFDerivWithinAt.congr' (hasFDerivWithinAt_apply w₀ c _) ?_ hc + exact fun _ h ↦ by simp_rw [abs_of_nonneg (Set.mem_setOf.mp h)] + · exact HasFDerivWithinAt.const_rpow (hasFDerivWithinAt_apply i c _) <| pos_iff.mpr (by aesop) + +open Classical in +/-- DOCSTRING -/ +abbrev mapToUnitsPow_jacobianCoeff (w i : InfinitePlace K) : (realSpace K) → ℝ := + fun c ↦ if hi : i = w₀ then 1 else |c w₀| * (w (fundSystem K (equivFinRank.symm ⟨i, hi⟩))).log + +/-- DOCSTRING -/ +abbrev mapToUnitsPow_jacobian (c : realSpace K) : (realSpace K) →L[ℝ] InfinitePlace K → ℝ := + pi fun i ↦ (mapToUnitsPow₀ K (fun w ↦ c w) i • + ∑ w, (mapToUnitsPow_jacobianCoeff i w c) • proj w) + +theorem hasFDeriv_mapToUnitsPow (c : InfinitePlace K → ℝ) (hc : 0 ≤ c w₀) : + HasFDerivWithinAt (mapToUnitsPow K) (mapToUnitsPow_jacobian c) {x | 0 ≤ x w₀} c := by + classical + refine hasFDerivWithinAt_pi'.mpr fun w ↦ ?_ + simp_rw [mapToUnitsPow_apply, Finset.prod_apply] + convert HasFDerivWithinAt.finset_prod fun i _ ↦ hasFDeriv_mapToUnitsPow_single c i w hc + rw [ContinuousLinearMap.proj_pi, Finset.smul_sum] + refine Fintype.sum_congr _ _ fun i ↦ ?_ + by_cases hi : i = w₀ + · simp_rw [hi, mapToUnitsPow_jacobianCoeff, dite_true, one_smul, dif_pos, + mapToUnitsPow₀_eq_prod_single] + · rw [mapToUnitsPow₀_eq_prod_single, mapToUnitsPow_jacobianCoeff, dif_neg hi, smul_smul, + ← mul_assoc, show |c w₀| = mapToUnitsPow_single c w₀ w by simp_rw [dif_pos rfl], + Finset.prod_erase_mul _ _ (Finset.mem_univ w₀), mapToUnitsPow_fDeriv_single, dif_neg hi, + smul_smul, ← mul_assoc, show w (algebraMap (𝓞 K) K + (fundSystem K (equivFinRank.symm ⟨i, hi⟩))) ^ c i = mapToUnitsPow_single c i w by + simp_rw [dif_neg hi], Finset.prod_erase_mul _ _ (Finset.mem_univ i)] + +open Classical in +theorem prod_mapToUnitsPow₀(c : {w : InfinitePlace K // w ≠ w₀} → ℝ) : + ∏ w : InfinitePlace K, mapToUnitsPow₀ K c w = + (∏ w : {w : InfinitePlace K // IsComplex w}, mapToUnitsPow₀ K c w)⁻¹ := by + have : ∏ w : { w // IsComplex w}, (mapToUnitsPow₀ K) c w.val ≠ 0 := + Finset.prod_ne_zero_iff.mpr fun w _ ↦ ne_of_gt (mapToUnitsPow₀_pos c w) + rw [← mul_eq_one_iff_eq_inv₀ this] + convert norm_mapToUnitsPow₀ c + simp_rw [norm_realToMixed, ← Fintype.prod_subtype_mul_prod_subtype (fun w ↦ IsReal w)] + rw [← (Equiv.subtypeEquivRight (fun _ ↦ not_isReal_iff_isComplex)).prod_comp] + simp_rw [Equiv.subtypeEquivRight_apply_coe] + rw [mul_assoc, ← sq, ← Finset.prod_pow] + congr with w + · rw [Real.norm_eq_abs, abs_of_pos (mapToUnitsPow₀_pos c _), mult, if_pos w.prop, pow_one] + · rw [Real.norm_eq_abs, abs_of_pos (mapToUnitsPow₀_pos c _), mult, if_neg w.prop] + +open Classical in +theorem mapToUnitsPow_jacobian_det {c : realSpace K} (hc : 0 ≤ c w₀) : + |(mapToUnitsPow_jacobian c).det| = + (∏ w : {w : InfinitePlace K // w.IsComplex }, mapToUnitsPow₀ K (fun w ↦ c w) w)⁻¹ * + 2⁻¹ ^ NrComplexPlaces K * |c w₀| ^ (rank K) * (finrank ℚ K) * regulator K := by + have : LinearMap.toMatrix' (mapToUnitsPow_jacobian c).toLinearMap = + Matrix.of fun w i ↦ mapToUnitsPow₀ K (fun w ↦ c w) w * + mapToUnitsPow_jacobianCoeff w i c := by + ext + simp_rw [mapToUnitsPow_jacobian, ContinuousLinearMap.coe_pi, ContinuousLinearMap.coe_smul, + ContinuousLinearMap.coe_sum, LinearMap.toMatrix'_apply, LinearMap.pi_apply, + LinearMap.smul_apply, LinearMap.coeFn_sum, Finset.sum_apply, ContinuousLinearMap.coe_coe, + ContinuousLinearMap.coe_smul', Pi.smul_apply, ContinuousLinearMap.proj_apply, smul_eq_mul, + mul_ite, mul_one, mul_zero, Finset.sum_ite_eq', Finset.mem_univ, if_true, Matrix.of_apply] + rw [ContinuousLinearMap.det, ← LinearMap.det_toMatrix', this] + rw [Matrix.det_mul_column, prod_mapToUnitsPow₀, ← Matrix.det_transpose] + simp_rw [mapToUnitsPow_jacobianCoeff] + rw [mul_assoc, finrank_mul_regulator_eq_det K w₀ equivFinRank.symm] + have : |c w₀| ^ rank K = |∏ w : InfinitePlace K, if w = w₀ then 1 else c w₀| := by + rw [Finset.prod_ite, Finset.prod_const_one, Finset.prod_const, one_mul, abs_pow] + rw [Finset.filter_ne', Finset.card_erase_of_mem (Finset.mem_univ _), Finset.card_univ, rank] + rw [this, mul_assoc, ← abs_mul, ← Matrix.det_mul_column] + have : (2 : ℝ)⁻¹ ^ NrComplexPlaces K = |∏ w : InfinitePlace K, (mult w : ℝ)⁻¹| := by + rw [Finset.abs_prod] + simp_rw [mult, Nat.cast_ite, Nat.cast_one, Nat.cast_ofNat, apply_ite, abs_inv, abs_one, inv_one, + Nat.abs_ofNat, Finset.prod_ite, Finset.prod_const_one, Finset.prod_const, one_mul] + rw [Finset.filter_not, Finset.card_univ_diff, ← Fintype.card_subtype] + rw [card_eq_nrRealPlaces_add_nrComplexPlaces, ← NrRealPlaces, add_tsub_cancel_left] + rw [this, mul_assoc, ← abs_mul, ← Matrix.det_mul_row, abs_mul] + congr + · rw [abs_eq_self.mpr] + rw [inv_nonneg] + exact Finset.prod_nonneg fun _ _ ↦ (mapToUnitsPow₀_pos _ _).le + · ext + simp only [Matrix.transpose_apply, Matrix.of_apply, ite_mul, one_mul, mul_ite] + split_ifs + · rw [inv_mul_cancel₀ mult_coe_ne_zero] + · rw [← mul_assoc, mul_comm _ (c w₀), mul_assoc, inv_mul_cancel_left₀ mult_coe_ne_zero, + abs_eq_self.mpr hc] + +open MeasureTheory + +private theorem setLIntegral_mapToUnitsPow_aux₁ {s : Set (realSpace K)} (hs : s ⊆ {x | 0 ≤ x w₀}) : + s \ (s ∩ {x | 0 < x w₀}) ⊆ {x | x w₀ = 0} := by + refine fun _ h ↦ eq_of_ge_of_not_gt (hs h.1) ?_ + have := h.2 + simp_rw [Set.mem_inter_iff, not_and, h.1, true_implies] at this + exact this + +private theorem setLIntegral_mapToUnitsPow_aux₂ : + volume {x : realSpace K | x w₀ = 0} = 0 := by + let A : AffineSubspace ℝ (realSpace K) := + Submodule.toAffineSubspace (Submodule.mk ⟨⟨{x | x w₀ = 0}, by aesop⟩, rfl⟩ (by aesop)) + convert Measure.addHaar_affineSubspace volume A fun h ↦ ?_ + have : 1 ∈ A := h ▸ Set.mem_univ _ + simp [A] at this + +private theorem setLIntegral_mapToUnitsPow_aux₃ {s : Set (realSpace K)} (hs : s ⊆ {x | 0 ≤ x w₀}) : + (mapToUnitsPow K) '' s =ᵐ[volume] (mapToUnitsPow K) '' (s ∩ {x | 0 < x w₀}) := by + refine measure_symmDiff_eq_zero_iff.mp ?_ + rw [symmDiff_of_ge (Set.image_mono Set.inter_subset_left)] + have : mapToUnitsPow K '' s \ mapToUnitsPow K '' (s ∩ {x | 0 < x w₀}) ⊆ { 0 } := by + rintro _ ⟨⟨x, hx, rfl⟩, hx'⟩ + have hx'' : x ∉ s ∩ {x | 0 < x w₀} := fun h ↦ hx' (Set.mem_image_of_mem _ h) + simp_rw [Set.mem_inter_iff, Set.mem_setOf_eq, not_and] at hx'' + rw [Set.mem_singleton_iff, mapToUnitsPow_zero_iff] + refine eq_of_ge_of_not_gt (hs hx) (hx'' hx) + exact measure_mono_null this (measure_singleton _) + +open ENNReal Classical in +theorem setLIntegral_mapToUnitsPow {s : Set (realSpace K)} (hs₀ : MeasurableSet s) + (hs₁ : s ⊆ {x | 0 ≤ x w₀}) (f : (InfinitePlace K → ℝ) → ℝ≥0∞) : + ∫⁻ x in (mapToUnitsPow K) '' s, f x = + (2 : ℝ≥0∞)⁻¹ ^ NrComplexPlaces K * ENNReal.ofReal (regulator K) * (finrank ℚ K) * + ∫⁻ x in s, ENNReal.ofReal |x w₀| ^ rank K * + ENNReal.ofReal (∏ i : {w : InfinitePlace K // IsComplex w}, + (mapToUnitsPow₀ K (fun w ↦ x w) i))⁻¹ * f (mapToUnitsPow K x) := by + rw [setLIntegral_congr (setLIntegral_mapToUnitsPow_aux₃ hs₁)] + have : s =ᵐ[volume] (s ∩ {x | 0 < x w₀} : Set (realSpace K)) := by + refine measure_symmDiff_eq_zero_iff.mp <| + measure_mono_null ?_ setLIntegral_mapToUnitsPow_aux₂ + rw [symmDiff_of_ge Set.inter_subset_left] + exact setLIntegral_mapToUnitsPow_aux₁ hs₁ + rw [setLIntegral_congr this] + have h : MeasurableSet (s ∩ {x | 0 < x w₀}) := + hs₀.inter (measurableSet_lt measurable_const (measurable_pi_apply w₀)) + rw [lintegral_image_eq_lintegral_abs_det_fderiv_mul volume h (fun c hc ↦ + HasFDerivWithinAt.mono (hasFDeriv_mapToUnitsPow c (hs₁ (Set.mem_of_mem_inter_left hc))) + (Set.inter_subset_left.trans hs₁)) ((mapToUnitsPow K).injOn.mono Set.inter_subset_right)] + rw [setLIntegral_congr_fun h + (ae_of_all volume fun c hc ↦ by rw [mapToUnitsPow_jacobian_det + (hs₁ (Set.mem_of_mem_inter_left hc))]), ← lintegral_const_mul'] + · congr with x + -- This will be useful for positivity goals + have : 0 ≤ (∏ w : {w : InfinitePlace K // IsComplex w}, mapToUnitsPow₀ K (fun w ↦ x w) w)⁻¹ := + inv_nonneg.mpr <| Finset.prod_nonneg fun w _ ↦ (mapToUnitsPow₀_pos _ _).le + rw [ofReal_mul (by positivity), ofReal_mul (by positivity), ofReal_mul (by positivity), + ofReal_mul (by positivity), ofReal_natCast, ofReal_pow (by positivity), ofReal_pow + (by positivity), ofReal_inv_of_pos zero_lt_two, ofReal_ofNat] + ring + · exact mul_ne_top (mul_ne_top (pow_ne_top (inv_ne_top.mpr two_ne_zero)) ofReal_ne_top) + (natCast_ne_top _) + +end mapToUnitsPow + +noncomputable section mapToUnitsPowComplex + +variable [NumberField K] + +/-- DOCSTRING. -/ +def mapToUnitsPowComplex : PartialHomeomorph + ((realSpace K) × ({w : InfinitePlace K // IsComplex w} → ℝ)) (mixedSpace K) := + PartialHomeomorph.trans + (PartialHomeomorph.prod (mapToUnitsPow K) (PartialHomeomorph.refl _)) + (polarCoordMixedSpace K).symm + +theorem mapToUnitsPowComplex_source : + (mapToUnitsPowComplex K).source = {x | 0 < x w₀} ×ˢ Set.univ.pi fun _ ↦ Set.Ioo (-π) π := by + ext + simp_rw [mapToUnitsPowComplex, PartialHomeomorph.trans_source, PartialHomeomorph.prod_source, + PartialHomeomorph.refl_source, Set.mem_inter_iff, Set.mem_prod, Set.mem_univ, and_true, + Set.mem_preimage, PartialHomeomorph.prod_apply, PartialHomeomorph.refl_apply, id_eq, + PartialHomeomorph.symm_source, polarCoordMixedSpace_target, Set.mem_prod, mapToUnitsPow_source] + rw [and_congr_right] + intro h + rw [and_iff_right_iff_imp] + intro _ + simp_rw [Set.mem_univ_pi, Set.mem_ite_univ_left, not_isReal_iff_isComplex] + intro w _ + rw [Set.mem_Ioi, lt_iff_le_and_ne] + refine ⟨mapToUnitsPow_nonneg _ _, ?_⟩ + rw [ne_comm, ne_eq, mapToUnitsPow_zero_iff'] + exact ne_of_gt h + +theorem mapToUnitsPowComplex_target : + (mapToUnitsPowComplex K).target = + (Set.univ.pi fun _ ↦ Set.Ioi 0) ×ˢ (Set.univ.pi fun _ ↦ Complex.slitPlane) := by + ext + simp_rw [mapToUnitsPowComplex, PartialHomeomorph.trans_target, PartialHomeomorph.symm_target, + polarCoordMixedSpace_source, PartialHomeomorph.prod_target, PartialHomeomorph.refl_target, + Set.mem_inter_iff, Set.mem_preimage, mapToUnitsPow_target, Set.mem_prod, Set.mem_univ, + true_and, and_true, and_comm] + rw [and_congr_right] + intro h + simp_rw [PartialHomeomorph.symm_symm, polarCoordMixedSpace_apply, realProdComplexProdEquiv_apply, + Set.mem_pi, Set.mem_univ, true_implies] + refine ⟨?_, ?_⟩ + · intro h' w + specialize h' w + simp_rw [dif_pos w.prop] at h' + exact h' + · intro h' w + by_cases hw : IsReal w + · simp_rw [dif_pos hw] + exact h' ⟨w, hw⟩ + · simp_rw [dif_neg hw] + rw [Complex.polarCoord_apply] + dsimp only + rw [Set.mem_pi] at h + specialize h ⟨w, not_isReal_iff_isComplex.mp hw⟩ (Set.mem_univ _) + rw [AbsoluteValue.pos_iff] + exact Complex.slitPlane_ne_zero h + +theorem continuous_mapToUnitsPowComplex : + Continuous (mapToUnitsPowComplex K) := by + simp [mapToUnitsPowComplex] + refine Continuous.comp ?_ ?_ + · exact continuous_polarCoordMixedSpace_symm K + · rw [continuous_prod_mk] + refine ⟨?_, continuous_snd⟩ + · exact (continuous_mapToUnitsPow K).comp' continuous_fst + +variable {K} + +theorem mapToUnitsPowComplex_apply (x : (InfinitePlace K → ℝ) × ({w // IsComplex w} → ℝ)) : + mapToUnitsPowComplex K x = + (fun w ↦ mapToUnitsPow K x.1 w.val, + fun w ↦ Complex.polarCoord.symm (mapToUnitsPow K x.1 w.val, x.2 w)) := rfl + +theorem mixedToReal_mapToUnitsPowComplex + (x : (InfinitePlace K → ℝ) × ({w // IsComplex w} → ℝ)) : + mixedToReal (mapToUnitsPowComplex K x) = mapToUnitsPow K x.1 := by + ext w + simp_rw [mapToUnitsPowComplex_apply] + obtain hw | hw := isReal_or_isComplex w + · rw [mixedToReal_apply_of_isReal _ ⟨w, hw⟩] + · rw [mixedToReal_apply_of_isComplex _ ⟨w, hw⟩, Complex.norm_eq_abs, Complex.polarCoord_symm_abs, + abs_of_nonneg (mapToUnitsPow_nonneg x.1 w)] + +theorem mapToUnitsPowComplex_image_prod (s : Set (InfinitePlace K → ℝ)) + (t : Set ({w : InfinitePlace K // IsComplex w} → ℝ)) : + mapToUnitsPowComplex K '' (s ×ˢ t) = + (polarCoordMixedSpace K).symm '' (mapToUnitsPow K '' s) ×ˢ t := by + ext + simp_rw [mapToUnitsPowComplex, PartialHomeomorph.coe_trans, Function.comp_apply, + PartialHomeomorph.prod_apply, PartialHomeomorph.refl_apply, id_eq, + polarCoordMixedSpace_symm_apply, Set.mem_image, Set.mem_prod, Prod.exists] + refine ⟨?_, ?_⟩ + · rintro ⟨x, y, ⟨hx, hy⟩, rfl⟩ + exact ⟨mapToUnitsPow K x, y, ⟨Set.mem_image_of_mem _ hx, hy⟩, rfl⟩ + · rintro ⟨_, y, ⟨⟨⟨x, hx, rfl⟩, hy⟩, rfl⟩⟩ + exact ⟨x, y, ⟨hx, hy⟩, rfl⟩ + +theorem mapToUnitsPowComplex_prod_indicator_aux (x y : ℝ × ℝ) (hx : x ∈ Set.Ici 0 ×ˢ Set.Icc (-π) π) + (hy : y ∈ Complex.polarCoord.target) + (hxy : Complex.polarCoord.symm x = Complex.polarCoord.symm y) : + x = y := by + suffices x ∈ Complex.polarCoord.target from Complex.polarCoord.symm.injOn this hy hxy + suffices x.1 ≠ 0 ∧ x.2 ≠ -π ∧ x.2 ≠ π by + simp only [Complex.polarCoord_target, Set.mem_prod, Set.mem_Ioi, Set.mem_Ioo] + simp at hx + refine ⟨?_, ?_, ?_⟩ + · rw [lt_iff_le_and_ne] + exact ⟨hx.1, this.1.symm⟩ + · rw [lt_iff_le_and_ne] + exact ⟨hx.2.1, this.2.1.symm⟩ + · rw [lt_iff_le_and_ne] + exact ⟨hx.2.2, this.2.2⟩ + have := (Complex.polarCoord_symm_mem_polarCoord_source x).mp ?_ + have hx₁ : 0 < x.1 := by + refine lt_iff_le_and_ne.mpr ⟨?_, ?_⟩ + exact hx.1 + exact this.1.symm + · refine ⟨?_, ?_, ?_⟩ + · exact this.1 + · have := this.2.1 hx₁ (-1) + rwa [show π + (-1 : ℤ) * (2 * π) = -π by ring, ne_comm] at this + · have := this.2.1 hx₁ 0 + rwa [Int.cast_zero, zero_mul, add_zero, ne_comm] at this + · rw [hxy] + exact Complex.polarCoord.map_target hy + +theorem mapToUnitsPowComplex_prod_indicator + {s : Set (InfinitePlace K → ℝ)} {t : Set ({w : InfinitePlace K // IsComplex w} → ℝ)} + (ht : t ⊆ Set.univ.pi fun _ ↦ Set.Icc (-π) π) + (x : (InfinitePlace K → ℝ) × ({w // IsComplex w} → ℝ)) + (hx : x ∈ (polarCoordMixedSpace K).target) : + (mapToUnitsPowComplex K '' s ×ˢ t).indicator (fun _ ↦ (1 : ENNReal)) + ((polarCoordMixedSpace K).symm x) = + (mapToUnitsPow K '' s).indicator 1 x.1 * t.indicator 1 x.2 := by + classical + simp_rw [mapToUnitsPowComplex_image_prod, ← Set.indicator_prod_one, Prod.mk.eta, + Set.indicator_apply, Set.mem_image, polarCoordMixedSpace_symm_apply, Prod.mk.inj_iff] + refine if_congr ⟨fun ⟨y, hy, ⟨hxy₁, hxy₂⟩⟩ ↦ ?_, fun h ↦ ⟨x, h, rfl, rfl⟩⟩ rfl rfl + suffices y = x by rwa [← this] + have hxy : ∀ i (hi : IsComplex i), y.1 i = x.1 i ∧ y.2 ⟨i, hi⟩ = x.2 ⟨i, hi⟩ := by + intro i hi + rw [← Prod.mk.inj_iff] + refine mapToUnitsPowComplex_prod_indicator_aux _ _ ?_ ?_ (congr_fun hxy₂ ⟨i, hi⟩) + · rw [Set.prod_mk_mem_set_prod_eq] + refine ⟨?_, ?_⟩ + · obtain ⟨t, _, ht⟩ := hy.1 + rw [← ht] + exact mapToUnitsPow_nonneg _ _ + · exact ht hy.2 ⟨i, hi⟩ trivial + · simp_rw [polarCoordMixedSpace_target, Set.mem_prod, Set.mem_univ_pi, Set.mem_ite_univ_left, + not_isReal_iff_isComplex] at hx + exact ⟨hx.1 i hi, hx.2 ⟨i, hi⟩⟩ + ext i + · obtain hi | hi := isReal_or_isComplex i + · exact congr_fun hxy₁ ⟨i, hi⟩ + · exact (hxy i hi).1 + · exact (hxy i.val i.prop).2 + +-- Isn't this result used already to compute the integral? +theorem mapToUnitsPow_image_minus_image_inter_aux {s : Set (InfinitePlace K → ℝ)} + (hs : s ⊆ {x | 0 ≤ x w₀}) : + s \ (s ∩ {x | 0 < x w₀}) ⊆ {x | x w₀ = 0} := by + refine fun _ h ↦ eq_of_ge_of_not_gt (hs h.1) ?_ + have := h.2 + simp_rw [Set.mem_inter_iff, not_and, h.1, true_implies] at this + exact this + +theorem mapToUnitsPow_image_minus_image_inter + {s : Set (InfinitePlace K → ℝ)} (hs : s ⊆ {x | 0 ≤ x w₀}) : + mapToUnitsPow K '' s \ mapToUnitsPow K '' (s ∩ {x | 0 < x w₀}) ⊆ {0} := by + rintro _ ⟨⟨x, hx, rfl⟩, hx'⟩ + have hx'' : x ∉ s ∩ {x | 0 < x w₀} := fun h ↦ hx' (Set.mem_image_of_mem _ h) + rw [Set.mem_singleton_iff, mapToUnitsPow_zero_iff] + exact mapToUnitsPow_image_minus_image_inter_aux hs ⟨hx, hx''⟩ + +theorem measurable_mapToUnitsPow_image {s : Set (InfinitePlace K → ℝ)} + (hs : MeasurableSet s) (hs' : s ⊆ {x | 0 ≤ x w₀}) : + MeasurableSet (mapToUnitsPow K '' s) := by + have hm : MeasurableSet (mapToUnitsPow K '' (s ∩ {x | 0 < x w₀})) := by + rw [PartialHomeomorph.image_eq_target_inter_inv_preimage _ (fun _ h ↦ h.2)] + refine (mapToUnitsPow K).open_target.measurableSet.inter ?_ + have : MeasurableSet (s ∩ {x | 0 < x w₀}) := + hs.inter (measurableSet_lt measurable_const (measurable_pi_apply w₀)) + exact MeasurableSet.preimage this (measurable_mapToUnitsPow_symm K) + obtain h | h : mapToUnitsPow K '' s = mapToUnitsPow K '' (s ∩ {x | 0 < x w₀}) ∨ + mapToUnitsPow K '' s = mapToUnitsPow K '' (s ∩ {x | 0 < x w₀}) ∪ { 0 } := by + have h₀ : mapToUnitsPow K '' (s ∩ {x | 0 < x w₀}) ⊆ mapToUnitsPow K '' s := + Set.image_mono Set.inter_subset_left + have h₂ : ∀ ⦃s t u : Set (realSpace K)⦄, t ⊆ s → s \ t = u → s = t ∪ u := by aesop + obtain h₁ | h₁ := Set.subset_singleton_iff_eq.mp (mapToUnitsPow_image_minus_image_inter hs') + · left + rw [h₂ h₀ h₁, Set.union_empty] + · right + rw [h₂ h₀ h₁] + · rwa [h] + · rw [h] + exact MeasurableSet.union hm (measurableSet_singleton 0) + +open MeasureTheory MeasureTheory.Measure + +open Classical in +theorem volume_mapToUnitsPowComplex_set_prod_set {s : Set (InfinitePlace K → ℝ)} + (hs : MeasurableSet s) (hs' : s ⊆ {x | 0 ≤ x w₀} ) + {t : Set ({w : InfinitePlace K // IsComplex w} → ℝ)} + (ht : MeasurableSet t) (ht' : t ⊆ Set.univ.pi fun _ ↦ Set.Icc (-π) π) + (hm : MeasurableSet (mapToUnitsPowComplex K '' s ×ˢ t)) : + volume (mapToUnitsPowComplex K '' (s ×ˢ t)) = + volume ((Set.univ.pi fun _ ↦ Set.Ioo (-π) π) ∩ t) * ∫⁻ x in mapToUnitsPow K '' s, + ∏ w : {w : InfinitePlace K // IsComplex w}, (x w).toNNReal := by + rw [← setLIntegral_one, ← lintegral_indicator hm, + lintegral_eq_lintegral_polarCoordMixedSpace_symm K _ + ((measurable_indicator_const_iff 1).mpr hm), + setLIntegral_congr (setLIntegral_mapToUnitsPow_aux₃ hs')] + calc + _ = ∫⁻ x in Set.univ.pi fun w ↦ if IsReal w then Set.univ else Set.Ioi 0, + ∫⁻ y in Set.univ.pi fun _ ↦ Set.Ioo (-π) π, + (∏ w : {w // IsComplex w}, (x w.val).toNNReal) * + ((mapToUnitsPow K '' s).indicator 1 x * t.indicator 1 y) := by + rw [lintegral_lintegral, Measure.prod_restrict, ← polarCoordMixedSpace_target] + · refine setLIntegral_congr_fun (measurableSet_polarCoordMixedSpace_target K) ?_ + filter_upwards with x hx + simp_rw [mapToUnitsPowComplex_prod_indicator ht' x hx] + · refine Measurable.aemeasurable ?_ + refine Measurable.mul ?_ ?_ + · exact measurable_coe_nnreal_ennreal_iff.mpr <| + Finset.measurable_prod _ fun _ _ ↦ by fun_prop + · refine Measurable.mul ?_ ?_ + · refine Measurable.ite ?_ ?_ ?_ + · change MeasurableSet (Prod.fst ⁻¹' (mapToUnitsPow K '' s)) + refine measurable_fst ?_ + refine measurable_mapToUnitsPow_image hs hs' + · exact measurable_const + · exact measurable_const + · refine Measurable.comp' ?_ measurable_snd + exact measurable_const.indicator ht + _ = volume ((Set.univ.pi fun _ ↦ Set.Ioo (-π) π) ∩ t) * + ∫⁻ x in Set.univ.pi fun w ↦ if IsReal w then Set.univ else Set.Ioi 0, + (∏ w : {w // IsComplex w}, (x w.val).toNNReal) * + (mapToUnitsPow K '' s).indicator 1 x := by + conv_lhs => + enter [2, x] + rw [lintegral_const_mul' _ _ ENNReal.coe_ne_top] + rw [lintegral_const_mul' _ _ (by + rw [Set.indicator_apply] + split_ifs + exacts [ENNReal.one_ne_top, ENNReal.zero_ne_top])] + rw [← lintegral_indicator (MeasurableSet.univ_pi fun _ ↦ measurableSet_Ioo), + Set.indicator_indicator, lintegral_indicator_one ((MeasurableSet.univ_pi + fun _ ↦ measurableSet_Ioo).inter ht)] + rw [← lintegral_const_mul'] + · congr with x + ring + · refine ne_top_of_le_ne_top ?_ (measure_mono Set.inter_subset_left) + simp_rw [volume_pi, pi_pi, Real.volume_Ioo, Finset.prod_const] + refine ENNReal.pow_ne_top ENNReal.ofReal_ne_top + _ = volume ((Set.univ.pi fun _ ↦ Set.Ioo (-π) π) ∩ t) * + ∫⁻ x in Set.univ.pi fun w ↦ if IsReal w then Set.univ else Set.Ioi 0, + (∏ w : {w // IsComplex w}, (x w.val).toNNReal) * + (mapToUnitsPow K '' (s ∩ {x | 0 < x w₀})).indicator 1 x := by + congr 1 + refine lintegral_congr_ae ?_ + refine Filter.EventuallyEq.mul ?_ ?_ + · exact Filter.EventuallyEq.rfl + · refine indicator_ae_eq_of_ae_eq_set ?_ + refine Filter.EventuallyEq.restrict ?_ + exact setLIntegral_mapToUnitsPow_aux₃ hs' + _ = volume ((Set.univ.pi fun _ ↦ Set.Ioo (-π) π) ∩ t) * + ∫⁻ x in mapToUnitsPow K '' (s ∩ {x | 0 < x w₀}), + ∏ w : {w // IsComplex w}, (x w.val).toNNReal := by + rw [← lintegral_indicator, ← lintegral_indicator] + congr with x + rw [Set.indicator_mul_right, Set.indicator_indicator, Set.inter_eq_right.mpr] + by_cases hx : x ∈ (mapToUnitsPow K) '' (s ∩ {x | 0 < x w₀}) + · rw [Set.indicator_of_mem hx, Set.indicator_of_mem hx, Pi.one_apply, mul_one, + ENNReal.coe_finset_prod] + · rw [Set.indicator_of_not_mem hx, Set.indicator_of_not_mem hx, mul_zero] + · rintro _ ⟨x, hx, rfl⟩ + refine Set.mem_univ_pi.mpr fun _ ↦ ?_ + rw [Set.mem_ite_univ_left] + intro _ + rw [Set.mem_Ioi] + exact mapToUnitsPow_pos (ne_of_gt hx.2) _ + · refine measurable_mapToUnitsPow_image ?_ ?_ + · exact hs.inter (measurableSet_lt measurable_const (measurable_pi_apply w₀)) + · exact Set.inter_subset_left.trans hs' + · refine MeasurableSet.univ_pi fun _ ↦ ?_ + refine MeasurableSet.ite' (fun _ ↦ ?_) (fun _ ↦ ?_) + exact MeasurableSet.univ + exact measurableSet_Ioi + +end mapToUnitsPowComplex + +namespace fundamentalCone + +open Pointwise Module + +variable [NumberField K] + +/-- DOCSTRING -/ +abbrev normLessThanOne : Set (mixedSpace K) := + {x | x ∈ fundamentalCone K ∧ mixedEmbedding.norm x ≤ 1} + +/-- DOCSTRING -/ +abbrev normEqOne : Set (mixedSpace K) := {x | x ∈ fundamentalCone K ∧ mixedEmbedding.norm x = 1} + +variable {K} + +theorem mem_normLessThanOne_of_normAtPlace_eq {x y : mixedSpace K} (hx : x ∈ normLessThanOne K) + (hy : ∀ w, normAtPlace w y = normAtPlace w x) : + y ∈ normLessThanOne K := by + have h₁ : mixedEmbedding.norm y = mixedEmbedding.norm x := by + simp_rw [mixedEmbedding.norm_apply, hy] + exact ⟨mem_of_normAtPlace_eq hx.1 hy, h₁ ▸ hx.2⟩ + +theorem mem_normLessThanOne_iff (x : mixedSpace K) : + x ∈ normLessThanOne K ↔ (fun w ↦ |x.1 w|, x.2) ∈ normLessThanOne K := by + have h_eq : ∀ w, normAtPlace w x = normAtPlace w (fun w ↦ |x.1 w|, x.2) := by + intro w + obtain hw | hw := isReal_or_isComplex w + · simp_rw [normAtPlace_apply_isReal hw, Real.norm_eq_abs, abs_abs] + · simp_rw [normAtPlace_apply_isComplex hw] + exact ⟨fun h ↦ mem_normLessThanOne_of_normAtPlace_eq h (fun w ↦ (h_eq w).symm), + fun h ↦ mem_normLessThanOne_of_normAtPlace_eq h h_eq⟩ + +theorem smul_normEqOne {c : ℝ} (hc : 0 < c) : + c • normEqOne K = {x | x ∈ fundamentalCone K ∧ mixedEmbedding.norm x = c ^ finrank ℚ K} := by + ext + rw [← Set.preimage_smul_inv₀ (ne_of_gt hc), Set.preimage_setOf_eq, Set.mem_setOf_eq, + mixedEmbedding.norm_smul, abs_inv, abs_eq_self.mpr hc.le, inv_pow, mul_comm, mul_inv_eq_one₀ + (pow_ne_zero _ (ne_of_gt hc)), Set.mem_setOf_eq, and_congr_left_iff] + exact fun _ ↦ smul_mem_iff_mem (inv_ne_zero (ne_of_gt hc)) + +theorem exists_mem_smul_normEqOne {x : mixedSpace K} (hx : x ∈ normLessThanOne K) : + ∃ c : ℝ, 0 < c ∧ c ≤ 1 ∧ x ∈ c • normEqOne K := by + have h₁ : (finrank ℚ K : ℝ) ≠ 0 := Nat.cast_ne_zero.mpr (ne_of_gt finrank_pos) + have h₂ : 0 < mixedEmbedding.norm x ^ (finrank ℚ K : ℝ)⁻¹ := + Real.rpow_pos_of_pos (norm_pos_of_mem hx.1) _ + refine ⟨(mixedEmbedding.norm x) ^ (finrank ℚ K : ℝ)⁻¹, h₂, ?_, ?_⟩ + · exact Real.rpow_le_one (mixedEmbedding.norm_nonneg _) hx.2 (inv_nonneg.mpr (Nat.cast_nonneg _)) + · rw [smul_normEqOne h₂] + refine ⟨hx.1, ?_⟩ + rw [← Real.rpow_natCast, ← Real.rpow_mul (mixedEmbedding.norm_nonneg _), inv_mul_cancel₀ h₁, + Real.rpow_one] + +theorem smul_normEqOne_subset {c : ℝ} (hc₁ : 0 < c) (hc₂ : c ≤ 1) : + c • normEqOne K ⊆ normLessThanOne K := by + rw [smul_normEqOne hc₁] + refine fun x hx ↦ ⟨hx.1, ?_⟩ + rw [hx.2] + exact pow_le_one₀ hc₁.le hc₂ + +theorem normLessThanOne_eq_union_smul_normEqOne : + normLessThanOne K = ⋃ c ∈ Set.Ioc (0 : ℝ) 1, c • normEqOne K := by + ext + simp_rw [Set.mem_iUnion, Set.mem_Ioc, exists_prop, and_assoc] + exact ⟨fun hx ↦ exists_mem_smul_normEqOne hx, + fun ⟨_, h₁, h₂, hx⟩ ↦ smul_normEqOne_subset h₁ h₂ hx⟩ + +variable (K) in +open Classical in +/-- DOCSTRING. -/ +abbrev box₁ : Set (realSpace K) := + Set.univ.pi fun w ↦ if w = w₀ then Set.Ioc 0 1 else Set.Ico 0 1 + +variable (K) in +theorem measurableSet_box₁ : + MeasurableSet (box₁ K) := + MeasurableSet.univ_pi fun _ ↦ + MeasurableSet.ite' (fun _ ↦ measurableSet_Ioc) (fun _ ↦ measurableSet_Ico) + +theorem pos_of_mem_box₁ {x : InfinitePlace K → ℝ} (hx : x ∈ box₁ K) : + 0 < x w₀ := by + have := hx w₀ (Set.mem_univ w₀) + simp_rw [if_true] at this + exact this.1 + +theorem mem_Ioc_of_mem_box₁ {c : realSpace K} (hc : c ∈ box₁ K) : + c w₀ ∈ Set.Ioc 0 1 := by + have := hc w₀ (Set.mem_univ _) + simp_rw [ite_true] at this + exact this + +theorem mem_Ico_of_mem_box₁ {c : realSpace K} (hc : c ∈ box₁ K) {w : InfinitePlace K} + (hw : w ≠ w₀) : + c w ∈ Set.Ico 0 1 := by + have := hc w (Set.mem_univ _) + simp_rw [if_neg hw] at this + exact this + +theorem mixedToReal_plusPart_normEqOne : + mixedToReal '' plusPart (normEqOne K) = + mapToUnitsPow₀ K '' (Set.univ.pi fun _ ↦ Set.Ico 0 1) := by + classical + ext + simp_rw [Set.mem_image, normEqOne, Set.mem_inter_iff, fundamentalCone, + Set.mem_diff, Set.mem_preimage, ZSpan.mem_fundamentalDomain] + constructor + · rintro ⟨x, ⟨⟨⟨h₁⟩, h₂⟩, h₃⟩, rfl⟩ + refine ⟨(mapToUnitsPow₀ K).symm (mixedToReal x), ?_, ?_⟩ + · intro i _ + rw [mapToUnitsPow₀_symm_apply_of_norm_one (by rwa [norm_mixedToRealToMixed]), + logMap_mixedToRealToMixed_of_norm_one h₂] + exact h₁ (equivFinRank.symm i) + · rw [PartialEquiv.right_inv] + exact mixedToReal_mem_target h₃ h₂ + · rintro ⟨c, hc, rfl⟩ + refine ⟨realToMixed (mapToUnitsPow₀ K c), ⟨⟨⟨fun i ↦ ?_, ?_⟩, + norm_mapToUnitsPow₀ c⟩, fun w ↦ mapToUnitsPow₀_pos c _⟩, + realToMixedToReal_eq_self_of_nonneg fun w _ ↦ ( mapToUnitsPow₀_pos c w).le⟩ + · rw [show i = equivFinRank.symm (equivFinRank i) by rw [Equiv.symm_apply_apply], + ← Basis.repr_reindex_apply, ← Basis.equivFun_apply, ← mapToUnitsPow₀_symm_apply_of_norm_one + (norm_mapToUnitsPow₀ c), PartialEquiv.left_inv _ (by trivial)] + exact hc _ trivial + · rw [Set.mem_setOf_eq, norm_mapToUnitsPow₀] + exact one_ne_zero + +omit [NumberField K] in +private theorem mixedToReal_normLessThanOne_aux₁ {c : ℝ} (hc : 0 < c) : + {x : mixedSpace K | ∀ w, 0 < x.1 w} = c • {x | ∀ w, 0 < x.1 w} := by + ext x + refine ⟨fun hx ↦ + ⟨c⁻¹ • x, fun w ↦ mul_pos (inv_pos.mpr hc) (hx w), by simp_rw [smul_inv_smul₀ hc.ne']⟩, ?_⟩ + rintro ⟨y, hy, rfl⟩ + exact fun w ↦ mul_pos hc (hy w) + +private theorem mixedToReal_normLessThanOne_aux₂ {c : ℝ} (hc : 0 < c) : + mixedToReal '' (c • normEqOne K ∩ {x | ∀ w, 0 < x.1 w}) = + c • mixedToReal '' plusPart (normEqOne K) := by + nth_rewrite 1 [mixedToReal_normLessThanOne_aux₁ hc, ← Set.smul_set_inter₀ hc.ne'] + ext + constructor + · rintro ⟨x, ⟨hx, rfl⟩⟩ + refine ⟨c⁻¹ • mixedToReal x, ⟨c⁻¹ • x, ?_, ?_⟩, ?_⟩ + · rwa [← Set.mem_smul_set_iff_inv_smul_mem₀ hc.ne'] + · rw [mixedToReal_smul _ (inv_nonneg.mpr hc.le)] + · simp_rw [smul_inv_smul₀ hc.ne'] + · rintro ⟨_, ⟨x, hx, rfl⟩, rfl⟩ + exact ⟨c • x, Set.smul_mem_smul_set hx, by rw [mixedToReal_smul _ hc.le]⟩ + +variable (K) + +theorem mixedToReal_plusPart_normLessThanOne : + mixedToReal '' plusPart (normLessThanOne K) = mapToUnitsPow K '' (box₁ K) := by + classical + rw [normLessThanOne_eq_union_smul_normEqOne, plusPart, Set.iUnion₂_inter, Set.image_iUnion₂, + Set.iUnion₂_congr (fun _ h ↦ by rw [mixedToReal_normLessThanOne_aux₂ h.1]), + mixedToReal_plusPart_normEqOne] + ext + simp_rw [Set.mem_iUnion, Set.mem_image, exists_prop', nonempty_prop, Set.mem_smul_set] + constructor + · rintro ⟨n, hn, ⟨x, ⟨c, hc, rfl⟩, rfl⟩⟩ + refine ⟨fun w ↦ if hw : w = w₀ then n else c ⟨w, hw⟩, ?_, ?_⟩ + · refine Set.mem_univ_pi.mpr ?_ + intro w + by_cases hw : w = w₀ + · rw [hw, dif_pos rfl, if_pos rfl] + exact hn + · rw [dif_neg hw, if_neg hw] + exact hc ⟨w, hw⟩ trivial + · rw [mapToUnitsPow_apply', dif_pos rfl, abs_of_pos hn.1] + conv_lhs => + enter [2, 2, w] + rw [dif_neg w.prop] + · rintro ⟨c, hc, rfl⟩ + refine ⟨c w₀, mem_Ioc_of_mem_box₁ hc, mapToUnitsPow₀ K fun w ↦ c w, ⟨fun w ↦ c w, + fun w _ ↦ mem_Ico_of_mem_box₁ hc w.prop, rfl⟩, ?_⟩ + rw [mapToUnitsPow_apply', abs_of_pos (mem_Ioc_of_mem_box₁ hc).1] + +/-- DOCSTRING. -/ +abbrev box₂ : Set ({w : InfinitePlace K // IsComplex w} → ℝ) := + Set.univ.pi fun _ ↦ Set.Ioc (-π) π + +theorem measurableSet_box₂ : + MeasurableSet (box₂ K) := MeasurableSet.univ_pi fun _ ↦ measurableSet_Ioc + +/-- DOCSTRING. -/ +abbrev box := (box₁ K) ×ˢ (box₂ K) + +theorem plusPart_normLessThanOne : + plusPart (normLessThanOne K) = mapToUnitsPowComplex K '' (box K) := by + have : plusPart (normLessThanOne K) = + mixedToReal⁻¹' (mixedToReal '' plusPart (normLessThanOne K)) := by + refine subset_antisymm (Set.subset_preimage_image mixedToReal _) + fun x ⟨y, ⟨hy₁, hy₂⟩, h⟩ ↦ ⟨mem_normLessThanOne_of_normAtPlace_eq hy₁ fun w ↦ ?_, fun w ↦ ?_⟩ + · rw [← norm_mixedToReal, ← congr_fun h w, norm_mixedToReal] + · rw [← mixedToReal_apply_of_isReal _ w, ← congr_fun h w, mixedToReal_apply_of_isReal _ w] + exact hy₂ w + rw [this, mixedToReal_plusPart_normLessThanOne] + ext x + refine ⟨fun ⟨c, hc₁, hc₂⟩ ↦ ⟨⟨c, fun w ↦ (x.2 w).arg⟩, ⟨hc₁, fun w _ ↦ (x.2 w).arg_mem_Ioc⟩, ?_⟩, + ?_⟩ + · simp_rw [mapToUnitsPowComplex_apply, Complex.polarCoord_symm_apply, hc₂] + ext w + · simp_rw [mixedToReal_apply_of_isReal] + · simp_rw [mixedToReal_apply_of_isComplex, Complex.norm_eq_abs, Complex.ofReal_cos, + Complex.ofReal_sin, Complex.abs_mul_cos_add_sin_mul_I] + · rintro ⟨c, hc, rfl⟩ + exact ⟨c.1, hc.1, by rw [mixedToReal_mapToUnitsPowComplex]⟩ + +theorem measurableSet_normLessThanOne : + MeasurableSet (normLessThanOne K) := + (measurableSet_fundamentalCone K).inter <| + measurableSet_le (mixedEmbedding.continuous_norm K).measurable measurable_const + +theorem isBounded_normEqOne : + IsBounded (normEqOne K) := by + classical + let B := (basisUnitLattice K).ofZLatticeBasis ℝ + obtain ⟨r, hr₁, hr₂⟩ := (ZSpan.fundamentalDomain_isBounded B).subset_closedBall_lt 0 0 + have h₁ : ∀ ⦃x w⦄, x ∈ normEqOne K → w ≠ w₀ → |mult w * Real.log (normAtPlace w x)| ≤ r := by + intro x w hx hw + rw [← logMap_apply_of_norm_one hx.2 ⟨w, hw⟩] + suffices ‖logMap x‖ ≤ r by exact (pi_norm_le_iff_of_nonneg hr₁.le).mp this ⟨w, hw⟩ + exact mem_closedBall_zero_iff.mp (hr₂ hx.1.1) + have h₂ : ∀ ⦃x⦄, x ∈ normEqOne K → mult (w₀ : InfinitePlace K) * Real.log (normAtPlace w₀ x) ≤ + (Finset.univ.erase (w₀ : InfinitePlace K)).card • r := by + intro x hx + suffices mult (w₀ : InfinitePlace K) * Real.log (normAtPlace w₀ x) = + - ∑ w ∈ Finset.univ.erase w₀, mult w * Real.log (normAtPlace w x) by + rw [this, ← Finset.sum_neg_distrib, ← Finset.sum_const] + exact Finset.sum_le_sum fun w hw ↦ + neg_le_of_neg_le (abs_le.mp (h₁ hx (Finset.mem_erase.mp hw).1)).1 + simp_rw [← Real.log_pow] + rw [← add_eq_zero_iff_eq_neg, Finset.univ.add_sum_erase (fun w ↦ + ((normAtPlace w x) ^ mult w).log) (Finset.mem_univ w₀), ← Real.log_prod, + ← mixedEmbedding.norm_apply, hx.2, Real.log_one] + exact fun w _ ↦ pow_ne_zero _ <| ne_of_gt (normAtPlace_pos_of_mem hx.1 w) + have h₃ : ∀ ⦃x w c⦄, 0 ≤ c → x ∈ fundamentalCone K → + mult w * Real.log (normAtPlace w x) ≤ c → normAtPlace w x ≤ Real.exp c := by + intro x w c hc hx + rw [← le_div_iff₀' (Nat.cast_pos.mpr mult_pos), + Real.log_le_iff_le_exp (normAtPlace_pos_of_mem hx w)] + exact fun h ↦ le_trans h <| Real.exp_le_exp.mpr (div_le_self hc one_le_mult) + refine (Metric.isBounded_iff_subset_closedBall 0).mpr + ⟨max (Real.exp r) (Real.exp ((Finset.univ.erase (w₀ : InfinitePlace K)).card • r)), + fun x hx ↦ mem_closedBall_zero_iff.mpr ?_⟩ + rw [norm_eq_sup'_normAtPlace] + refine Finset.sup'_le _ _ fun w _ ↦ ?_ + by_cases hw : w = w₀ + · rw [hw] + exact (h₃ (nsmul_nonneg (hr₁.le) _) hx.1 (h₂ hx)).trans (le_max_right _ _) + · exact le_trans (h₃ hr₁.le hx.1 (abs_le.mp (h₁ hx hw)).2) (le_max_left _ _) + +theorem isBounded_normLessThanOne : + IsBounded (normLessThanOne K) := by + classical + obtain ⟨r, hr₁, hr₂⟩ := (isBounded_normEqOne K).subset_ball_lt 0 0 + refine (Metric.isBounded_iff_subset_ball 0).mpr ⟨r, fun x hx ↦ ?_⟩ + obtain ⟨c, hc₁, _, hc₂⟩ := exists_mem_smul_normEqOne hx + suffices x ∈ c • Metric.ball (0 : (mixedSpace K)) r by + rw [smul_ball (ne_of_gt hc₁), smul_zero] at this + refine Set.mem_of_subset_of_mem (Metric.ball_subset_ball ?_) this + rwa [mul_le_iff_le_one_left hr₁, Real.norm_eq_abs, abs_eq_self.mpr hc₁.le] + exact (Set.image_subset _ hr₂) hc₂ + +open MeasureTheory MeasureTheory.Measure + +theorem volume_normLessThanOnePlus_aux (n : ℕ) : + ∫⁻ x in box₁ K, ENNReal.ofReal |x w₀| ^ n = (n + 1 : ENNReal)⁻¹ := by + classical + rw [volume_pi, box₁, restrict_pi_pi, lintegral_eq_lmarginal_univ 0, + lmarginal_erase' _ ?_ (Finset.mem_univ w₀)] + simp_rw [if_true, Function.update_same] + have : ∫⁻ (xᵢ : ℝ) in Set.Ioc 0 1, ENNReal.ofReal |xᵢ| ^ n = (n + 1 : ENNReal)⁻¹ := by + convert congr_arg ENNReal.ofReal (integral_pow (a := 0) (b := 1) n) + · rw [intervalIntegral.integral_of_le zero_le_one] + rw [ofReal_integral_eq_lintegral_ofReal] + · refine setLIntegral_congr_fun measurableSet_Ioc ?_ + filter_upwards with _ h using by rw [abs_of_pos h.1, ENNReal.ofReal_pow h.1.le] + · refine IntegrableOn.integrable ?_ + rw [← Set.uIoc_of_le zero_le_one, ← intervalIntegrable_iff] + exact intervalIntegral.intervalIntegrable_pow n + · exact ae_restrict_of_forall_mem measurableSet_Ioc fun _ h ↦ pow_nonneg h.1.le _ + · rw [one_pow, zero_pow (by linarith), sub_zero, ENNReal.ofReal_div_of_pos (by positivity), + ENNReal.ofReal_add (by positivity) zero_le_one, ENNReal.ofReal_one, ENNReal.ofReal_natCast, + one_div] + rw [this] + rw [lmarginal] + rw [lintegral_const] + rw [pi_univ] + rw [Finset.prod_congr rfl (g := fun _ ↦ 1) (fun x _ ↦ by rw [if_neg (by aesop), restrict_apply + MeasurableSet.univ, Set.univ_inter, Real.volume_Ico, sub_zero, ENNReal.ofReal_one])] + rw [Finset.prod_const_one, mul_one] + fun_prop + +open Classical in +theorem volume_plusPart_normLessThanOne : volume (plusPart (normLessThanOne K)) = + NNReal.pi ^ NrComplexPlaces K * (regulator K).toNNReal := by + rw [plusPart_normLessThanOne, volume_mapToUnitsPowComplex_set_prod_set + (measurableSet_box₁ K) (fun _ h ↦ le_of_lt (pos_of_mem_box₁ h)) (measurableSet_box₂ K) + (Set.pi_mono fun _ _ ↦ Set.Ioc_subset_Icc_self) + (by rw [← plusPart_normLessThanOne]; exact + measurableSet_plusPart (measurableSet_normLessThanOne K)), + setLIntegral_mapToUnitsPow (measurableSet_box₁ K) (fun _ h ↦ ((if_pos rfl) ▸ + Set.mem_univ_pi.mp h w₀).1.le), Set.inter_eq_left.mpr (Set.pi_mono fun _ _ ↦ + Set.Ioo_subset_Ioc_self), volume_pi_pi] + simp_rw [Real.volume_Ioo, sub_neg_eq_add, ← two_mul, Finset.prod_const, ENNReal.ofReal_mul + zero_le_two, ENNReal.ofReal_ofNat, mul_pow] + have h₁ : ∀ x : InfinitePlace K → ℝ, + 0 < ∏ i : {w // IsComplex w}, (mapToUnitsPow₀ K) (fun w ↦ x w) i.val := + fun _ ↦ Finset.prod_pos fun _ _ ↦ mapToUnitsPow₀_pos _ _ + have h₂ : rank K + NrComplexPlaces K + 1 = finrank ℚ K := by + rw [rank, add_comm _ 1, ← add_assoc, add_tsub_cancel_of_le NeZero.one_le, + card_eq_nrRealPlaces_add_nrComplexPlaces, ← card_add_two_mul_card_eq_rank] + ring + calc + _ = (NNReal.pi : ENNReal) ^ NrComplexPlaces K * (regulator K).toNNReal * (finrank ℚ K) * + ∫⁻ x in box₁ K, ENNReal.ofReal |x w₀| ^ (rank K + NrComplexPlaces K) := by + simp_rw [← mul_assoc] + congr + · rw [mul_comm, ← mul_assoc, NrComplexPlaces, Finset.card_univ, ← mul_pow, + ENNReal.inv_mul_cancel two_ne_zero ENNReal.two_ne_top, one_pow, one_mul, + ← ENNReal.ofReal_coe_nnreal, NNReal.coe_real_pi] + · ext x + simp_rw [mapToUnitsPow_apply', Pi.smul_apply, smul_eq_mul, Real.toNNReal_mul (abs_nonneg _), + ENNReal.coe_mul, ENNReal.ofNNReal_toNNReal] + rw [Finset.prod_mul_distrib, Finset.prod_const, mul_mul_mul_comm, + ← ENNReal.ofReal_prod_of_nonneg (fun _ _ ↦ (mapToUnitsPow₀_pos _ _).le), + ENNReal.ofReal_inv_of_pos (h₁ x), ENNReal.inv_mul_cancel + (ENNReal.ofReal_ne_zero_iff.mpr (h₁ x)) ENNReal.ofReal_ne_top, mul_one, pow_add, + NrComplexPlaces, Finset.card_univ] + _ = NNReal.pi ^ NrComplexPlaces K * (regulator K).toNNReal := by + rw [volume_normLessThanOnePlus_aux, ← Nat.cast_add_one, h₂, mul_assoc, ENNReal.mul_inv_cancel, + mul_one] + · rw [Nat.cast_ne_zero] + refine ne_of_gt ?_ + exact finrank_pos + · exact ENNReal.natCast_ne_top _ + +open Classical in +theorem volume_normLessThanOne : + volume (normLessThanOne K) = + 2 ^ NrRealPlaces K * NNReal.pi ^ NrComplexPlaces K * (regulator K).toNNReal := by + rw [volume_eq_two_pow_mul_volume_plusPart (normLessThanOne K) mem_normLessThanOne_iff + (measurableSet_normLessThanOne K), volume_plusPart_normLessThanOne, mul_assoc] + +theorem isBounded_box₁ : Bornology.IsBounded (box₁ K) := by + refine Bornology.IsBounded.pi ?_ + intro i + by_cases hi : i = w₀ + · rw [hi, if_pos rfl] + exact Metric.isBounded_Ioc 0 1 + · rw [if_neg hi] + exact Metric.isBounded_Ico 0 1 + +omit [NumberField K] in +theorem isBounded_box₂ : Bornology.IsBounded (box₂ K) := by + refine Bornology.IsBounded.pi ?_ + intro _ + exact Metric.isBounded_Ioc _ _ + +theorem closure_box₁ : + closure (box₁ K) = Set.Icc 0 1 := by + rw [closure_pi_set] + simp_rw [← Set.pi_univ_Icc, Pi.zero_apply, Pi.one_apply] + refine Set.pi_congr rfl ?_ + intro i _ + by_cases hi : i = w₀ + · rw [hi, if_pos rfl] + exact closure_Ioc zero_ne_one + · rw [if_neg hi] + exact closure_Ico zero_ne_one + +omit [NumberField K] in +theorem closure_box₂ : + closure (box₂ K) = Set.univ.pi fun _ ↦ Set.Icc (-π) π := by + rw [closure_pi_set] + refine Set.pi_congr rfl ?_ + intro _ _ + refine closure_Ioc ?_ + rw [ne_eq, CharZero.neg_eq_self_iff] + exact Real.pi_ne_zero + +theorem interior_box₁ : + interior (box₁ K) = Set.univ.pi fun _ ↦ Set.Ioo 0 1 := by + rw [interior_pi_set Set.finite_univ] + refine Set.pi_congr rfl ?_ + intro i _ + by_cases hi : i = w₀ + · rw [hi, if_pos rfl] + exact interior_Ioc + · rw [if_neg hi] + exact interior_Ico + +theorem interior_box₂ : + interior (box₂ K) = Set.univ.pi fun _ ↦ Set.Ioo (-π) π := by + rw [interior_pi_set Set.finite_univ] + refine Set.pi_congr rfl ?_ + intro _ _ + exact interior_Ioc + +theorem isClosed_mapToUnitsPowComplex_closure : + IsClosed (mapToUnitsPowComplex K '' (closure (box K))) := by + classical + refine (IsCompact.image_of_continuousOn ?_ ?_).isClosed + · exact Metric.isCompact_iff_isClosed_bounded.mpr + ⟨isClosed_closure, + Metric.isBounded_closure_iff.mpr ((isBounded_box₁ K).prod (isBounded_box₂ K))⟩ + · exact (continuous_mapToUnitsPowComplex K).continuousOn + +theorem closure_subset_mapToUnitsPowComplex_closure : + closure (plusPart (normLessThanOne K)) ⊆ mapToUnitsPowComplex K '' (closure (box K)) := by + rw [plusPart_normLessThanOne] + exact closure_minimal (Set.image_mono subset_closure) (isClosed_mapToUnitsPowComplex_closure K) + +theorem interior_subset_mapToUnitsPowComplex_source : + interior (box K) ⊆ (mapToUnitsPowComplex K).source := by + rw [interior_prod_eq, interior_box₁, interior_box₂, mapToUnitsPowComplex_source] + exact Set.prod_mono (fun _ h ↦ (h w₀ (Set.mem_univ _)).1) subset_rfl + +theorem isOpen_mapToUnitsPowComplex_interior : + IsOpen (mapToUnitsPowComplex K '' (interior (box K))) := + (mapToUnitsPowComplex K).isOpen_image_of_subset_source isOpen_interior + (interior_subset_mapToUnitsPowComplex_source K) + +theorem mapToUnitsPowComplex_interior_subset_interior : + mapToUnitsPowComplex K '' (interior (box K)) ⊆ interior (plusPart (normLessThanOne K)) := by + rw [plusPart_normLessThanOne] + exact interior_maximal (Set.image_mono interior_subset) (isOpen_mapToUnitsPowComplex_interior K) + +open Classical in +theorem volume_mapToUnitsPowComplex_interior_eq_volume_mapToUnitsPowComplex_closure : + volume (mapToUnitsPowComplex K '' (interior (box K))) = + volume (mapToUnitsPowComplex K '' (closure (box K))) := by + have hm₁ : MeasurableSet + (mapToUnitsPowComplex K '' (Set.univ.pi fun x ↦ Set.Ioo 0 1) ×ˢ + Set.univ.pi fun _ ↦ Set.Ioo (-π) π) := by + rw [← interior_box₁, ← interior_box₂, ← interior_prod_eq] + exact (isOpen_mapToUnitsPowComplex_interior K).measurableSet + have hm₂ : MeasurableSet + (mapToUnitsPowComplex K '' Set.Icc 0 1 ×ˢ Set.univ.pi fun _ ↦ Set.Icc (-π) π) := by + rw [← closure_box₁, ← closure_box₂, ← closure_prod_eq] + exact (isClosed_mapToUnitsPowComplex_closure K).measurableSet + have h₁ : Set.Icc 0 1 ⊆ {x : InfinitePlace K → ℝ | 0 ≤ x w₀} := + fun _ h ↦ (Set.mem_Icc.mp h).1 w₀ + have h₂ : Set.univ.pi (fun _ : InfinitePlace K ↦ Set.Ioo (0 : ℝ) 1) ⊆ {x | 0 ≤ x w₀} := + fun _ h ↦ (h w₀ (Set.mem_univ _)).1.le + have h₃ : MeasurableSet (Set.univ.pi fun _ : InfinitePlace K ↦ Set.Ioo (0 : ℝ) 1) := + MeasurableSet.univ_pi fun _ ↦ measurableSet_Ioo + rw [closure_prod_eq, interior_prod_eq, closure_box₁, closure_box₂, interior_box₁, interior_box₂, + volume_mapToUnitsPowComplex_set_prod_set h₃ h₂ (MeasurableSet.univ_pi fun _ ↦ + measurableSet_Ioo) (Set.pi_mono fun _ _ ↦ Set.Ioo_subset_Icc_self) hm₁, + volume_mapToUnitsPowComplex_set_prod_set measurableSet_Icc h₁ (MeasurableSet.univ_pi fun _ ↦ + measurableSet_Icc) le_rfl hm₂] + simp_rw [setLIntegral_mapToUnitsPow h₃ h₂, setLIntegral_mapToUnitsPow measurableSet_Icc h₁, + mul_assoc, ← Set.pi_inter_distrib, Set.inter_self, Set.inter_eq_left.mpr + Set.Ioo_subset_Icc_self] + congr 5 + refine Measure.restrict_congr_set ?_ + rw [show (Set.univ.pi fun _ ↦ Set.Ioo (0 : ℝ) 1) = interior (Set.Icc 0 1) by + simp_rw [← Set.pi_univ_Icc, interior_pi_set Set.finite_univ, Pi.zero_apply, Pi.one_apply, + interior_Icc]] + exact interior_ae_eq_of_null_frontier ((convex_Icc 0 1).addHaar_frontier volume) + +open Classical in +theorem volume_frontier_plusPart_normLessThanOne : + volume (frontier (plusPart (normLessThanOne K))) = 0 := by + rw [frontier, measure_diff] + have : volume (closure (plusPart (normLessThanOne K))) = + volume (interior (plusPart (normLessThanOne K))) := by + refine le_antisymm ?_ (measure_mono interior_subset_closure) + refine le_trans ?_ (measure_mono (mapToUnitsPowComplex_interior_subset_interior K)) + rw [volume_mapToUnitsPowComplex_interior_eq_volume_mapToUnitsPowComplex_closure] + exact measure_mono (closure_subset_mapToUnitsPowComplex_closure K) + refine tsub_eq_zero_iff_le.mpr (le_of_eq this) + · exact interior_subset_closure + · exact measurableSet_interior.nullMeasurableSet + · rw [← lt_top_iff_ne_top] + refine lt_of_le_of_lt (measure_mono interior_subset) ?_ + rw [volume_plusPart_normLessThanOne] + exact Batteries.compareOfLessAndEq_eq_lt.mp rfl + +open Classical in +theorem volume_frontier_normLessThanOne : + volume (frontier (normLessThanOne K)) = 0 := by + have h : normLessThanOne K ∩ ⋃ w, {x | x.1 w = 0} = ∅ := by + refine (Set.disjoint_left.mpr fun x hx₁ hx₂ ↦ ?_).inter_eq + obtain ⟨w, hw⟩ := Set.mem_iUnion.mp hx₂ + have := hx₁.1.2 + refine this ?_ + simp_rw [mixedEmbedding.norm_eq_zero_iff] + use w + rw [normAtPlace_apply_isReal, hw, norm_zero] + rw [← iUnion_negAtPlusPart_union (normLessThanOne K) mem_normLessThanOne_iff, h, Set.union_empty] + refine measure_mono_null (frontier_iUnion _) (measure_iUnion_null_iff.mpr fun s ↦ ?_) + rw [negAtPlusPart, ← negAt_preimage] + rw [← ContinuousLinearEquiv.coe_toHomeomorph, ← Homeomorph.preimage_frontier, + ContinuousLinearEquiv.coe_toHomeomorph, (volume_preserving_negAt s).measure_preimage + measurableSet_frontier.nullMeasurableSet] + exact volume_frontier_plusPart_normLessThanOne K + +end NumberField.mixedEmbedding.fundamentalCone + +set_option linter.style.longFile 1700 diff --git a/Mathlib/NumberTheory/NumberField/ClassNumber.lean b/Mathlib/NumberTheory/NumberField/ClassNumber.lean index 06ed53e2b5316..0cdae4676f32d 100644 --- a/Mathlib/NumberTheory/NumberField/ClassNumber.lean +++ b/Mathlib/NumberTheory/NumberField/ClassNumber.lean @@ -35,6 +35,10 @@ end RingOfIntegers noncomputable def classNumber : ℕ := Fintype.card (ClassGroup (𝓞 K)) +theorem classNumber_ne_zero : classNumber K ≠ 0 := Fintype.card_ne_zero + +theorem classNumber_pos : 0 < classNumber K := Fintype.card_pos + variable {K} /-- The class number of a number field is `1` iff the ring of integers is a PID. -/ diff --git a/Mathlib/NumberTheory/NumberField/DedekindZeta.lean b/Mathlib/NumberTheory/NumberField/DedekindZeta.lean new file mode 100644 index 0000000000000..b5859a158d9e9 --- /dev/null +++ b/Mathlib/NumberTheory/NumberField/DedekindZeta.lean @@ -0,0 +1,70 @@ +/- +Copyright (c) 2024 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.NumberTheory.LSeries.Residue +import Mathlib.NumberTheory.NumberField.Ideal + +/-! +# Docstring + +-/ + +variable (K : Type*) [Field K] [NumberField K] + +noncomputable section + +namespace NumberField + +open Filter Ideal NumberField.InfinitePlace NumberField.Units Topology NumberTheory.LSeries + +open scoped Real + +/-- Docstring. -/ +def dedekindZeta (s : ℂ) := + LSeries (fun n ↦ Nat.card {I : Ideal (𝓞 K) // absNorm I = n}) s + +/-- Docstring. -/ +def residue : ℝ := + (2 ^ NrRealPlaces K * (2 * π) ^ NrComplexPlaces K * regulator K * classNumber K) / + (torsionOrder K * Real.sqrt |discr K|) + +theorem residue_pos : 0 < residue K := by + refine div_pos ?_ ?_ + · exact mul_pos (mul_pos (by positivity) (regulator_pos K)) (Nat.cast_pos.mpr (classNumber_pos K)) + · refine mul_pos ?_ ?_ + · rw [Nat.cast_pos] + exact PNat.pos (torsionOrder K) + · exact Real.sqrt_pos_of_pos <| abs_pos.mpr (Int.cast_ne_zero.mpr (discr_ne_zero K)) + +theorem residue_ne_zero : residue K ≠ 0 := (residue_pos K).ne' + +theorem dedekindZeta_residue : + Tendsto (fun s : ℝ ↦ (s - 1) * dedekindZeta K s) (𝓝[>] 1) (𝓝 (residue K)) := by + refine tendsto_mul_of_sum_div_tendsto (residue_pos K) ?_ + convert (ideal.tendsto_norm_le_div_atop K).comp tendsto_natCast_atTop_atTop with n + simp_rw [Function.comp_apply, Nat.cast_le] + congr + have : ∀ i, Fintype {I : Ideal (𝓞 K) | absNorm I = i} := by + intro i + refine Set.Finite.fintype ?_ + exact finite_setOf_absNorm_eq i + have : ∀ i, Fintype {I : Ideal (𝓞 K) | absNorm I ≤ i} := by + intro i + refine Set.Finite.fintype ?_ + exact finite_setOf_absNorm_le i + simp_rw (config := {singlePass := true}) [← Set.coe_setOf, Nat.card_eq_card_toFinset] + rw [← Nat.cast_sum, Finset.card_eq_sum_card_fiberwise (t := Finset.range (n + 1)) + (f := fun I ↦ absNorm I)] + · congr! with n hn + ext + simp only [Set.mem_toFinset, Set.mem_setOf_eq, Finset.mem_filter, iff_and_self] + intro h + rw [h] + exact Finset.mem_range_succ_iff.mp hn + · intro x hx + simp at hx + exact Finset.mem_range_succ_iff.mpr hx + +end NumberField diff --git a/Mathlib/NumberTheory/NumberField/Discriminant.lean b/Mathlib/NumberTheory/NumberField/Discriminant.lean index 758d8a99b48ed..b4001671630a6 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant.lean @@ -3,6 +3,7 @@ Copyright (c) 2023 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot -/ +import Mathlib.Algebra.Module.ZLattice.Covolume import Mathlib.Data.Real.Pi.Bounds import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.ConvexBody import Mathlib.Tactic.Rify @@ -103,6 +104,24 @@ theorem _root_.NumberField.mixedEmbedding.volume_fundamentalDomain_latticeBasis stdBasis_repr_eq_matrixToStdBasis_mul K _ (fun _ => rfl)] rfl +theorem _root_.NumberField.mixedEmbedding.covolume_integerLattice : + ZLattice.covolume (mixedEmbedding.integerLattice K) = + (2 ⁻¹) ^ NrComplexPlaces K * √|discr K| := by + rw [ZLattice.covolume_eq_measure_fundamentalDomain _ _ (fundamentalDomain_integerLattice K), + volume_fundamentalDomain_latticeBasis, ENNReal.toReal_mul, ENNReal.toReal_pow, + ENNReal.toReal_inv, toReal_ofNat, ENNReal.coe_toReal, Real.coe_sqrt, coe_nnnorm, + Int.norm_eq_abs] + +theorem _root_.NumberField.mixedEmbedding.covolume_idealLattice (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) : + ZLattice.covolume (mixedEmbedding.idealLattice K I) = + (FractionalIdeal.absNorm (I : FractionalIdeal (𝓞 K)⁰ K)) * + (2 ⁻¹) ^ NrComplexPlaces K * √|discr K| := by + rw [ZLattice.covolume_eq_measure_fundamentalDomain _ _ (fundamentalDomain_idealLattice K I), + volume_fundamentalDomain_fractionalIdealLatticeBasis, volume_fundamentalDomain_latticeBasis, + ENNReal.toReal_mul, ENNReal.toReal_mul, ENNReal.toReal_pow, ENNReal.toReal_inv, toReal_ofNat, + ENNReal.coe_toReal, Real.coe_sqrt, coe_nnnorm, Int.norm_eq_abs, + ENNReal.toReal_ofReal (Rat.cast_nonneg.mpr (FractionalIdeal.absNorm_nonneg I.val)), mul_assoc] + theorem exists_ne_zero_mem_ideal_of_norm_le_mul_sqrt_discr (I : (FractionalIdeal (𝓞 K)⁰ K)ˣ) : ∃ a ∈ (I : FractionalIdeal (𝓞 K)⁰ K), a ≠ 0 ∧ |Algebra.norm ℚ (a : K)| ≤ FractionalIdeal.absNorm I.1 * (4 / π) ^ NrComplexPlaces K * diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index d713d6bce1442..9ab5362e58492 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -426,7 +426,10 @@ theorem mult_pos {w : InfinitePlace K} : 0 < mult w := by split_ifs <;> norm_num @[simp] -theorem mult_ne_zero {w : InfinitePlace K} : mult w ≠ 0 := ne_of_gt mult_pos +theorem mult_ne_zero {w : InfinitePlace K} : mult w ≠ 0 := mult_pos.ne' + +theorem mult_coe_ne_zero {w : InfinitePlace K} : (mult w : ℝ) ≠ 0 := + Nat.cast_ne_zero.mpr mult_ne_zero theorem one_le_mult {w : InfinitePlace K} : (1 : ℝ) ≤ mult w := by rw [← Nat.cast_one, Nat.cast_le] diff --git a/Mathlib/NumberTheory/NumberField/Ideal.lean b/Mathlib/NumberTheory/NumberField/Ideal.lean new file mode 100644 index 0000000000000..cf030425ddc1e --- /dev/null +++ b/Mathlib/NumberTheory/NumberField/Ideal.lean @@ -0,0 +1,139 @@ +/- +Copyright (c) 2024 Xavier Roblot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Xavier Roblot +-/ +import Mathlib.NumberTheory.NumberField.CanonicalEmbedding.NormLessThanOne +import Mathlib.NumberTheory.NumberField.ClassNumber + +/-! +# Docstring + +-/ + +variable (K : Type*) [Field K] [NumberField K] + +namespace NumberField + +open Filter Ideal Submodule Topology NumberField NumberField.InfinitePlace + NumberField.mixedEmbedding NumberField.mixedEmbedding.fundamentalCone + NumberField.mixedEmbedding.euclidean NumberField.Units + +open scoped nonZeroDivisors Real + +open Classical in +private noncomputable def ideal.tendsto_mk_eq_norm_le_div_atop_aux₁ (J : (Ideal (𝓞 K))⁰) (s : ℝ) : + ↑({x | x ∈ (toMixed K) ⁻¹' fundamentalCone K ∧ mixedEmbedding.norm ((toMixed K) x) ≤ s} ∩ + (ZLattice.comap ℝ (idealLattice K ((FractionalIdeal.mk0 K) J)) (toMixed K).toLinearMap)) + ≃ {a : idealPoint K J // mixedEmbedding.norm (a : mixedSpace K) ≤ s} := by + rw [ZLattice.coe_comap] + refine (((toMixed K).toEquiv.image _).trans (Equiv.setCongr ?_)).trans + (Equiv.subtypeSubtypeEquivSubtypeInter _ (mixedEmbedding.norm · ≤ s)).symm + ext + simp_rw [mem_idealPoint, Set.mem_image, Set.mem_inter_iff, Set.mem_preimage, SetLike.mem_coe, + mem_idealLattice, FractionalIdeal.coe_mk0] + constructor + · rintro ⟨_, ⟨⟨hx₁, hx₂⟩, _, ⟨x, hx₃, rfl⟩, rfl⟩, rfl⟩ + exact ⟨⟨hx₁, x, hx₃, rfl⟩, hx₂⟩ + · rintro ⟨⟨hx₁, ⟨x, hx₂, rfl⟩⟩, hx₃⟩ + exact ⟨(toMixed K).symm (mixedEmbedding K x), ⟨⟨hx₁, hx₃⟩, ⟨(x : K), by simp [hx₂], rfl⟩⟩, rfl⟩ + +private theorem ideal.tendsto_mk_eq_norm_le_div_atop_aux₂ (C : ClassGroup (𝓞 K)) + (J : (Ideal (𝓞 K))⁰) (hJ : ClassGroup.mk0 J = C⁻¹) (s : ℝ) : + Nat.card {I : (Ideal (𝓞 K))⁰ // absNorm (I : Ideal (𝓞 K)) ≤ s ∧ ClassGroup.mk0 I = C} + = Nat.card {I : (Ideal (𝓞 K))⁰ // (J : Ideal (𝓞 K)) ∣ I ∧ IsPrincipal (I : Ideal (𝓞 K)) ∧ + absNorm (I : Ideal (𝓞 K)) ≤ s * absNorm (J : Ideal (𝓞 K))} := by + let e : (Ideal (𝓞 K))⁰ ≃ {I : (Ideal (𝓞 K))⁰ // J ∣ I} := by + refine Equiv.ofBijective (fun I ↦ ⟨J * I, dvd_mul_right J I⟩) ⟨?_, ?_⟩ + · intro _ _ h + simp_rw [Subtype.ext_iff, Submonoid.coe_mul, mul_eq_mul_left_iff] at h + exact Subtype.ext_iff_val.mpr (h.resolve_right (nonZeroDivisors.coe_ne_zero J)) + · rintro ⟨_, ⟨I, hI⟩⟩ + exact ⟨I, Subtype.mk_eq_mk.mpr hI.symm⟩ + refine Nat.card_congr ?_ + let g := Equiv.subtypeSubtypeEquivSubtypeInter (fun I : (Ideal (𝓞 K))⁰ ↦ J ∣ I) + simp_rw [← nonZeroDivisors_dvd_iff_dvd_coe] + refine Equiv.trans ?_ (g _) + refine Equiv.subtypeEquiv e ?_ + intro I + simp_rw [← ClassGroup.mk0_eq_one_iff (SetLike.coe_mem (e I).1), e, Equiv.ofBijective_apply, + Submonoid.coe_mul, ← Submonoid.mul_def, _root_.map_mul, Nat.cast_mul, hJ, and_comm, + ← inv_eq_iff_mul_eq_one, inv_inv, eq_comm, mul_comm s _, _root_.mul_le_mul_left + (Nat.cast_pos.mpr (Nat.zero_lt_of_ne_zero (absNorm_ne_zero_of_nonZeroDivisors J)))] + +theorem ideal.tendsto_mk_eq_norm_le_div_atop (C : ClassGroup (𝓞 K)) : + Tendsto (fun s : ℝ ↦ + (Nat.card {I : (Ideal (𝓞 K))⁰ // + absNorm (I : Ideal (𝓞 K)) ≤ s ∧ ClassGroup.mk0 I = C} : ℝ) / s) atTop + (𝓝 ((2 ^ NrRealPlaces K * (2 * π) ^ NrComplexPlaces K * regulator K) / + (torsionOrder K * Real.sqrt |discr K|))) := by + classical + have h : ∀ s : ℝ, + {x | x ∈ toMixed K ⁻¹' fundamentalCone K ∧ mixedEmbedding.norm (toMixed K x) ≤ s} = + toMixed K ⁻¹' {x | x ∈ fundamentalCone K ∧ mixedEmbedding.norm x ≤ s} := fun _ ↦ rfl + obtain ⟨J, hJ⟩ := ClassGroup.mk0_surjective C⁻¹ + convert ((ZLattice.covolume.tendsto_card_le_div' + (ZLattice.comap ℝ (mixedEmbedding.idealLattice K (FractionalIdeal.mk0 K J)) + (toMixed K).toLinearMap) + (F := fun x ↦ mixedEmbedding.norm (toMixed K x)) + (X := (toMixed K)⁻¹' (fundamentalCone K)) (fun _ _ _ h ↦ ?_) (fun _ _ h ↦ ?_) + (isBounded_normLessThanOne K) ?_ ?_).mul (tendsto_const_nhds + (x := (absNorm (J : Ideal (𝓞 K)) : ℝ) * (torsionOrder K : ℝ)⁻¹))).comp + (tendsto_id.atTop_mul_const' <| Nat.cast_pos.mpr (absNorm_pos_of_nonZeroDivisors J)) + using 2 with s + · simp_rw [ideal.tendsto_mk_eq_norm_le_div_atop_aux₂ K C J hJ, Function.comp_def, id_eq, + Nat.card_congr (ideal.tendsto_mk_eq_norm_le_div_atop_aux₁ K J (s * (absNorm J.1))), + ← card_isPrincipal_dvd_norm_le, Nat.cast_mul, div_eq_mul_inv, mul_inv, ← mul_assoc] + rw [inv_mul_cancel_right₀ (Nat.cast_ne_zero.mpr (absNorm_ne_zero_of_nonZeroDivisors J)), + mul_right_comm, mul_inv_cancel_right₀ (Nat.cast_ne_zero.mpr (torsionOrder K).ne_zero)] + · simp_rw [h, (volumePreserving_toMixed K).measure_preimage + (measurableSet_normLessThanOne K).nullMeasurableSet, volume_normLessThanOne, + ZLattice.covolume_comap _ _ _ (volumePreserving_toMixed K), covolume_idealLattice, + ENNReal.toReal_mul, ENNReal.toReal_pow, ENNReal.toReal_ofNat, ENNReal.coe_toReal, + NNReal.coe_real_pi, Real.coe_toNNReal _ (regulator_pos K).le, FractionalIdeal.coe_mk0, + FractionalIdeal.coeIdeal_absNorm, Rat.cast_natCast, inv_pow, div_eq_mul_inv, mul_inv, + inv_inv, mul_assoc, mul_comm (absNorm J.1 : ℝ), mul_comm (absNorm J.1 : ℝ)⁻¹, mul_assoc] + rw [mul_inv_cancel₀ (Nat.cast_ne_zero.mpr (absNorm_ne_zero_of_nonZeroDivisors J)), mul_one] + ring + · rwa [Set.mem_preimage, map_smul, smul_mem_iff_mem h.ne'] + · simp_rw [map_smul, mixedEmbedding.norm_smul, euclidean.finrank, abs_of_nonneg h] + · exact (toMixed K).continuous.measurable (measurableSet_normLessThanOne K) + · rw [h, ← ContinuousLinearEquiv.coe_toHomeomorph, ← Homeomorph.preimage_frontier, + ContinuousLinearEquiv.coe_toHomeomorph, (volumePreserving_toMixed K).measure_preimage + measurableSet_frontier.nullMeasurableSet, volume_frontier_normLessThanOne] + +theorem ideal.tendsto_norm_le_div_atop₀ : + Tendsto (fun s : ℝ ↦ + (Nat.card {I : (Ideal (𝓞 K))⁰ // absNorm (I : Ideal (𝓞 K)) ≤ s} : ℝ) / s) atTop + (𝓝 ((2 ^ NrRealPlaces K * (2 * π) ^ NrComplexPlaces K * regulator K * classNumber K) / + (torsionOrder K * Real.sqrt |discr K|))) := by + classical + convert Filter.Tendsto.congr' ?_ + (tendsto_finset_sum Finset.univ (fun C _ ↦ ideal.tendsto_mk_eq_norm_le_div_atop K C)) + · rw [Finset.sum_const, Finset.card_univ, nsmul_eq_mul, classNumber] + ring + · filter_upwards [eventually_ge_atTop 0] with s hs + have : Fintype {I : (Ideal (𝓞 K))⁰ // absNorm (I : Ideal (𝓞 K)) ≤ s} := by + simp_rw [← Nat.le_floor_iff hs] + refine @Fintype.ofFinite _ ?_ + refine (finite_setOf_absNorm_le (S := 𝓞 K) ⌊s⌋₊).of_injective (fun ⟨I, h⟩ ↦ ⟨I.1, h⟩) ?_ + intro _ _ h + rwa [Subtype.mk.injEq, SetLike.coe_eq_coe, SetCoe.ext_iff] at h + let e := fun C : ClassGroup (𝓞 K) ↦ Equiv.subtypeSubtypeEquivSubtypeInter + (fun I : (Ideal (𝓞 K))⁰ ↦ absNorm I.1 ≤ s) (fun I ↦ ClassGroup.mk0 I = C) + simp_rw [← Nat.card_congr (e _), Nat.card_eq_fintype_card, Fintype.subtype_card] + rw [Fintype.card, Finset.card_eq_sum_card_fiberwise (f := fun I ↦ ClassGroup.mk0 I.1) + (t := Finset.univ) (fun _ _ ↦ Finset.mem_univ _), Nat.cast_sum, Finset.sum_div] + +theorem ideal.tendsto_norm_le_div_atop : + Tendsto (fun s : ℝ ↦ (Nat.card {I : Ideal (𝓞 K) // absNorm I ≤ s} : ℝ) / s) atTop + (𝓝 ((2 ^ NrRealPlaces K * (2 * π) ^ NrComplexPlaces K * regulator K * classNumber K) / + (torsionOrder K * Real.sqrt |discr K|))) := by + have := (ideal.tendsto_norm_le_div_atop₀ K).add tendsto_inv_atTop_zero + rw [add_zero] at this + refine Tendsto.congr' ?_ this + filter_upwards [eventually_ge_atTop 0] with s hs + simp_rw [← Nat.le_floor_iff hs] + rw [Ideal.card_norm_le_eq_card_norm_le_add_one, Nat.cast_add, Nat.cast_one, add_div, one_div] + +end NumberField diff --git a/Mathlib/NumberTheory/NumberField/Units/Basic.lean b/Mathlib/NumberTheory/NumberField/Units/Basic.lean index 7907293a0c950..ea607c7cd1f74 100644 --- a/Mathlib/NumberTheory/NumberField/Units/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Units/Basic.lean @@ -90,13 +90,21 @@ end coe open NumberField.InfinitePlace +variable {K} + @[simp] protected theorem norm [NumberField K] (x : (𝓞 K)ˣ) : |Algebra.norm ℚ (x : K)| = 1 := by rw [← RingOfIntegers.coe_norm, isUnit_iff_norm.mp x.isUnit] +@[simp] +theorem pos_at_place (x : (𝓞 K)ˣ) (w : InfinitePlace K) : + 0 < w x := pos_iff.mpr (coe_ne_zero x) + section torsion +variable (K) + /-- The torsion subgroup of the group of units. -/ def torsion : Subgroup (𝓞 K)ˣ := CommGroup.torsion (𝓞 K)ˣ diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 4970ba0001a5a..d0a64ae08ad20 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -76,9 +76,9 @@ variable (K) /-- The logarithmic embedding of the units (seen as an `Additive` group). -/ def _root_.NumberField.Units.logEmbedding : Additive ((𝓞 K)ˣ) →+ ({w : InfinitePlace K // w ≠ w₀} → ℝ) := -{ toFun := fun x w => mult w.val * Real.log (w.val ↑(Additive.toMul x)) +{ toFun := fun x w ↦ mult w.val * Real.log (w.val ↑(Additive.toMul x)) map_zero' := by simp; rfl - map_add' := fun _ _ => by simp [Real.log_mul, mul_add]; rfl } + map_add' := fun _ _ ↦ by simp [Real.log_mul, mul_add]; rfl } variable {K} @@ -95,10 +95,10 @@ theorem sum_logEmbedding_component (x : (𝓞 K)ˣ) : rw [← insert_erase (mem_univ w₀), sum_insert (not_mem_erase w₀ univ), add_comm, add_eq_zero_iff_eq_neg] at h convert h using 1 - · refine (sum_subtype _ (fun w => ?_) (fun w => (mult w) * (Real.log (w (x : K))))).symm - exact ⟨ne_of_mem_erase, fun h => mem_erase_of_ne_of_mem h (mem_univ w)⟩ + · refine (sum_subtype _ (fun w ↦ ?_) (fun w ↦ (mult w) * (Real.log (w (x : K))))).symm + exact ⟨ne_of_mem_erase, fun h ↦ mem_erase_of_ne_of_mem h (mem_univ w)⟩ · norm_num - · exact fun w _ => pow_ne_zero _ (AbsoluteValue.ne_zero _ (coe_ne_zero x)) + · refine fun w _ ↦ pow_ne_zero _ (Units.pos_at_place _ _).ne' end NumberField @@ -106,22 +106,21 @@ theorem mult_log_place_eq_zero {x : (𝓞 K)ˣ} {w : InfinitePlace K} : mult w * Real.log (w x) = 0 ↔ w x = 1 := by rw [mul_eq_zero, or_iff_right, Real.log_eq_zero, or_iff_right, or_iff_left] · linarith [(apply_nonneg _ _ : 0 ≤ w x)] - · simp only [ne_eq, map_eq_zero, coe_ne_zero x, not_false_eq_true] - · refine (ne_of_gt ?_) - rw [mult]; split_ifs <;> norm_num + · exact (Units.pos_at_place _ _).ne' + · exact mult_coe_ne_zero variable [NumberField K] theorem logEmbedding_eq_zero_iff {x : (𝓞 K)ˣ} : logEmbedding K (Additive.ofMul x) = 0 ↔ x ∈ torsion K := by rw [mem_torsion] - refine ⟨fun h w => ?_, fun h => ?_⟩ + refine ⟨fun h w ↦ ?_, fun h ↦ ?_⟩ · by_cases hw : w = w₀ · suffices -mult w₀ * Real.log (w₀ (x : K)) = 0 by rw [neg_mul, neg_eq_zero, ← hw] at this exact mult_log_place_eq_zero.mp this rw [← sum_logEmbedding_component, sum_eq_zero] - exact fun w _ => congrFun h w + exact fun w _ ↦ congrFun h w · exact mult_log_place_eq_zero.mp (congrFun h ⟨w, hw⟩) · ext w rw [logEmbedding_component, h w.val, Real.log_one, mul_zero, Pi.zero_apply] @@ -147,7 +146,7 @@ theorem log_le_of_logEmbedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) · rw [← hw] exact tool _ (abs_nonneg _) · refine (sum_le_card_nsmul univ _ _ - (fun w _ => logEmbedding_component_le hr h w)).trans ?_ + (fun w _ ↦ logEmbedding_component_le hr h w)).trans ?_ rw [nsmul_eq_mul] refine mul_le_mul ?_ le_rfl hr (Fintype.card (InfinitePlace K)).cast_nonneg simp [card_univ] @@ -250,9 +249,9 @@ def seq : ℕ → { x : 𝓞 K // x ≠ 0 } theorem seq_ne_zero (n : ℕ) : algebraMap (𝓞 K) K (seq K w₁ hB n) ≠ 0 := RingOfIntegers.coe_ne_zero_iff.mpr (seq K w₁ hB n).prop -/-- The terms of the sequence have nonzero norm. -/ -theorem seq_norm_ne_zero (n : ℕ) : Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K) ≠ 0 := - Algebra.norm_ne_zero_iff.mpr (Subtype.coe_ne_coe.1 (seq_ne_zero K w₁ hB n)) +-- /-- The terms of the sequence have nonzero norm. -/ +-- theorem seq_norm_ne_zero (n : ℕ) : Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K) ≠ 0 := +-- Algebra.norm_ne_zero_iff.mpr (Subtype.coe_ne_coe.1 (seq_ne_zero K w₁ hB n)) /-- The sequence is strictly decreasing at infinite places distinct from `w₁`. -/ theorem seq_decreasing {n m : ℕ} (h : n < m) (w : InfinitePlace K) (hw : w ≠ w₁) : @@ -295,26 +294,20 @@ theorem exists_unit (w₁ : InfinitePlace K) : rsuffices ⟨n, m, hnm, h⟩ : ∃ n m, n < m ∧ (Ideal.span ({ (seq K w₁ hB n : 𝓞 K) }) = Ideal.span ({ (seq K w₁ hB m : 𝓞 K) })) · have hu := Ideal.span_singleton_eq_span_singleton.mp h - refine ⟨hu.choose, fun w hw => Real.log_neg ?_ ?_⟩ - · exact pos_iff.mpr (coe_ne_zero _) - · calc - _ = w (algebraMap (𝓞 K) K (seq K w₁ hB m) * (algebraMap (𝓞 K) K (seq K w₁ hB n))⁻¹) := by - rw [← congr_arg (algebraMap (𝓞 K) K) hu.choose_spec, mul_comm, map_mul (algebraMap _ _), + refine ⟨hu.choose, fun w hw => Real.log_neg (pos_at_place (Exists.choose hu) w) ?_⟩ + calc + _ = w (algebraMap (𝓞 K) K (seq K w₁ hB m) * (algebraMap (𝓞 K) K (seq K w₁ hB n))⁻¹) := by + rw [← congr_arg (algebraMap (𝓞 K) K) hu.choose_spec, mul_comm, map_mul (algebraMap _ _), ← mul_assoc, inv_mul_cancel₀ (seq_ne_zero K w₁ hB n), one_mul] - _ = w (algebraMap (𝓞 K) K (seq K w₁ hB m)) * w (algebraMap (𝓞 K) K (seq K w₁ hB n))⁻¹ := - _root_.map_mul _ _ _ - _ < 1 := by - rw [map_inv₀, mul_inv_lt_iff₀ (pos_iff.mpr (seq_ne_zero K w₁ hB n)), one_mul] - exact seq_decreasing K w₁ hB hnm w hw - refine Set.Finite.exists_lt_map_eq_of_forall_mem - (t := { I : Ideal (𝓞 K) | 1 ≤ Ideal.absNorm I ∧ Ideal.absNorm I ≤ B }) - (fun n => ?_) ?_ - · rw [Set.mem_setOf_eq, Ideal.absNorm_span_singleton] - refine ⟨?_, seq_norm_le K w₁ hB n⟩ - exact Nat.one_le_iff_ne_zero.mpr (Int.natAbs_ne_zero.mpr (seq_norm_ne_zero K w₁ hB n)) - · rw [show { I : Ideal (𝓞 K) | 1 ≤ Ideal.absNorm I ∧ Ideal.absNorm I ≤ B } = - (⋃ n ∈ Set.Icc 1 B, { I : Ideal (𝓞 K) | Ideal.absNorm I = n }) by ext; simp] - exact Set.Finite.biUnion (Set.finite_Icc _ _) (fun n hn => Ideal.finite_setOf_absNorm_eq hn.1) + _ = w (algebraMap (𝓞 K) K (seq K w₁ hB m)) * w (algebraMap (𝓞 K) K (seq K w₁ hB n))⁻¹ := + _root_.map_mul _ _ _ + _ < 1 := by + rw [map_inv₀, mul_inv_lt_iff₀' (pos_iff.mpr (seq_ne_zero K w₁ hB n)), mul_one] + exact seq_decreasing K w₁ hB hnm w hw + refine Set.Finite.exists_lt_map_eq_of_forall_mem (t := {I : Ideal (𝓞 K) | Ideal.absNorm I ≤ B}) + (fun n ↦ ?_) (Ideal.finite_setOf_absNorm_le B) + rw [Set.mem_setOf_eq, Ideal.absNorm_span_singleton] + exact seq_norm_le K w₁ hB n theorem unitLattice_span_eq_top : Submodule.span ℝ (unitLattice K : Set ({w : InfinitePlace K // w ≠ w₀} → ℝ)) = ⊤ := by @@ -366,7 +359,7 @@ instance instDiscrete_unitLattice : DiscreteTopology (unitLattice K) := by · refine Set.Finite.of_finite_image ?_ (Set.injOn_of_injective Subtype.val_injective) convert unitLattice_inter_ball_finite K 1 ext x - refine ⟨?_, fun ⟨hx1, hx2⟩ => ⟨⟨x, hx1⟩, hx2, rfl⟩⟩ + refine ⟨?_, fun ⟨hx1, hx2⟩ ↦ ⟨⟨x, hx1⟩, hx2, rfl⟩⟩ rintro ⟨x, hx, rfl⟩ exact ⟨Subtype.mem x, hx⟩ @@ -478,7 +471,7 @@ def basisUnitLattice : Basis (Fin (rank K)) ℤ (unitLattice K) := units in `basisModTorsion`. -/ def fundSystem : Fin (rank K) → (𝓞 K)ˣ := -- `:)` prevents the `⧸` decaying to a quotient by `leftRel` when we unfold this later - fun i => Quotient.out' (Additive.toMul (basisModTorsion K i):) + fun i ↦ Quotient.out' (Additive.toMul (basisModTorsion K i):) theorem fundSystem_mk (i : Fin (rank K)) : Additive.ofMul (QuotientGroup.mk (fundSystem K i)) = (basisModTorsion K i) := by diff --git a/Mathlib/NumberTheory/NumberField/Units/Regulator.lean b/Mathlib/NumberTheory/NumberField/Units/Regulator.lean index 34d7ce2339491..a4cbacbf173bf 100644 --- a/Mathlib/NumberTheory/NumberField/Units/Regulator.lean +++ b/Mathlib/NumberTheory/NumberField/Units/Regulator.lean @@ -109,6 +109,56 @@ theorem regulator_eq_det (w' : InfinitePlace K) (e : {w // w ≠ w'} ≃ Fin (ra simp_rw [regulator_eq_det' K e', logEmbedding, AddMonoidHom.coe_mk, ZeroHom.coe_mk] exact abs_det_eq_abs_det K (fun i ↦ fundSystem K i) e' e +open Module in +theorem finrank_mul_regulator_eq_det (w' : InfinitePlace K) (e : {w // w ≠ w'} ≃ Fin (rank K)) : + finrank ℚ K * regulator K = + |(Matrix.of (fun i w : InfinitePlace K ↦ + if h : i = w' then (w.mult : ℝ) else w.mult * (w (fundSystem K (e ⟨i, h⟩))).log)).det| := by + rw [show |Matrix.det _| = |(1 : ℝ) • Matrix.det _| by rw [one_smul], + ← Matrix.det_updateColumn_sum _ w' (fun _ ↦ 1)] + let M := Matrix.of fun i w : InfinitePlace K ↦ if w = w' then + (if i = w' then (finrank ℚ K : ℝ) else 0) else + (if h : i = w' then w.mult else w.mult * (w (fundSystem K (e ⟨i, h⟩))).log) + have : |M.det| = finrank ℚ K * regulator K := by + simp only [M] + let e' : Fin (rank K + 1) ≃ InfinitePlace K := + (finSuccEquiv _).trans ((Equiv.optionSubtype _).symm e.symm).val + have h₁ : ∀ j, e' ((e'.symm w').succAbove j) = e.symm j := by + intro _ + have : e'.symm w' = 0 := by + rw [Equiv.symm_apply_eq, Equiv.trans_apply, finSuccEquiv_zero, + Equiv.optionSubtype_symm_apply_apply_none] + rw [this] + simp [ne_eq, Fin.zero_succAbove, Equiv.trans_apply, finSuccEquiv_succ, + Equiv.optionSubtype_symm_apply_apply_coe, e'] + have h₂ : ∀ j, e' ((e'.symm w').succAbove j) ≠ w' := by + intro _ + rw [ne_eq, Equiv.apply_eq_iff_eq_symm_apply] + exact Fin.succAbove_ne (e'.symm w') _ + rw [← Matrix.det_reindex_self e'.symm, Matrix.det_succ_column _ (e'.symm w')] + simp [Function.comp_def] + simp_rw [Equiv.apply_eq_iff_eq_symm_apply] + rw [Fintype.sum_ite_eq', abs_mul, abs_mul, Nat.abs_cast, abs_pow, abs_neg, abs_one, one_pow, + one_mul, regulator_eq_det K w' e, ← Matrix.det_reindex_self e] + rw [Matrix.reindex_apply] + congr + ext + simp_rw [Matrix.submatrix_apply, Matrix.of_apply] + simp_rw [Equiv.apply_symm_apply] + simp_rw [if_neg (h₂ _), dif_neg (h₂ _), h₁] + simp only [Subtype.coe_eta, Equiv.apply_symm_apply] + rw [← this] + congr + ext + have : ∀ (w : InfinitePlace K) i, w ((algebraMap (𝓞 K) K) (fundSystem K (e i))) ^ w.mult ≠ 0 := by + intro _ _ + refine pow_ne_zero _ ((map_ne_zero _).mpr (coe_ne_zero _)) + simp_rw [M, Matrix.of_apply, smul_eq_mul, one_mul, Finset.sum_dite_irrel, + Matrix.updateColumn_apply, ← Real.log_pow, ← Real.log_prod _ _ (fun _ _ ↦ this _ _), + prod_eq_abs_norm, + Units.norm, Rat.cast_one, Real.log_one, ← Nat.cast_sum, sum_mult_eq, dite_eq_ite, + Matrix.of_apply] + end Units end NumberField diff --git a/Mathlib/RingTheory/Ideal/Norm.lean b/Mathlib/RingTheory/Ideal/Norm.lean index 6e1a0aa1de64f..73a8177bcc104 100644 --- a/Mathlib/RingTheory/Ideal/Norm.lean +++ b/Mathlib/RingTheory/Ideal/Norm.lean @@ -391,24 +391,66 @@ theorem absNorm_eq_zero_iff {I : Ideal S} : Ideal.absNorm I = 0 ↔ I = ⊥ := b · rintro rfl exact absNorm_bot -theorem absNorm_ne_zero_of_nonZeroDivisors (I : (Ideal S)⁰) : Ideal.absNorm (I : Ideal S) ≠ 0 := - Ideal.absNorm_eq_zero_iff.not.mpr <| nonZeroDivisors.coe_ne_zero _ +theorem absNorm_ne_zero_iff_mem_nonZeroDivisors {I : Ideal S} : + absNorm I ≠ 0 ↔ I ∈ (Ideal S)⁰ := by + simp_rw [ne_eq, Ideal.absNorm_eq_zero_iff, mem_nonZeroDivisors_iff_ne_zero, Submodule.zero_eq_bot] -theorem finite_setOf_absNorm_eq [CharZero S] {n : ℕ} (hn : 0 < n) : +theorem absNorm_pos_iff_mem_nonZeroDivisors {I : Ideal S} : + 0 < absNorm I ↔ I ∈ (Ideal S)⁰ := by + rw [← absNorm_ne_zero_iff_mem_nonZeroDivisors, Nat.pos_iff_ne_zero] + +theorem absNorm_ne_zero_of_nonZeroDivisors (I : (Ideal S)⁰) : absNorm (I : Ideal S) ≠ 0 := + absNorm_ne_zero_iff_mem_nonZeroDivisors.mpr (SetLike.coe_mem I) + +theorem absNorm_pos_of_nonZeroDivisors (I : (Ideal S)⁰) : 0 < absNorm (I : Ideal S) := + absNorm_pos_iff_mem_nonZeroDivisors.mpr (SetLike.coe_mem I) + +theorem finite_setOf_absNorm_eq [CharZero S] (n : ℕ) : {I : Ideal S | Ideal.absNorm I = n}.Finite := by - let f := fun I : Ideal S => Ideal.map (Ideal.Quotient.mk (@Ideal.span S _ {↑n})) I - refine @Set.Finite.of_finite_image _ _ _ f ?_ ?_ - · suffices Finite (S ⧸ @Ideal.span S _ {↑n}) by - let g := ((↑) : Ideal (S ⧸ @Ideal.span S _ {↑n}) → Set (S ⧸ @Ideal.span S _ {↑n})) - refine @Set.Finite.of_finite_image _ _ _ g ?_ SetLike.coe_injective.injOn - exact Set.Finite.subset (@Set.finite_univ _ (@Set.finite' _ this)) (Set.subset_univ _) - rw [← absNorm_ne_zero_iff, absNorm_span_singleton] - simpa only [Ne, Int.natAbs_eq_zero, Algebra.norm_eq_zero_iff, Nat.cast_eq_zero] using - ne_of_gt hn - · intro I hI J hJ h - rw [← comap_map_mk (span_singleton_absNorm_le I), ← hI.symm, ← - comap_map_mk (span_singleton_absNorm_le J), ← hJ.symm] - congr + obtain hn | hn := Nat.eq_zero_or_pos n + · simp only [hn, absNorm_eq_zero_iff, Set.setOf_eq_eq_singleton, Set.finite_singleton] + · let f := fun I : Ideal S => Ideal.map (Ideal.Quotient.mk (@Ideal.span S _ {↑n})) I + refine @Set.Finite.of_finite_image _ _ _ f ?_ ?_ + · suffices Finite (S ⧸ @Ideal.span S _ {↑n}) by + let g := ((↑) : Ideal (S ⧸ @Ideal.span S _ {↑n}) → Set (S ⧸ @Ideal.span S _ {↑n})) + refine @Set.Finite.of_finite_image _ _ _ g ?_ SetLike.coe_injective.injOn + exact Set.Finite.subset (@Set.finite_univ _ (@Set.finite' _ this)) (Set.subset_univ _) + rw [← absNorm_ne_zero_iff, absNorm_span_singleton] + simpa only [Ne, Int.natAbs_eq_zero, Algebra.norm_eq_zero_iff, Nat.cast_eq_zero] using + ne_of_gt hn + · intro I hI J hJ h + rw [← comap_map_mk (span_singleton_absNorm_le I), ← hI.symm, ← + comap_map_mk (span_singleton_absNorm_le J), ← hJ.symm] + congr + +theorem finite_setOf_absNorm_le [CharZero S] (n : ℕ) : + {I : Ideal S | Ideal.absNorm I ≤ n}.Finite := by + rw [show {I : Ideal S | Ideal.absNorm I ≤ n} = + (⋃ i ∈ Set.Icc 0 n, {I : Ideal S | Ideal.absNorm I = i}) by ext; simp] + refine Set.Finite.biUnion (Set.finite_Icc 0 n) (fun i _ => Ideal.finite_setOf_absNorm_eq i) + +theorem card_norm_le_eq_card_norm_le_add_one (n : ℕ) [CharZero S] : + Nat.card {I : Ideal S // absNorm I ≤ n} = + Nat.card {I : (Ideal S)⁰ // absNorm (I : Ideal S) ≤ n} + 1 := by + classical + have : Finite {I : Ideal S // I ∈ (Ideal S)⁰ ∧ absNorm I ≤ n} := + (finite_setOf_absNorm_le n).subset fun _ ⟨_, h⟩ ↦ h + have : Finite {I : Ideal S // I ∉ (Ideal S)⁰ ∧ absNorm I ≤ n} := + (finite_setOf_absNorm_le n).subset fun _ ⟨_, h⟩ ↦ h + rw [Nat.card_congr (Equiv.subtypeSubtypeEquivSubtypeInter (fun I ↦ I ∈ (Ideal S)⁰) + (fun I ↦ absNorm I ≤ n))] + let e : {I : Ideal S // absNorm I ≤ n} ≃ {I : Ideal S // I ∈ (Ideal S)⁰ ∧ absNorm I ≤ n} ⊕ + {I : Ideal S // I ∉ (Ideal S)⁰ ∧ absNorm I ≤ n} := by + refine (Equiv.subtypeEquivRight ?_).trans (subtypeOrEquiv _ _ ?_) + · intro _ + simp_rw [← or_and_right, em, true_and] + · exact Pi.disjoint_iff.mpr fun I ↦ Prop.disjoint_iff.mpr (by tauto) + simp_rw [Nat.card_congr e, Nat.card_sum, add_right_inj] + conv_lhs => + enter [1, 1, I] + rw [← absNorm_ne_zero_iff_mem_nonZeroDivisors, ne_eq, not_not, and_iff_left_iff_imp.mpr + (fun h ↦ by rw [h]; exact Nat.zero_le n), absNorm_eq_zero_iff] + rw [Nat.card_unique] theorem norm_dvd_iff {x : S} (hx : Prime (Algebra.norm ℤ x)) {y : ℤ} : Algebra.norm ℤ x ∣ y ↔ x ∣ y := by diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean index b47e95b56f4ad..22685776831b5 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean @@ -198,6 +198,17 @@ lemma HasProd.tprod_fiberwise [T2Space α] {f : β → α} {a : α} (hf : HasPro (((Equiv.sigmaFiberEquiv g).hasProd_iff).mpr hf).sigma <| fun _ ↦ ((hf.multipliable.subtype _).hasProd_iff).mpr rfl +@[to_additive] +theorem tprod_card_pow_eq_tprod [T2Space α] {u : β → γ} (hu : ∀ n, {k | u k = n}.Finite) + (f : γ → α) (hf : Multipliable (fun n ↦ f (u n))) : + ∏' n, f n ^ Nat.card {k | u k = n} = ∏' n, f (u n) := by + apply HasProd.tprod_eq + convert (HasProd.tprod_fiberwise hf.hasProd u) with n + have : Fintype {k | u k = n} := (hu n).fintype + rw [← Equiv.tprod_eq (Equiv.setCongr (by rfl :{k | u k = n} = u ⁻¹' {n})), tprod_fintype, + Finset.prod_congr rfl (fun x _ ↦ by rw [Equiv.setCongr_apply, x.prop]), Finset.prod_const, + Nat.card_eq_fintype_card, Finset.card_univ] + section CompleteT0Space variable [T0Space α] diff --git a/Mathlib/Topology/Algebra/Module/Basic.lean b/Mathlib/Topology/Algebra/Module/Basic.lean index 3a3b76145aaeb..0bc139cd0b327 100644 --- a/Mathlib/Topology/Algebra/Module/Basic.lean +++ b/Mathlib/Topology/Algebra/Module/Basic.lean @@ -1757,6 +1757,10 @@ protected def refl : M₁ ≃L[R₁] M₁ := end +@[simp] +theorem refl_apply (x : M₁) : + ContinuousLinearEquiv.refl R₁ M₁ x = x := rfl + @[simp, norm_cast] theorem coe_refl : ↑(ContinuousLinearEquiv.refl R₁ M₁) = ContinuousLinearMap.id R₁ M₁ := rfl @@ -2025,6 +2029,32 @@ def arrowCongrEquiv (e₁₂ : M₁ ≃SL[σ₁₂] M₂) (e₄₃ : M₄ ≃SL[ ContinuousLinearMap.ext fun x => by simp only [ContinuousLinearMap.comp_apply, apply_symm_apply, coe_coe] +/-- Combine a family of continuous linear equivalences into a continuous linear equivalence of +pi-types. -/ +def piCongrRight {ι : Type*} {M : ι → Type*} [∀ i, TopologicalSpace (M i)] + [∀ i, AddCommMonoid (M i)] [∀ i, Module R₁ (M i)] {N : ι → Type*} [∀ i, TopologicalSpace (N i)] + [∀ i, AddCommMonoid (N i)] [∀ i, Module R₁ (N i)] (f : (i : ι) → M i ≃L[R₁] N i) : + ((i : ι) → M i) ≃L[R₁] (i : ι) → N i := + { LinearEquiv.piCongrRight fun i ↦ f i with + continuous_toFun := by + exact continuous_pi fun i ↦ (f i).continuous_toFun.comp (continuous_apply i) + continuous_invFun := by + exact continuous_pi fun i => (f i).continuous_invFun.comp (continuous_apply i) } + +@[simp] +theorem piCongrRight_apply {ι : Type*} {M : ι → Type*} [∀ i, TopologicalSpace (M i)] + [∀ i, AddCommMonoid (M i)] [∀ i, Module R₁ (M i)] {N : ι → Type*} [∀ i, TopologicalSpace (N i)] + [∀ i, AddCommMonoid (N i)] [∀ i, Module R₁ (N i)] (f : (i : ι) → M i ≃L[R₁] N i) + (m : (i : ι) → M i) (i : ι) : + ContinuousLinearEquiv.piCongrRight f m i = (f i) (m i) := rfl + +@[simp] +theorem piCongrRight_symm_apply {ι : Type*} {M : ι → Type*} [∀ i, TopologicalSpace (M i)] + [∀ i, AddCommMonoid (M i)] [∀ i, Module R₁ (M i)] {N : ι → Type*} [∀ i, TopologicalSpace (N i)] + [∀ i, AddCommMonoid (N i)] [∀ i, Module R₁ (N i)] (f : (i : ι) → M i ≃L[R₁] N i) + (n : (i : ι) → N i) (i : ι) : + (ContinuousLinearEquiv.piCongrRight f).symm n i = (f i).symm (n i) := rfl + end AddCommMonoid section AddCommGroup @@ -2060,6 +2090,26 @@ theorem skewProd_symm_apply (e : M ≃L[R] M₂) (e' : M₃ ≃L[R] M₄) (f : M (e.skewProd e' f).symm x = (e.symm x.1, e'.symm (x.2 - f (e.symm x.1))) := rfl +variable (R) in +/-- The negation map as a continuous linear equivalence. -/ +def neg [ContinuousNeg M] : + M ≃L[R] M := + { LinearEquiv.neg R with + continuous_toFun := continuous_neg + continuous_invFun := continuous_neg } + +@[simp] +theorem coe_neg [ContinuousNeg M] : + (neg R : M → M) = -id := rfl + +@[simp] +theorem neg_apply [ContinuousNeg M] (x : M) : + neg R x = -x := by simp + +@[simp] +theorem symm_neg [ContinuousNeg M] : + (neg R : M ≃L[R] M).symm = neg R := rfl + end AddCommGroup section Ring diff --git a/Mathlib/Topology/Basic.lean b/Mathlib/Topology/Basic.lean index 1c5fc5ac20a5c..ad216bac8cc58 100644 --- a/Mathlib/Topology/Basic.lean +++ b/Mathlib/Topology/Basic.lean @@ -635,6 +635,25 @@ theorem frontier_union_subset (s t : Set X) : frontier (s ∪ t) ⊆ frontier s ∩ closure tᶜ ∪ closure sᶜ ∩ frontier t := by simpa only [frontier_compl, ← compl_union] using frontier_inter_subset sᶜ tᶜ +/-- See `frontier_union_subset` for a stronger version. -/ +theorem frontier_union_subset' (s t : Set X) : + frontier (s ∪ t) ⊆ frontier s ∪ frontier t := + (frontier_union_subset s t).trans <| union_subset_union inter_subset_left inter_subset_right + +theorem Finset.frontier_biUnion {ι : Type*} (s : Finset ι) (t : ι → Set X) : + frontier (⋃ i ∈ s, t i) ⊆ ⋃ i ∈ s, frontier (t i) := by + classical + induction s using Finset.induction with + | empty => simp + | insert _ h_ind => + simp_rw [Finset.mem_insert, iUnion_iUnion_eq_or_left] + exact (frontier_union_subset' _ _).trans (Set.union_subset_union subset_rfl h_ind) + +theorem frontier_iUnion {ι : Type*} [Fintype ι] (t : ι → Set X) : + frontier (⋃ i, t i) ⊆ ⋃ i, frontier (t i) := by + have := Finset.frontier_biUnion Finset.univ t + simpa only [Finset.mem_univ, Set.iUnion_true] using this + theorem IsClosed.frontier_eq (hs : IsClosed s) : frontier s = s \ interior s := by rw [frontier, hs.closure_eq] diff --git a/Mathlib/Topology/ContinuousOn.lean b/Mathlib/Topology/ContinuousOn.lean index a09c8362c391f..6135b6c61d418 100644 --- a/Mathlib/Topology/ContinuousOn.lean +++ b/Mathlib/Topology/ContinuousOn.lean @@ -1140,6 +1140,18 @@ protected lemma Continuous.mulIndicator (hs : ∀ a ∈ frontier s, f a = 1) (hf Continuous (mulIndicator s f) := by classical exact hf.piecewise hs continuous_const +@[to_additive] +theorem ContinuousOn.continuousAt_mulIndicator (hf : ContinuousOn f (interior s)) {x : α} + (hx : x ∉ frontier s) : + ContinuousAt (s.mulIndicator f) x := by + rw [← Set.not_mem_compl_iff, Set.not_not_mem, compl_frontier_eq_union_interior] at hx + obtain h | h := hx + · have hs : interior s ∈ 𝓝 x := mem_interior_iff_mem_nhds.mp (by rwa [interior_interior]) + exact ContinuousAt.congr (hf.continuousAt hs) <| Filter.eventuallyEq_iff_exists_mem.mpr + ⟨interior s, hs, Set.eqOn_mulIndicator.symm.mono interior_subset⟩ + · exact ContinuousAt.congr continuousAt_const <| Filter.eventuallyEq_iff_exists_mem.mpr + ⟨sᶜ, mem_interior_iff_mem_nhds.mp h, Set.eqOn_mulIndicator'.symm⟩ + end Indicator theorem IsOpen.ite' {s s' t : Set α} (hs : IsOpen s) (hs' : IsOpen s') diff --git a/Mathlib/Topology/PartialHomeomorph.lean b/Mathlib/Topology/PartialHomeomorph.lean index 3cecb43406b4e..d61961dc925b0 100644 --- a/Mathlib/Topology/PartialHomeomorph.lean +++ b/Mathlib/Topology/PartialHomeomorph.lean @@ -907,7 +907,7 @@ variable {ι : Type*} [Finite ι] {X Y : ι → Type*} [∀ i, TopologicalSpace [∀ i, TopologicalSpace (Y i)] (ei : ∀ i, PartialHomeomorph (X i) (Y i)) /-- The product of a finite family of `PartialHomeomorph`s. -/ -@[simps toPartialEquiv] +@[simps! toPartialEquiv apply symm_apply source target] def pi : PartialHomeomorph (∀ i, X i) (∀ i, Y i) where toPartialEquiv := PartialEquiv.pi fun i => (ei i).toPartialEquiv open_source := isOpen_set_pi finite_univ fun i _ => (ei i).open_source From 0dc4ddb50acb3e8c9e136745e520cf28952b81e7 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Fri, 25 Oct 2024 13:33:24 +0200 Subject: [PATCH 2/4] Fix after merge --- .../Analysis/BoxIntegral/UnitPartition.lean | 6 +- .../Analysis/SpecialFunctions/PolarCoord.lean | 2 +- Mathlib/MeasureTheory/Constructions/Pi.lean | 4 +- Mathlib/NumberTheory/LSeries/Residue.lean | 8 +- .../CanonicalEmbedding/NormLessThanOne.lean | 126 +++++++++--------- 5 files changed, 72 insertions(+), 74 deletions(-) diff --git a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean index fdd2ba91a9caf..e8838e4dc91ad 100644 --- a/Mathlib/Analysis/BoxIntegral/UnitPartition.lean +++ b/Mathlib/Analysis/BoxIntegral/UnitPartition.lean @@ -121,7 +121,7 @@ abbrev tag (ν : ι → ℤ) : ι → ℝ := fun i ↦ (ν i + 1) / n theorem tag_apply (ν : ι → ℤ) (i : ι) : tag n ν i = (ν i + 1) / n := rfl theorem tag_injective : Function.Injective (fun ν : ι → ℤ ↦ tag n ν) := by - refine fun _ _ h ↦ Function.funext_iff.mpr fun i ↦ ?_ + refine fun _ _ h ↦ funext_iff.mpr fun i ↦ ?_ have := congr_arg (fun x ↦ x i) h field_simp at this exact this @@ -143,7 +143,7 @@ theorem index_apply {x : ι → ℝ} (i : ι) : variable {n} in theorem mem_box_iff_index {x : ι → ℝ} {ν : ι → ℤ} : x ∈ box n ν ↔ index n x = ν := by - simp_rw [mem_box_iff', Function.funext_iff, index_apply, sub_eq_iff_eq_add, Int.ceil_eq_iff, + simp_rw [mem_box_iff', funext_iff, index_apply, sub_eq_iff_eq_add, Int.ceil_eq_iff, Int.cast_add, Int.cast_one, add_sub_cancel_right] @[simp] @@ -393,7 +393,7 @@ theorem tendsto_tsum_div_pow (hF : Continuous F) (hs₁ : Bornology.IsBounded s) rw [← integralSum_eq_tsum_div _ s F hB hs₀, ← Measure.restrict_restrict_of_subset hs₀, ← MeasureTheory.integral_indicator hs₂] refine hr₂ 0 _ ⟨?_, fun _ ↦ ?_, fun h ↦ ?_, fun h ↦ ?_⟩ (prepartition_isPartition _ hB) - · rw [show r 0 = fun _ ↦ r 0 0 from Function.funext_iff.mpr (hr₁ 0 rfl)] + · rw [show r 0 = fun _ ↦ r 0 0 from funext_iff.mpr (hr₁ 0 rfl)] apply prepartition_isSubordinate n B rw [one_div, inv_le_comm₀ (Nat.cast_pos.mpr <| PNat.pos n) (by convert (r 0 0).prop)] exact le_trans (Nat.le_ceil _) (Nat.cast_le.mpr hn) diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index fbcd4627f44d1..e4b1e94dc63d3 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -163,7 +163,7 @@ theorem lintegral_comp_polarCoord_symm (f : ℝ × ℝ → ENNReal) : _ = ∫⁻ (p : ℝ × ℝ) in polarCoord.target, |p.1|.toNNReal • f (polarCoord.symm p) := by rw [lintegral_image_eq_lintegral_abs_det_fderiv_mul volume _ (fun p _ ↦ (hasFDerivAt_polarCoord_symm p).hasFDerivWithinAt)] - simp_rw [FDerivAt_polarCoord_symm_det]; rfl + · simp_rw [FDerivAt_polarCoord_symm_det]; rfl exacts [polarCoord.symm.injOn, measurableSet_Ioi.prod measurableSet_Ioo] _ = ∫⁻ (p : ℝ × ℝ) in polarCoord.target, (p.1).toNNReal • f (polarCoord.symm p) := by refine setLIntegral_congr_fun polarCoord.open_target.measurableSet ?_ diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index 1b16155ed9328..e0fb5385811c9 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -774,8 +774,8 @@ theorem measurePreserving_arrowProdEquivProdArrow (α β γ : Type*) [Measurable ((FiniteSpanningSetsIn.pi fun i ↦ (μ i).toFiniteSpanningSetsIn).prod (FiniteSpanningSetsIn.pi (fun i ↦ (ν i).toFiniteSpanningSetsIn))) ?_).symm · refine (generateFrom_eq_prod generateFrom_pi generateFrom_pi ?_ ?_).symm - exact (FiniteSpanningSetsIn.pi (fun i ↦ (μ i).toFiniteSpanningSetsIn)).isCountablySpanning - exact (FiniteSpanningSetsIn.pi (fun i ↦ (ν i).toFiniteSpanningSetsIn)).isCountablySpanning + · exact (FiniteSpanningSetsIn.pi (fun i ↦ (μ i).toFiniteSpanningSetsIn)).isCountablySpanning + · exact (FiniteSpanningSetsIn.pi (fun i ↦ (ν i).toFiniteSpanningSetsIn)).isCountablySpanning · rintro _ ⟨s, ⟨s, _, rfl⟩, ⟨_, ⟨t, _, rfl⟩, rfl⟩⟩ rw [MeasurableEquiv.map_apply, MeasurableEquiv.arrowProdEquivProdArrow, MeasurableEquiv.coe_mk] diff --git a/Mathlib/NumberTheory/LSeries/Residue.lean b/Mathlib/NumberTheory/LSeries/Residue.lean index a56dd53225e7f..2e79f9d55d4d3 100644 --- a/Mathlib/NumberTheory/LSeries/Residue.lean +++ b/Mathlib/NumberTheory/LSeries/Residue.lean @@ -183,7 +183,7 @@ theorem lt_u_rpow_lt {ε : ℝ} (hε₁ : 0 < ε) (hε₂ : ε ≤ l) : rw [← Real.mul_rpow, ← Real.mul_rpow, Real.rpow_lt_rpow_iff, Real.rpow_lt_rpow_iff, mul_inv_lt_iff₀, lt_mul_inv_iff₀, ← neg_add_lt_iff_lt_add, sub_eq_add_neg, ← lt_neg_add_iff_add_lt (a := l), neg_add_eq_sub, ← abs_lt, mul_comm] - exact h + · exact h all_goals positivity theorem summable_u_rpow {s : ℝ} (hs : 1 < s) : @@ -198,9 +198,9 @@ theorem summable_u_rpow {s : ℝ} (hs : 1 < s) : refine Eventually.isBigO ?_ filter_upwards [this] with n hn rw [Real.norm_eq_abs, abs_of_nonneg] - exact (hn s (lt_trans zero_lt_one hs)).2.le - refine Real.rpow_nonneg ?_ _ - exact Nat.cast_nonneg _ + · exact (hn s (lt_trans zero_lt_one hs)).2.le + · refine Real.rpow_nonneg ?_ _ + exact Nat.cast_nonneg _ theorem exist_finset_lt_tsum_u_lt {ε : ℝ} (hε₁ : 0 < ε) (hε₂ : ε ≤ l) : ∃ T : Finset ℕ, ∀ s, 1 < s → diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean index a4cde6fece926..268f90214426a 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean @@ -326,10 +326,10 @@ def mapToUnitsPow₀_aux : refine ⟨Set.mem_setOf.mpr fun w ↦ by split_ifs <;> exact Real.exp_pos _, ?_⟩ simp_rw [Set.mem_setOf_eq, ← Finset.univ.sum_erase_add _ (mem_univ w₀), dif_pos] rw [sum_subtype _ (by aesop : ∀ w, w ∈ univ.erase w₀ ↔ w ≠ w₀)] - conv_lhs => enter [1,2,w]; rw [dif_neg w.prop] - simp_rw [Real.log_exp, neg_mul, mul_neg, mul_inv_cancel_left₀ mult_coe_ne_zero, - add_neg_eq_zero] - infer_instance + · conv_lhs => enter [1,2,w]; rw [dif_neg w.prop] + simp_rw [Real.log_exp, neg_mul, mul_neg, mul_inv_cancel_left₀ mult_coe_ne_zero, + add_neg_eq_zero] + · infer_instance map_target' _ _ := trivial left_inv' := by intro x _ @@ -544,11 +544,9 @@ def mapToUnitsPow : PartialHomeomorph (realSpace K) (realSpace K) where enter [2, 2, w] rw [dif_neg w.prop] ext w - rw [Pi.smul_apply, ← logMap_real_smul] - rw [← _root_.map_smul] - rw [← mapToUnitsPow₀_symm_apply_of_norm_one h₀] - rw [PartialEquiv.right_inv, Pi.smul_apply, smul_eq_mul, smul_eq_mul] - rw [abs_of_nonneg, Real.rpow_neg, mul_inv_cancel_left₀] + rw [Pi.smul_apply, ← logMap_real_smul, ← _root_.map_smul, + ← mapToUnitsPow₀_symm_apply_of_norm_one h₀, PartialEquiv.right_inv, Pi.smul_apply, + smul_eq_mul, smul_eq_mul, abs_of_nonneg, Real.rpow_neg, mul_inv_cancel_left₀] · refine Real.rpow_ne_zero_of_pos ?_ _ exact pos_norm_realToMixed hx' · exact mixedEmbedding.norm_nonneg (realToMixed x) @@ -578,8 +576,8 @@ def mapToUnitsPow : PartialHomeomorph (realSpace K) (realSpace K) where refine Continuous.continuousOn ?_ refine Continuous.rpow_const ?_ ?_ · refine Continuous.comp' ?_ ?_ - exact mixedEmbedding.continuous_norm K - exact ContinuousLinearMap.continuous realToMixed + · exact mixedEmbedding.continuous_norm K + · exact ContinuousLinearMap.continuous realToMixed · intro _ right rw [inv_nonneg] @@ -588,11 +586,11 @@ def mapToUnitsPow : PartialHomeomorph (realSpace K) (realSpace K) where refine Continuous.comp_continuousOn' (continuous_apply _) <| (continuous_equivFun_basis _).comp_continuousOn' ?_ refine ContinuousOn.comp'' (continuousOn_logMap K) ?_ ?_ - refine Continuous.continuousOn ?_ - exact ContinuousLinearMap.continuous realToMixed - intro x hx - refine ne_of_gt ?_ - exact pos_norm_realToMixed (fun w ↦ (hx w).ne') + · refine Continuous.continuousOn ?_ + exact ContinuousLinearMap.continuous realToMixed + · intro x hx + refine ne_of_gt ?_ + exact pos_norm_realToMixed (fun w ↦ (hx w).ne') theorem measurable_mapToUnitsPow_symm : Measurable (mapToUnitsPow K).symm := by @@ -929,16 +927,16 @@ theorem mapToUnitsPowComplex_prod_indicator_aux (x y : ℝ × ℝ) (hx : x ∈ S · rw [lt_iff_le_and_ne] exact ⟨hx.2.2, this.2.2⟩ have := (Complex.polarCoord_symm_mem_polarCoord_source x).mp ?_ - have hx₁ : 0 < x.1 := by - refine lt_iff_le_and_ne.mpr ⟨?_, ?_⟩ - exact hx.1 - exact this.1.symm - · refine ⟨?_, ?_, ?_⟩ - · exact this.1 - · have := this.2.1 hx₁ (-1) - rwa [show π + (-1 : ℤ) * (2 * π) = -π by ring, ne_comm] at this - · have := this.2.1 hx₁ 0 - rwa [Int.cast_zero, zero_mul, add_zero, ne_comm] at this + · have hx₁ : 0 < x.1 := by + refine lt_iff_le_and_ne.mpr ⟨?_, ?_⟩ + exact hx.1 + exact this.1.symm + · refine ⟨?_, ?_, ?_⟩ + · exact this.1 + · have := this.2.1 hx₁ (-1) + rwa [show π + (-1 : ℤ) * (2 * π) = -π by ring, ne_comm] at this + · have := this.2.1 hx₁ 0 + rwa [Int.cast_zero, zero_mul, add_zero, ne_comm] at this · rw [hxy] exact Complex.polarCoord.map_target hy @@ -1086,25 +1084,25 @@ theorem volume_mapToUnitsPowComplex_set_prod_set {s : Set (InfinitePlace K → ∫⁻ x in mapToUnitsPow K '' (s ∩ {x | 0 < x w₀}), ∏ w : {w // IsComplex w}, (x w.val).toNNReal := by rw [← lintegral_indicator, ← lintegral_indicator] - congr with x - rw [Set.indicator_mul_right, Set.indicator_indicator, Set.inter_eq_right.mpr] - by_cases hx : x ∈ (mapToUnitsPow K) '' (s ∩ {x | 0 < x w₀}) - · rw [Set.indicator_of_mem hx, Set.indicator_of_mem hx, Pi.one_apply, mul_one, - ENNReal.coe_finset_prod] - · rw [Set.indicator_of_not_mem hx, Set.indicator_of_not_mem hx, mul_zero] - · rintro _ ⟨x, hx, rfl⟩ - refine Set.mem_univ_pi.mpr fun _ ↦ ?_ - rw [Set.mem_ite_univ_left] - intro _ - rw [Set.mem_Ioi] - exact mapToUnitsPow_pos (ne_of_gt hx.2) _ + · congr with x + rw [Set.indicator_mul_right, Set.indicator_indicator, Set.inter_eq_right.mpr] + · by_cases hx : x ∈ (mapToUnitsPow K) '' (s ∩ {x | 0 < x w₀}) + · rw [Set.indicator_of_mem hx, Set.indicator_of_mem hx, Pi.one_apply, mul_one, + ENNReal.coe_finset_prod] + · rw [Set.indicator_of_not_mem hx, Set.indicator_of_not_mem hx, mul_zero] + · rintro _ ⟨x, hx, rfl⟩ + refine Set.mem_univ_pi.mpr fun _ ↦ ?_ + rw [Set.mem_ite_univ_left] + intro _ + rw [Set.mem_Ioi] + exact mapToUnitsPow_pos (ne_of_gt hx.2) _ · refine measurable_mapToUnitsPow_image ?_ ?_ · exact hs.inter (measurableSet_lt measurable_const (measurable_pi_apply w₀)) · exact Set.inter_subset_left.trans hs' · refine MeasurableSet.univ_pi fun _ ↦ ?_ refine MeasurableSet.ite' (fun _ ↦ ?_) (fun _ ↦ ?_) - exact MeasurableSet.univ - exact measurableSet_Ioi + · exact MeasurableSet.univ + · exact measurableSet_Ioi end mapToUnitsPowComplex @@ -1378,28 +1376,28 @@ theorem volume_normLessThanOnePlus_aux (n : ℕ) : classical rw [volume_pi, box₁, restrict_pi_pi, lintegral_eq_lmarginal_univ 0, lmarginal_erase' _ ?_ (Finset.mem_univ w₀)] - simp_rw [if_true, Function.update_same] - have : ∫⁻ (xᵢ : ℝ) in Set.Ioc 0 1, ENNReal.ofReal |xᵢ| ^ n = (n + 1 : ENNReal)⁻¹ := by - convert congr_arg ENNReal.ofReal (integral_pow (a := 0) (b := 1) n) - · rw [intervalIntegral.integral_of_le zero_le_one] - rw [ofReal_integral_eq_lintegral_ofReal] - · refine setLIntegral_congr_fun measurableSet_Ioc ?_ - filter_upwards with _ h using by rw [abs_of_pos h.1, ENNReal.ofReal_pow h.1.le] - · refine IntegrableOn.integrable ?_ - rw [← Set.uIoc_of_le zero_le_one, ← intervalIntegrable_iff] - exact intervalIntegral.intervalIntegrable_pow n - · exact ae_restrict_of_forall_mem measurableSet_Ioc fun _ h ↦ pow_nonneg h.1.le _ - · rw [one_pow, zero_pow (by linarith), sub_zero, ENNReal.ofReal_div_of_pos (by positivity), - ENNReal.ofReal_add (by positivity) zero_le_one, ENNReal.ofReal_one, ENNReal.ofReal_natCast, - one_div] - rw [this] - rw [lmarginal] - rw [lintegral_const] - rw [pi_univ] - rw [Finset.prod_congr rfl (g := fun _ ↦ 1) (fun x _ ↦ by rw [if_neg (by aesop), restrict_apply - MeasurableSet.univ, Set.univ_inter, Real.volume_Ico, sub_zero, ENNReal.ofReal_one])] - rw [Finset.prod_const_one, mul_one] - fun_prop + · simp_rw [if_true, Function.update_same] + have : ∫⁻ (xᵢ : ℝ) in Set.Ioc 0 1, ENNReal.ofReal |xᵢ| ^ n = (n + 1 : ENNReal)⁻¹ := by + convert congr_arg ENNReal.ofReal (integral_pow (a := 0) (b := 1) n) + · rw [intervalIntegral.integral_of_le zero_le_one] + rw [ofReal_integral_eq_lintegral_ofReal] + · refine setLIntegral_congr_fun measurableSet_Ioc ?_ + filter_upwards with _ h using by rw [abs_of_pos h.1, ENNReal.ofReal_pow h.1.le] + · refine IntegrableOn.integrable ?_ + rw [← Set.uIoc_of_le zero_le_one, ← intervalIntegrable_iff] + exact intervalIntegral.intervalIntegrable_pow n + · exact ae_restrict_of_forall_mem measurableSet_Ioc fun _ h ↦ pow_nonneg h.1.le _ + · rw [one_pow, zero_pow (by linarith), sub_zero, ENNReal.ofReal_div_of_pos (by positivity), + ENNReal.ofReal_add (by positivity) zero_le_one, ENNReal.ofReal_one, + ENNReal.ofReal_natCast, one_div] + rw [this] + rw [lmarginal] + rw [lintegral_const] + rw [pi_univ] + rw [Finset.prod_congr rfl (g := fun _ ↦ 1) (fun x _ ↦ by rw [if_neg (by aesop), restrict_apply + MeasurableSet.univ, Set.univ_inter, Real.volume_Ico, sub_zero, ENNReal.ofReal_one])] + rw [Finset.prod_const_one, mul_one] + · fun_prop open Classical in theorem volume_plusPart_normLessThanOne : volume (plusPart (normLessThanOne K)) = @@ -1573,14 +1571,14 @@ theorem volume_mapToUnitsPowComplex_interior_eq_volume_mapToUnitsPowComplex_clos open Classical in theorem volume_frontier_plusPart_normLessThanOne : volume (frontier (plusPart (normLessThanOne K))) = 0 := by - rw [frontier, measure_diff] have : volume (closure (plusPart (normLessThanOne K))) = volume (interior (plusPart (normLessThanOne K))) := by refine le_antisymm ?_ (measure_mono interior_subset_closure) refine le_trans ?_ (measure_mono (mapToUnitsPowComplex_interior_subset_interior K)) rw [volume_mapToUnitsPowComplex_interior_eq_volume_mapToUnitsPowComplex_closure] exact measure_mono (closure_subset_mapToUnitsPowComplex_closure K) - refine tsub_eq_zero_iff_le.mpr (le_of_eq this) + rw [frontier, measure_diff] + · exact tsub_eq_zero_iff_le.mpr (le_of_eq this) · exact interior_subset_closure · exact measurableSet_interior.nullMeasurableSet · rw [← lt_top_iff_ne_top] From 0661e20234a1441eeba0a6475de665beaebc20c0 Mon Sep 17 00:00:00 2001 From: Xavier Roblot Date: Sat, 26 Oct 2024 16:22:40 +0200 Subject: [PATCH 3/4] Fix after merge --- .../Analysis/SpecialFunctions/PolarCoord.lean | 9 ++++--- Mathlib/NumberTheory/LSeries/Residue.lean | 2 +- .../NumberField/CanonicalEmbedding/Basic.lean | 4 +-- .../CanonicalEmbedding/NormLessThanOne.lean | 26 +++++++++---------- .../NumberField/DedekindZeta.lean | 2 +- Mathlib/NumberTheory/NumberField/Ideal.lean | 6 ++--- 6 files changed, 25 insertions(+), 24 deletions(-) diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index e4b1e94dc63d3..21a84f33ecf87 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -210,7 +210,7 @@ theorem polarCoord_symm_abs (p : ℝ × ℝ) : open scoped ComplexOrder -theorem polarCoord_symm_mem_polarCoord_source (x : ℝ × ℝ) : +theorem polarCoord_symm_mem_polarCoord_source_iff {x : ℝ × ℝ} : Complex.polarCoord.symm x ∈ Complex.polarCoord.source ↔ x.1 ≠ 0 ∧ (x.1 > 0 → ∀ k : ℤ, π + k * (2 * π) ≠ x.2) ∧ (x.1 < 0 → ∀ k : ℤ, k * (2 * π) ≠ x.2) := by @@ -221,9 +221,10 @@ theorem polarCoord_symm_mem_polarCoord_source (x : ℝ × ℝ) : obtain hx | hx | hx := lt_trichotomy x.1 0 · simp_rw [hx, hx.ne, not_lt_of_gt hx, false_and, false_or, true_and, or_false] have : (x.1 * cexp (x.2 * I)).arg = π ↔ (cexp (x.2 * I)).arg = 0 := by - simp_rw [arg_eq_pi_iff_lt_zero, arg_eq_zero_iff_zero_le, ofReal_mul_neg_iff, hx, - not_lt_of_gt hx, true_and, false_and, or_false, lt_iff_le_and_ne, ne_eq, eq_comm, - exp_ne_zero, not_false_eq_true, and_true] + simp_rw [arg_eq_pi_iff_lt_zero, arg_eq_zero_iff_zero_le, ← coe_algebraMap, + RCLike.ofReal_mul_neg_iff, hx, not_lt_of_gt hx, true_and, false_and, or_false, + lt_iff_le_and_ne, ne_eq, eq_comm (a := 0) (b := cexp _), exp_ne_zero, not_false_eq_true, + and_true] simp_rw [this, arg_exp_mul_I, toIocMod_eq_iff, zero_add, zsmul_eq_mul, eq_comm, and_iff_right_iff_imp] exact fun _ ↦ ⟨neg_neg_iff_pos.mpr Real.pi_pos, by ring_nf; positivity⟩ diff --git a/Mathlib/NumberTheory/LSeries/Residue.lean b/Mathlib/NumberTheory/LSeries/Residue.lean index 2e79f9d55d4d3..251abfa201d02 100644 --- a/Mathlib/NumberTheory/LSeries/Residue.lean +++ b/Mathlib/NumberTheory/LSeries/Residue.lean @@ -62,7 +62,7 @@ theorem tendsto_rpow_mul_tsum_rpow {c : ℝ} (hc : c ≠ 0) (T : Finset ℕ) : refine Tendsto.mul (tendsto_nhdsWithin_of_tendsto_nhds ?_) ?_ · refine Continuous.tendsto' ?_ 1 c (by rw [Real.rpow_one]) exact continuous_const.rpow continuous_id fun _ ↦ Or.inl hc - · refine (riemannZeta_residue_one'.sub (tendsto_mul_sum_rpow T (fun n ↦ n))).congr' ?_ + · refine (tendsto_sub_mul_tsum_nat_rpow.sub (tendsto_mul_sum_rpow T (fun n ↦ n))).congr' ?_ filter_upwards [eventually_mem_nhdsWithin] with s hs simp_rw [sub_eq_iff_eq_add', ← mul_add, sum_add_tsum_compl (Real.summable_nat_rpow.mpr (neg_lt_neg_iff.mpr hs)), Real.rpow_neg (Nat.cast_nonneg _), one_div] diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index b10668e65019b..acf87061aa4ae 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -993,10 +993,10 @@ open Classical in `2^r₁` times the volume of its `plusPart` where `r₁` is the number of real places. -/ theorem volume_eq_two_pow_mul_volume_plusPart (hA : ∀ x, x ∈ A ↔ (fun w ↦ |x.1 w|, x.2) ∈ A) (hm : MeasurableSet A) : - volume A = 2 ^ NrRealPlaces K * volume (plusPart A) := by + volume A = 2 ^ nrRealPlaces K * volume (plusPart A) := by simp_rw [← measure_congr (iUnion_negAtPart_ae A hA), measure_iUnion (disjoint_negAtPlusPart A) (fun _ ↦ measurableSet_negAtPlusPart A _ hm), volume_negAtPlusPart _ hm, tsum_fintype, - Finset.sum_const, Finset.card_univ, NrRealPlaces, nsmul_eq_mul, Fintype.card_set, Nat.cast_pow, + Finset.sum_const, Finset.card_univ, nrRealPlaces, nsmul_eq_mul, Fintype.card_set, Nat.cast_pow, Nat.cast_ofNat] end plusPart diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean index 268f90214426a..9d9908eca1778 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/NormLessThanOne.lean @@ -712,7 +712,7 @@ open Classical in theorem mapToUnitsPow_jacobian_det {c : realSpace K} (hc : 0 ≤ c w₀) : |(mapToUnitsPow_jacobian c).det| = (∏ w : {w : InfinitePlace K // w.IsComplex }, mapToUnitsPow₀ K (fun w ↦ c w) w)⁻¹ * - 2⁻¹ ^ NrComplexPlaces K * |c w₀| ^ (rank K) * (finrank ℚ K) * regulator K := by + 2⁻¹ ^ nrComplexPlaces K * |c w₀| ^ (rank K) * (finrank ℚ K) * regulator K := by have : LinearMap.toMatrix' (mapToUnitsPow_jacobian c).toLinearMap = Matrix.of fun w i ↦ mapToUnitsPow₀ K (fun w ↦ c w) w * mapToUnitsPow_jacobianCoeff w i c := by @@ -730,12 +730,12 @@ theorem mapToUnitsPow_jacobian_det {c : realSpace K} (hc : 0 ≤ c w₀) : rw [Finset.prod_ite, Finset.prod_const_one, Finset.prod_const, one_mul, abs_pow] rw [Finset.filter_ne', Finset.card_erase_of_mem (Finset.mem_univ _), Finset.card_univ, rank] rw [this, mul_assoc, ← abs_mul, ← Matrix.det_mul_column] - have : (2 : ℝ)⁻¹ ^ NrComplexPlaces K = |∏ w : InfinitePlace K, (mult w : ℝ)⁻¹| := by + have : (2 : ℝ)⁻¹ ^ nrComplexPlaces K = |∏ w : InfinitePlace K, (mult w : ℝ)⁻¹| := by rw [Finset.abs_prod] simp_rw [mult, Nat.cast_ite, Nat.cast_one, Nat.cast_ofNat, apply_ite, abs_inv, abs_one, inv_one, Nat.abs_ofNat, Finset.prod_ite, Finset.prod_const_one, Finset.prod_const, one_mul] rw [Finset.filter_not, Finset.card_univ_diff, ← Fintype.card_subtype] - rw [card_eq_nrRealPlaces_add_nrComplexPlaces, ← NrRealPlaces, add_tsub_cancel_left] + rw [card_eq_nrRealPlaces_add_nrComplexPlaces, ← nrRealPlaces, add_tsub_cancel_left] rw [this, mul_assoc, ← abs_mul, ← Matrix.det_mul_row, abs_mul] congr · rw [abs_eq_self.mpr] @@ -781,7 +781,7 @@ open ENNReal Classical in theorem setLIntegral_mapToUnitsPow {s : Set (realSpace K)} (hs₀ : MeasurableSet s) (hs₁ : s ⊆ {x | 0 ≤ x w₀}) (f : (InfinitePlace K → ℝ) → ℝ≥0∞) : ∫⁻ x in (mapToUnitsPow K) '' s, f x = - (2 : ℝ≥0∞)⁻¹ ^ NrComplexPlaces K * ENNReal.ofReal (regulator K) * (finrank ℚ K) * + (2 : ℝ≥0∞)⁻¹ ^ nrComplexPlaces K * ENNReal.ofReal (regulator K) * (finrank ℚ K) * ∫⁻ x in s, ENNReal.ofReal |x w₀| ^ rank K * ENNReal.ofReal (∏ i : {w : InfinitePlace K // IsComplex w}, (mapToUnitsPow₀ K (fun w ↦ x w) i))⁻¹ * f (mapToUnitsPow K x) := by @@ -926,7 +926,7 @@ theorem mapToUnitsPowComplex_prod_indicator_aux (x y : ℝ × ℝ) (hx : x ∈ S exact ⟨hx.2.1, this.2.1.symm⟩ · rw [lt_iff_le_and_ne] exact ⟨hx.2.2, this.2.2⟩ - have := (Complex.polarCoord_symm_mem_polarCoord_source x).mp ?_ + have := (Complex.polarCoord_symm_mem_polarCoord_source_iff (x := x)).mp ?_ · have hx₁ : 0 < x.1 := by refine lt_iff_le_and_ne.mpr ⟨?_, ?_⟩ exact hx.1 @@ -1401,7 +1401,7 @@ theorem volume_normLessThanOnePlus_aux (n : ℕ) : open Classical in theorem volume_plusPart_normLessThanOne : volume (plusPart (normLessThanOne K)) = - NNReal.pi ^ NrComplexPlaces K * (regulator K).toNNReal := by + NNReal.pi ^ nrComplexPlaces K * (regulator K).toNNReal := by rw [plusPart_normLessThanOne, volume_mapToUnitsPowComplex_set_prod_set (measurableSet_box₁ K) (fun _ h ↦ le_of_lt (pos_of_mem_box₁ h)) (measurableSet_box₂ K) (Set.pi_mono fun _ _ ↦ Set.Ioc_subset_Icc_self) @@ -1415,16 +1415,16 @@ theorem volume_plusPart_normLessThanOne : volume (plusPart (normLessThanOne K)) have h₁ : ∀ x : InfinitePlace K → ℝ, 0 < ∏ i : {w // IsComplex w}, (mapToUnitsPow₀ K) (fun w ↦ x w) i.val := fun _ ↦ Finset.prod_pos fun _ _ ↦ mapToUnitsPow₀_pos _ _ - have h₂ : rank K + NrComplexPlaces K + 1 = finrank ℚ K := by + have h₂ : rank K + nrComplexPlaces K + 1 = finrank ℚ K := by rw [rank, add_comm _ 1, ← add_assoc, add_tsub_cancel_of_le NeZero.one_le, card_eq_nrRealPlaces_add_nrComplexPlaces, ← card_add_two_mul_card_eq_rank] ring calc - _ = (NNReal.pi : ENNReal) ^ NrComplexPlaces K * (regulator K).toNNReal * (finrank ℚ K) * - ∫⁻ x in box₁ K, ENNReal.ofReal |x w₀| ^ (rank K + NrComplexPlaces K) := by + _ = (NNReal.pi : ENNReal) ^ nrComplexPlaces K * (regulator K).toNNReal * (finrank ℚ K) * + ∫⁻ x in box₁ K, ENNReal.ofReal |x w₀| ^ (rank K + nrComplexPlaces K) := by simp_rw [← mul_assoc] congr - · rw [mul_comm, ← mul_assoc, NrComplexPlaces, Finset.card_univ, ← mul_pow, + · rw [mul_comm, ← mul_assoc, nrComplexPlaces, Finset.card_univ, ← mul_pow, ENNReal.inv_mul_cancel two_ne_zero ENNReal.two_ne_top, one_pow, one_mul, ← ENNReal.ofReal_coe_nnreal, NNReal.coe_real_pi] · ext x @@ -1434,8 +1434,8 @@ theorem volume_plusPart_normLessThanOne : volume (plusPart (normLessThanOne K)) ← ENNReal.ofReal_prod_of_nonneg (fun _ _ ↦ (mapToUnitsPow₀_pos _ _).le), ENNReal.ofReal_inv_of_pos (h₁ x), ENNReal.inv_mul_cancel (ENNReal.ofReal_ne_zero_iff.mpr (h₁ x)) ENNReal.ofReal_ne_top, mul_one, pow_add, - NrComplexPlaces, Finset.card_univ] - _ = NNReal.pi ^ NrComplexPlaces K * (regulator K).toNNReal := by + nrComplexPlaces, Finset.card_univ] + _ = NNReal.pi ^ nrComplexPlaces K * (regulator K).toNNReal := by rw [volume_normLessThanOnePlus_aux, ← Nat.cast_add_one, h₂, mul_assoc, ENNReal.mul_inv_cancel, mul_one] · rw [Nat.cast_ne_zero] @@ -1446,7 +1446,7 @@ theorem volume_plusPart_normLessThanOne : volume (plusPart (normLessThanOne K)) open Classical in theorem volume_normLessThanOne : volume (normLessThanOne K) = - 2 ^ NrRealPlaces K * NNReal.pi ^ NrComplexPlaces K * (regulator K).toNNReal := by + 2 ^ nrRealPlaces K * NNReal.pi ^ nrComplexPlaces K * (regulator K).toNNReal := by rw [volume_eq_two_pow_mul_volume_plusPart (normLessThanOne K) mem_normLessThanOne_iff (measurableSet_normLessThanOne K), volume_plusPart_normLessThanOne, mul_assoc] diff --git a/Mathlib/NumberTheory/NumberField/DedekindZeta.lean b/Mathlib/NumberTheory/NumberField/DedekindZeta.lean index b5859a158d9e9..be10854e77710 100644 --- a/Mathlib/NumberTheory/NumberField/DedekindZeta.lean +++ b/Mathlib/NumberTheory/NumberField/DedekindZeta.lean @@ -27,7 +27,7 @@ def dedekindZeta (s : ℂ) := /-- Docstring. -/ def residue : ℝ := - (2 ^ NrRealPlaces K * (2 * π) ^ NrComplexPlaces K * regulator K * classNumber K) / + (2 ^ nrRealPlaces K * (2 * π) ^ nrComplexPlaces K * regulator K * classNumber K) / (torsionOrder K * Real.sqrt |discr K|) theorem residue_pos : 0 < residue K := by diff --git a/Mathlib/NumberTheory/NumberField/Ideal.lean b/Mathlib/NumberTheory/NumberField/Ideal.lean index cf030425ddc1e..df31c730a3918 100644 --- a/Mathlib/NumberTheory/NumberField/Ideal.lean +++ b/Mathlib/NumberTheory/NumberField/Ideal.lean @@ -65,7 +65,7 @@ theorem ideal.tendsto_mk_eq_norm_le_div_atop (C : ClassGroup (𝓞 K)) : Tendsto (fun s : ℝ ↦ (Nat.card {I : (Ideal (𝓞 K))⁰ // absNorm (I : Ideal (𝓞 K)) ≤ s ∧ ClassGroup.mk0 I = C} : ℝ) / s) atTop - (𝓝 ((2 ^ NrRealPlaces K * (2 * π) ^ NrComplexPlaces K * regulator K) / + (𝓝 ((2 ^ nrRealPlaces K * (2 * π) ^ nrComplexPlaces K * regulator K) / (torsionOrder K * Real.sqrt |discr K|))) := by classical have h : ∀ s : ℝ, @@ -105,7 +105,7 @@ theorem ideal.tendsto_mk_eq_norm_le_div_atop (C : ClassGroup (𝓞 K)) : theorem ideal.tendsto_norm_le_div_atop₀ : Tendsto (fun s : ℝ ↦ (Nat.card {I : (Ideal (𝓞 K))⁰ // absNorm (I : Ideal (𝓞 K)) ≤ s} : ℝ) / s) atTop - (𝓝 ((2 ^ NrRealPlaces K * (2 * π) ^ NrComplexPlaces K * regulator K * classNumber K) / + (𝓝 ((2 ^ nrRealPlaces K * (2 * π) ^ nrComplexPlaces K * regulator K * classNumber K) / (torsionOrder K * Real.sqrt |discr K|))) := by classical convert Filter.Tendsto.congr' ?_ @@ -127,7 +127,7 @@ theorem ideal.tendsto_norm_le_div_atop₀ : theorem ideal.tendsto_norm_le_div_atop : Tendsto (fun s : ℝ ↦ (Nat.card {I : Ideal (𝓞 K) // absNorm I ≤ s} : ℝ) / s) atTop - (𝓝 ((2 ^ NrRealPlaces K * (2 * π) ^ NrComplexPlaces K * regulator K * classNumber K) / + (𝓝 ((2 ^ nrRealPlaces K * (2 * π) ^ nrComplexPlaces K * regulator K * classNumber K) / (torsionOrder K * Real.sqrt |discr K|))) := by have := (ideal.tendsto_norm_le_div_atop₀ K).add tendsto_inv_atTop_zero rw [add_zero] at this From d6b3f8d0b3dc54385e93bb58b7873c536c59faf8 Mon Sep 17 00:00:00 2001 From: Xavier Roblot <46200072+xroblot@users.noreply.github.com> Date: Sat, 26 Oct 2024 16:24:57 +0200 Subject: [PATCH 4/4] Update Mathlib/Analysis/SpecialFunctions/PolarCoord.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/Analysis/SpecialFunctions/PolarCoord.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean index 21a84f33ecf87..11510b0539d2a 100644 --- a/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean +++ b/Mathlib/Analysis/SpecialFunctions/PolarCoord.lean @@ -210,7 +210,7 @@ theorem polarCoord_symm_abs (p : ℝ × ℝ) : open scoped ComplexOrder -theorem polarCoord_symm_mem_polarCoord_source_iff {x : ℝ × ℝ} : +theorem polarCoord_symm_mem_polarCoord_source_iff {x : ℝ × ℝ} : Complex.polarCoord.symm x ∈ Complex.polarCoord.source ↔ x.1 ≠ 0 ∧ (x.1 > 0 → ∀ k : ℤ, π + k * (2 * π) ≠ x.2) ∧ (x.1 < 0 → ∀ k : ℤ, k * (2 * π) ≠ x.2) := by