diff --git a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean index 69a827b17b1c0..013f14fce60b1 100644 --- a/Mathlib/NumberTheory/NumberField/Discriminant/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Discriminant/Basic.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 @@ -72,6 +73,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 *