From 6079e9f391de2bb9fb71d73b02677501731044fe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 27 Jun 2023 09:38:16 +0200 Subject: [PATCH 01/67] First version --- .../NumberField/CanonicalEmbedding.lean | 206 ++++++------------ 1 file changed, 72 insertions(+), 134 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 8cc8355934c9d..965efcb66f4b9 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -9,14 +9,14 @@ Authors: Xavier Roblot ! if you have ported upstream changes. -/ import Mathlib.NumberTheory.NumberField.Embeddings +import Mathlib.RingTheory.Discriminant /-! # Canonical embedding of a number field -The canonical embedding of a number field `K` of signature `(r₁, r₂)` is the ring homomorphism -`K →+* ℝ^r₁ × ℂ^r₂` that sends `x ∈ K` to `(φ_₁(x),...,φ_r₁(x)) × (ψ_₁(x),..., ψ_r₂(x))` where -`φ_₁,...,φ_r₁` are its real embeddings and `ψ_₁,..., ψ_r₂` are its complex embeddings (up to -complex conjugation). +The canonical embedding of a number field `K` of degree `n` is the ring homomorphism +`K →+* ℂ^n` that sends `x ∈ K` to `(φ_₁(x),...,φ_n(x))` where the `φ_i`'s are the complex +embeddings of `K`. ## Main definitions and results @@ -29,159 +29,97 @@ radius is finite. number field, infinite places -/ - -noncomputable section - -open Function FiniteDimensional Finset Fintype NumberField NumberField.InfinitePlace Metric Module - -open scoped Classical NumberField - variable (K : Type _) [Field K] +open scoped NumberField + namespace NumberField.canonicalEmbedding --- The ambient space `ℝ^r₁ × ℂ^r₂` with `(r₁, r₂)` the signature of `K`. -set_option quotPrecheck false in -- Porting note: Added. -scoped[CanonicalEmbedding] - notation "E" => - ({ w : InfinitePlace K // IsReal w } → ℝ) × ({ w : InfinitePlace K // IsComplex w } → ℂ) - -open CanonicalEmbedding - -theorem space_rank [NumberField K] : finrank ℝ E = finrank ℚ K := by - haveI : Module.Free ℝ ℂ := inferInstance - rw [finrank_prod, finrank_pi, finrank_pi_fintype, Complex.finrank_real_complex, Finset.sum_const, - Finset.card_univ, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, - ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, - Nat.add_sub_of_le (Fintype.card_subtype_le _)] -#align number_field.canonical_embedding.space_rank NumberField.canonicalEmbedding.space_rank - -theorem nontrivial_space [NumberField K] : Nontrivial E := by - have : Nonempty <| InfinitePlace K := inferInstance - rcases this with ⟨w⟩ - obtain hw | hw := w.isReal_or_isComplex - · haveI : Nonempty { w : InfinitePlace K // IsReal w } := ⟨⟨w, hw⟩⟩ - exact nontrivial_prod_left - · haveI : Nonempty { w : InfinitePlace K // IsComplex w } := ⟨⟨w, hw⟩⟩ - exact nontrivial_prod_right -#align number_field.canonical_embedding.non_trivial_space NumberField.canonicalEmbedding.nontrivial_space - -/-- The canonical embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ -def _root_.NumberField.canonicalEmbedding : K →+* E := - RingHom.prod (Pi.ringHom fun w => w.prop.embedding) (Pi.ringHom fun w => w.val.embedding) -#align number_field.canonical_embedding NumberField.canonicalEmbedding - -theorem _root_.NumberField.canonicalEmbedding_injective [NumberField K] : - Injective (NumberField.canonicalEmbedding K) := - @RingHom.injective _ _ _ _ (nontrivial_space K) _ -#align number_field.canonical_embedding_injective NumberField.canonicalEmbedding_injective - -open NumberField +/-- The canonical embedding of a number field `K` of degree `n` into `ℂ^n`. -/ +def _root_.NumberField.canonicalEmbedding : K →+* ((K →+* ℂ) → ℂ) := Pi.ringHom fun φ => φ -variable {K} +lemma _root_.NumberField.canonicalEmbedding_injective [NumberField K] : + Function.Injective (NumberField.canonicalEmbedding K) := RingHom.injective _ -@[simp] -theorem apply_at_real_infinitePlace (w : { w : InfinitePlace K // IsReal w }) (x : K) : - (NumberField.canonicalEmbedding K x).1 w = w.prop.embedding x := by - simp only [canonicalEmbedding, RingHom.prod_apply, Pi.ringHom_apply] -#align number_field.canonical_embedding.apply_at_real_infinite_place NumberField.canonicalEmbedding.apply_at_real_infinitePlace +variable {K} @[simp] -theorem apply_at_complex_infinitePlace (w : { w : InfinitePlace K // IsComplex w }) (x : K) : - (NumberField.canonicalEmbedding K x).2 w = embedding w.val x := by - simp only [canonicalEmbedding, RingHom.prod_apply, Pi.ringHom_apply] -#align number_field.canonical_embedding.apply_at_complex_infinite_place NumberField.canonicalEmbedding.apply_at_complex_infinitePlace +lemma apply_at (φ : K →+* ℂ) (x : K) : + (NumberField.canonicalEmbedding K x) φ = φ x := rfl -theorem nnnorm_eq [NumberField K] (x : K) : +lemma nnnorm_eq [NumberField K] (x : K) : ‖canonicalEmbedding K x‖₊ = - Finset.univ.sup fun w : InfinitePlace K => ⟨w x, map_nonneg w x⟩ := by - rw [Prod.nnnorm_def', Pi.nnnorm_def, Pi.nnnorm_def] - rw [(_ : Finset.univ = - {w : InfinitePlace K | IsReal w}.toFinset ∪ {w : InfinitePlace K | IsComplex w}.toFinset)] - · rw [Finset.sup_union, sup_eq_max] - refine' congr_arg₂ _ _ _ - · convert - (Finset.univ.sup_map (Function.Embedding.subtype fun w : InfinitePlace K => IsReal w) - fun w => (⟨w x, map_nonneg w x⟩ : NNReal)).symm using 2 - ext w - dsimp - simp only [apply_at_real_infinitePlace, coe_nnnorm, Real.norm_eq_abs, - Function.Embedding.coe_subtype, Subtype.coe_mk, IsReal.abs_embedding_apply] - · convert - (Finset.univ.sup_map (Function.Embedding.subtype fun w : InfinitePlace K => IsComplex w) - fun w => (⟨w x, map_nonneg w x⟩ : NNReal)).symm using 2 - ext w - dsimp - simp only [apply_at_complex_infinitePlace, coe_nnnorm, - Complex.norm_eq_abs, Function.Embedding.coe_subtype, Subtype.coe_mk, abs_embedding] - · ext w - simp_rw [Finset.mem_univ, Finset.mem_union, Set.mem_toFinset, Set.mem_setOf_eq, - w.isReal_or_isComplex] -#align number_field.canonical_embedding.nnnorm_eq NumberField.canonicalEmbedding.nnnorm_eq - -theorem norm_le_iff [NumberField K] (x : K) (r : ℝ) : - ‖canonicalEmbedding K x‖ ≤ r ↔ ∀ w : InfinitePlace K, w x ≤ r := by + Finset.univ.sup (fun φ : K →+* ℂ => ‖φ x‖₊) := by + simp_rw [Pi.nnnorm_def, apply_at] + +lemma norm_le_iff [NumberField K] (x : K) (r : ℝ) : + ‖canonicalEmbedding K x‖ ≤ r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by obtain hr | hr := lt_or_le r 0 - · have : Nonempty <| InfinitePlace K := inferInstance - rcases this with ⟨w⟩ - exact iff_of_false - (hr.trans_le <| norm_nonneg _).not_le fun h => hr.not_le <| (map_nonneg w _).trans <| h _ + · obtain ⟨φ : K →+* ℂ⟩ := (inferInstance : Nonempty _) + refine iff_of_false ?_ ?_ + exact (hr.trans_le (norm_nonneg _)).not_le + exact fun h => hr.not_le (le_trans (norm_nonneg _) (h φ)) · lift r to NNReal using hr simp_rw [← coe_nnnorm, nnnorm_eq, NNReal.coe_le_coe, Finset.sup_le_iff, Finset.mem_univ, - forall_true_left, ← NNReal.coe_le_coe, NNReal.toReal] -#align number_field.canonical_embedding.norm_le_iff NumberField.canonicalEmbedding.norm_le_iff + forall_true_left] variable (K) -/-- The image of `𝓞 K` as a subring of `ℝ^r₁ × ℂ^r₂`. -/ -def integerLattice : Subring E := +/-- The image of `𝓞 K` as a subring of `ℂ^n`. -/ +def integerLattice : Subring ((K →+* ℂ) → ℂ) := (RingHom.range (algebraMap (𝓞 K) K)).map (canonicalEmbedding K) -#align number_field.canonical_embedding.integer_lattice NumberField.canonicalEmbedding.integerLattice - --- Porting note: See https://github.com/leanprover-community/mathlib4/issues/5028 -set_option maxHeartbeats 400000 in -set_option synthInstance.maxHeartbeats 50000 in -/-- The linear equiv between `𝓞 K` and the integer lattice. -/ -def equivIntegerLattice [NumberField K] : 𝓞 K ≃ₗ[ℤ] integerLattice K := - LinearEquiv.ofBijective - { toFun := fun x => (by - refine ⟨canonicalEmbedding K (algebraMap (𝓞 K) K x), ⟨algebraMap (𝓞 K) K x, ⟨?_, rfl⟩⟩⟩ - simp only [Subsemiring.coe_carrier_toSubmonoid, Subring.coe_toSubsemiring, - RingHom.coe_range, Set.mem_range, exists_apply_eq_apply] ) - map_add' := fun x y => (by - apply Subtype.eq - simp [map_add] ) - map_smul' := fun c x => (by - simp only [RingHom.id_apply, zsmul_eq_mul, RingHom.map_mul, map_intCast] - rfl ) } - (by - refine ⟨fun _ _ h => ?_, fun ⟨_, ⟨_, ⟨⟨a, rfl⟩, rfl⟩⟩⟩ => ⟨a, rfl⟩⟩ - dsimp only at h - rw [LinearMap.coe_mk, Subtype.mk_eq_mk] at h - exact IsFractionRing.injective (𝓞 K) K (canonicalEmbedding_injective K h)) -#align number_field.canonical_embedding.equiv_integer_lattice NumberField.canonicalEmbedding.equivIntegerLattice - -theorem integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : - ((integerLattice K : Set E) ∩ closedBall 0 r).Finite := by + +lemma integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : + ((integerLattice K : Set ((K →+* ℂ) → ℂ)) ∩ Metric.closedBall 0 r).Finite := by obtain hr | _ := lt_or_le r 0 - · simp [closedBall_eq_empty.2 hr] - have heq : ∀ x, canonicalEmbedding K x ∈ closedBall (0 : E) r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by - simp only [← place_apply, ← InfinitePlace.coe_mk, mem_closedBall_zero_iff, norm_le_iff] - exact fun x => le_iff_le x r - convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) - ext; constructor - · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx2⟩ - exact ⟨x, ⟨SetLike.coe_mem x, (heq x).mp hx2⟩, rfl⟩ - · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ - exact ⟨⟨x, ⟨⟨x, hx1⟩, rfl⟩, rfl⟩, (heq x).mpr hx2⟩ -#align number_field.canonical_embedding.integer_lattice.inter_ball_finite NumberField.canonicalEmbedding.integerLattice.inter_ball_finite + · simp [Metric.closedBall_eq_empty.2 hr] + · have heq : ∀ x, canonicalEmbedding K x ∈ Metric.closedBall 0 r ↔ + ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by + intro x; rw [← norm_le_iff, mem_closedBall_zero_iff] + convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) + ext; constructor + · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx⟩ + exact ⟨↑x, ⟨SetLike.coe_mem x, fun φ => (heq x).mp hx φ⟩, rfl⟩ + · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ + exact ⟨⟨x, ⟨⟨x, hx1⟩, rfl⟩, rfl⟩, (heq x).mpr hx2⟩ instance [NumberField K] : Countable (integerLattice K) := by - have : (⋃ n : ℕ, (integerLattice K : Set E) ∩ closedBall 0 n).Countable := + have : (⋃ n : ℕ, (integerLattice K : Set ((K →+* ℂ) → ℂ)) ∩ Metric.closedBall 0 n).Countable := Set.countable_iUnion fun n => (integerLattice.inter_ball_finite K n).countable - refine' (this.mono _).to_subtype + refine (this.mono ?_).to_subtype rintro _ ⟨x, hx, rfl⟩ rw [Set.mem_iUnion] exact ⟨⌈‖canonicalEmbedding K x‖⌉₊, ⟨x, hx, rfl⟩, mem_closedBall_zero_iff.2 (Nat.le_ceil _)⟩ +open Module Fintype FiniteDimensional + +/-- A `ℂ`-basis of `ℂ^n` that is also a `ℤ`-basis of the `integerLattice`. -/ +noncomputable def lattice_basis [NumberField K] : + Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℂ ((K →+* ℂ) → ℂ) := by + classical + -- Let `B` be the canonical basis of `(K →+* ℂ) → ℂ`. We prove that the determinant of + -- the image by `canonicalEmbedding` of the integral basis of `K` is non-zero. This + -- will imply the result. + let B := Pi.basisFun ℂ (K →+* ℂ) + let e : (K →+* ℂ) ≃ Free.ChooseBasisIndex ℤ (𝓞 K) := + equivOfCardEq ((Embeddings.card K ℂ).trans (finrank_eq_card_basis (integralBasis K))) + let M := B.toMatrix (fun i => canonicalEmbedding K (integralBasis K (e i))) + suffices M.det ≠ 0 by + rw [← isUnit_iff_ne_zero, ← Basis.det_apply, ← is_basis_iff_det] at this + refine basisOfLinearIndependentOfCardEqFinrank + ((linearIndependent_equiv e.symm).mpr this.1) ?_ + rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_fintype_fun_eq_card, + Embeddings.card] + -- In order to prove that the determinant is non-zero, we show that it is equal to the + -- square of the discriminant of the integral basis and thus it is not zero + let N := Algebra.embeddingsMatrixReindex ℚ ℂ (fun i => integralBasis K (e i)) + RingHom.equivRatAlgHom + rw [show M = N.transpose by { ext : 2; rfl }] + rw [Matrix.det_transpose, ← @pow_ne_zero_iff ℂ _ _ _ 2 (by norm_num)] + convert (map_ne_zero_iff _ (algebraMap ℚ ℂ).injective).mpr + (Algebra.discr_not_zero_of_basis ℚ (integralBasis K)) + rw [← Algebra.discr_reindex ℚ (integralBasis K) e.symm] + exact (Algebra.discr_eq_det_embeddingsMatrixReindex_pow_two ℚ ℂ + (fun i => integralBasis K (e i)) RingHom.equivRatAlgHom).symm + end NumberField.canonicalEmbedding From f1245a31bf92c1301c44a0f3213f3899b0d9b572 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 27 Jun 2023 14:35:43 +0200 Subject: [PATCH 02/67] Remove unneeded instance --- Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean | 8 -------- 1 file changed, 8 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 965efcb66f4b9..fc02f9a839c76 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -83,14 +83,6 @@ lemma integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ exact ⟨⟨x, ⟨⟨x, hx1⟩, rfl⟩, rfl⟩, (heq x).mpr hx2⟩ -instance [NumberField K] : Countable (integerLattice K) := by - have : (⋃ n : ℕ, (integerLattice K : Set ((K →+* ℂ) → ℂ)) ∩ Metric.closedBall 0 n).Countable := - Set.countable_iUnion fun n => (integerLattice.inter_ball_finite K n).countable - refine (this.mono ?_).to_subtype - rintro _ ⟨x, hx, rfl⟩ - rw [Set.mem_iUnion] - exact ⟨⌈‖canonicalEmbedding K x‖⌉₊, ⟨x, hx, rfl⟩, mem_closedBall_zero_iff.2 (Nat.le_ceil _)⟩ - open Module Fintype FiniteDimensional /-- A `ℂ`-basis of `ℂ^n` that is also a `ℤ`-basis of the `integerLattice`. -/ From a32c2fc39cc4cda4d5b3d7f2a56c5c5e54a33bf8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 28 Jun 2023 15:09:45 +0200 Subject: [PATCH 03/67] Add mem_span_integralBasis --- Mathlib/NumberTheory/NumberField/Basic.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index e4227695d2345..b299d5199459f 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -171,6 +171,21 @@ theorem integralBasis_apply (i : Free.ChooseBasisIndex ℤ (𝓞 K)) : Basis.localizationLocalization_apply ℚ (nonZeroDivisors ℤ) K (RingOfIntegers.basis K) i #align number_field.integral_basis_apply NumberField.integralBasis_apply +theorem mem_span_integralBasis (x : K) : + x ∈ Submodule.span ℤ (Set.range (integralBasis K)) ↔ x ∈ 𝓞 K := by + suffices Submodule.span ℤ (Set.range (integralBasis K)) = + AddSubgroup.toIntSubmodule (𝓞 K).toSubring.toAddSubgroup by rw [this]; rfl + convert congrArg + (Submodule.map (Submodule.subtype (AddSubgroup.toIntSubmodule (𝓞 K).toSubring.toAddSubgroup))) + (Basis.span_eq (RingOfIntegers.basis K)) + · rw [Submodule.map_span] + congr + ext + simp only [Set.mem_range, integralBasis_apply, Submodule.coeSubtype, Set.mem_image, + exists_exists_eq_and] + rfl + · simp only [Submodule.map_top, Submodule.range_subtype] + theorem RingOfIntegers.rank : FiniteDimensional.finrank ℤ (𝓞 K) = FiniteDimensional.finrank ℚ K := IsIntegralClosure.rank ℤ ℚ K (𝓞 K) #align number_field.ring_of_integers.rank NumberField.RingOfIntegers.rank From 3695a30e96357d3f2d74f24e13c5bcf2bd539758 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 28 Jun 2023 15:17:12 +0200 Subject: [PATCH 04/67] Add apply_mem_span_image_iff_mem_span --- Mathlib/LinearAlgebra/Span.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/LinearAlgebra/Span.lean b/Mathlib/LinearAlgebra/Span.lean index 6dea36914cf5a..1d683316c19a7 100644 --- a/Mathlib/LinearAlgebra/Span.lean +++ b/Mathlib/LinearAlgebra/Span.lean @@ -583,6 +583,17 @@ theorem apply_mem_span_image_of_mem_span [RingHomSurjective σ₁₂] (f : M → exact Submodule.mem_map_of_mem h #align submodule.apply_mem_span_image_of_mem_span Submodule.apply_mem_span_image_of_mem_span +theorem mem_span_of_apply_mem_span_image [RingHomSurjective σ₁₂] {f : M →ₛₗ[σ₁₂] M₂} {x : M} + {s : Set M} (hf : Function.Injective f) (h : f x ∈ Submodule.span R₂ (f '' s)) : + x ∈ Submodule.span R s := by + rwa [← Submodule.mem_comap, ← Submodule.map_span, Submodule.comap_map_eq_of_injective hf] at h + +theorem apply_mem_span_image_iff_mem_span [RingHomSurjective σ₁₂] {f : M →ₛₗ[σ₁₂] M₂} {x : M} + {s : Set M} (hf : Function.Injective f) : + f x ∈ Submodule.span R₂ (f '' s) ↔ x ∈ Submodule.span R s := + ⟨fun h => Submodule.mem_span_of_apply_mem_span_image hf h, + fun h => Submodule.apply_mem_span_image_of_mem_span _ h⟩ + @[simp] theorem map_subtype_span_singleton {p : Submodule R M} (x : p) : map p.subtype (R ∙ x) = R ∙ (x : M) := by simp [← span_image] From 45552e82fad84c91a6137f6c3d58a723887900db Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 28 Jun 2023 15:48:40 +0200 Subject: [PATCH 05/67] Add mem_span_latticeBasis --- .../NumberField/CanonicalEmbedding.lean | 32 ++++++++++++++----- 1 file changed, 24 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index fc02f9a839c76..6417e1132727b 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -31,28 +31,28 @@ number field, infinite places variable (K : Type _) [Field K] -open scoped NumberField - namespace NumberField.canonicalEmbedding +open NumberField + /-- The canonical embedding of a number field `K` of degree `n` into `ℂ^n`. -/ def _root_.NumberField.canonicalEmbedding : K →+* ((K →+* ℂ) → ℂ) := Pi.ringHom fun φ => φ -lemma _root_.NumberField.canonicalEmbedding_injective [NumberField K] : +theorem _root_.NumberField.canonicalEmbedding_injective [NumberField K] : Function.Injective (NumberField.canonicalEmbedding K) := RingHom.injective _ variable {K} @[simp] -lemma apply_at (φ : K →+* ℂ) (x : K) : +theorem apply_at (φ : K →+* ℂ) (x : K) : (NumberField.canonicalEmbedding K x) φ = φ x := rfl -lemma nnnorm_eq [NumberField K] (x : K) : +theorem nnnorm_eq [NumberField K] (x : K) : ‖canonicalEmbedding K x‖₊ = Finset.univ.sup (fun φ : K →+* ℂ => ‖φ x‖₊) := by simp_rw [Pi.nnnorm_def, apply_at] -lemma norm_le_iff [NumberField K] (x : K) (r : ℝ) : +theorem norm_le_iff [NumberField K] (x : K) (r : ℝ) : ‖canonicalEmbedding K x‖ ≤ r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by obtain hr | hr := lt_or_le r 0 · obtain ⟨φ : K →+* ℂ⟩ := (inferInstance : Nonempty _) @@ -69,7 +69,7 @@ variable (K) def integerLattice : Subring ((K →+* ℂ) → ℂ) := (RingHom.range (algebraMap (𝓞 K) K)).map (canonicalEmbedding K) -lemma integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : +theorem integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : ((integerLattice K : Set ((K →+* ℂ) → ℂ)) ∩ Metric.closedBall 0 r).Finite := by obtain hr | _ := lt_or_le r 0 · simp [Metric.closedBall_eq_empty.2 hr] @@ -86,7 +86,7 @@ lemma integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : open Module Fintype FiniteDimensional /-- A `ℂ`-basis of `ℂ^n` that is also a `ℤ`-basis of the `integerLattice`. -/ -noncomputable def lattice_basis [NumberField K] : +noncomputable def latticeBasis [NumberField K] : Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℂ ((K →+* ℂ) → ℂ) := by classical -- Let `B` be the canonical basis of `(K →+* ℂ) → ℂ`. We prove that the determinant of @@ -114,4 +114,20 @@ noncomputable def lattice_basis [NumberField K] : exact (Algebra.discr_eq_det_embeddingsMatrixReindex_pow_two ℚ ℂ (fun i => integralBasis K (e i)) RingHom.equivRatAlgHom).symm +@[simp] +theorem latticeBasis_apply [NumberField K] (i : Free.ChooseBasisIndex ℤ (𝓞 K)) : + latticeBasis K i = (canonicalEmbedding K) (integralBasis K i) := by + simp only [latticeBasis, integralBasis_apply, coe_basisOfLinearIndependentOfCardEqFinrank, + Function.comp_apply, Equiv.apply_symm_apply] + +theorem mem_span_latticeBasis [NumberField K] (x : K) : + canonicalEmbedding K x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ 𝓞 K := by + rw [← mem_span_integralBasis, show Set.range (latticeBasis K) = + (canonicalEmbedding K).toIntAlgHom.toLinearMap '' (Set.range (integralBasis K)) by + rw [← Set.range_comp] + exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] + refine Submodule.apply_mem_span_image_iff_mem_span ?_ + change Function.Injective (canonicalEmbedding K) + exact RingHom.injective _ + end NumberField.canonicalEmbedding From b34f46494d14f1ce97b276fbb416c18fb059b420 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Thu, 29 Jun 2023 10:30:39 +0200 Subject: [PATCH 06/67] Small golf --- Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 6417e1132727b..856b80e9c505b 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -126,8 +126,6 @@ theorem mem_span_latticeBasis [NumberField K] (x : K) : (canonicalEmbedding K).toIntAlgHom.toLinearMap '' (Set.range (integralBasis K)) by rw [← Set.range_comp] exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] - refine Submodule.apply_mem_span_image_iff_mem_span ?_ - change Function.Injective (canonicalEmbedding K) - exact RingHom.injective _ + exact Submodule.apply_mem_span_image_iff_mem_span (canonicalEmbedding_injective K) end NumberField.canonicalEmbedding From c6c9ddbab99c6ba0c69ba45248b92110fdceea85 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Thu, 29 Jun 2023 11:42:02 +0200 Subject: [PATCH 07/67] Add conj_apply --- .../NumberField/CanonicalEmbedding.lean | 15 +++++++++++++++ 1 file changed, 15 insertions(+) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 856b80e9c505b..7966e276977f2 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -47,6 +47,21 @@ variable {K} theorem apply_at (φ : K →+* ℂ) (x : K) : (NumberField.canonicalEmbedding K x) φ = φ x := rfl +open scoped ComplexConjugate + +/-- The image of `canonicalEmbedding` lives in the `ℝ`-submodule of the `x ∈ ((K →+* ℂ) → ℂ)` such +that `conj x_φ = x_(conj φ)` for all `∀ φ : K →+* ℂ`. -/ +theorem conj_apply {x : ((K →+* ℂ) → ℂ)} (φ : K →+* ℂ) + (hx : x ∈ Submodule.span ℝ (Set.range (canonicalEmbedding K))) : + conj (x φ) = x (ComplexEmbedding.conjugate φ) := by + refine Submodule.span_induction hx ?_ ?_ (fun _ _ hx hy => ?_) (fun a _ hx => ?_) + · rintro _ ⟨x, rfl⟩ + rw [apply_at, apply_at, ComplexEmbedding.conjugate_coe_eq] + · rw [Pi.zero_apply, Pi.zero_apply, map_zero] + · rw [Pi.add_apply, Pi.add_apply, map_add, hx, hy] + · rw [Pi.smul_apply, Complex.real_smul, map_mul, Complex.conj_ofReal] + exact congrArg ((a : ℂ) * ·) hx + theorem nnnorm_eq [NumberField K] (x : K) : ‖canonicalEmbedding K x‖₊ = Finset.univ.sup (fun φ : K →+* ℂ => ‖φ x‖₊) := by From a8f96774f840f845e7c2904a198f2b6240dc24c9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 1 Jul 2023 16:50:59 +0200 Subject: [PATCH 08/67] 1st commit --- Mathlib/LinearAlgebra/Finsupp.lean | 21 +- Mathlib/NumberTheory/NumberField/Basic.lean | 15 + .../NumberField/CanonicalEmbedding.lean | 471 +++++++++++++----- .../NumberTheory/NumberField/Embeddings.lean | 11 +- 4 files changed, 378 insertions(+), 140 deletions(-) diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index e7fd9ab263998..78d10a2d374aa 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -8,6 +8,7 @@ Authors: Johannes Hölzl ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ +import Mathlib.Data.Set.Countable import Mathlib.Data.Finsupp.Defs import Mathlib.LinearAlgebra.Pi import Mathlib.LinearAlgebra.Span @@ -134,7 +135,7 @@ theorem ker_lsingle (a : α) : ker (lsingle a : M →ₗ[R] α →₀ M) = ⊥ : #align finsupp.ker_lsingle Finsupp.ker_lsingle theorem lsingle_range_le_ker_lapply (s t : Set α) (h : Disjoint s t) : - ⨆ a ∈ s, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M) ≤ + (⨆ a ∈ s, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M)) ≤ ⨅ a ∈ t, ker (lapply a : (α →₀ M) →ₗ[R] M) := by refine' iSup_le fun a₁ => iSup_le fun h₁ => range_le_iff_comap.2 _ simp only [(ker_comp _ _).symm, eq_top_iff, SetLike.le_def, mem_ker, comap_iInf, mem_iInf] @@ -143,12 +144,12 @@ theorem lsingle_range_le_ker_lapply (s t : Set α) (h : Disjoint s t) : exact single_eq_of_ne this #align finsupp.lsingle_range_le_ker_lapply Finsupp.lsingle_range_le_ker_lapply -theorem iInf_ker_lapply_le_bot : ⨅ a, ker (lapply a : (α →₀ M) →ₗ[R] M) ≤ ⊥ := by +theorem iInf_ker_lapply_le_bot : (⨅ a, ker (lapply a : (α →₀ M) →ₗ[R] M)) ≤ ⊥ := by simp only [SetLike.le_def, mem_iInf, mem_ker, mem_bot, lapply_apply] exact fun a h => Finsupp.ext h #align finsupp.infi_ker_lapply_le_bot Finsupp.iInf_ker_lapply_le_bot -theorem iSup_lsingle_range : ⨆ a, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M) = ⊤ := by +theorem iSup_lsingle_range : (⨆ a, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M)) = ⊤ := by refine' eq_top_iff.2 <| SetLike.le_def.2 fun f _ => _ rw [← sum_single f] exact sum_mem fun a _ => Submodule.mem_iSup_of_mem a ⟨_, rfl⟩ @@ -160,8 +161,8 @@ theorem disjoint_lsingle_lsingle (s t : Set α) (hs : Disjoint s t) : -- Porting note: 2 placeholders are added to prevent timeout. refine' (Disjoint.mono - (lsingle_range_le_ker_lapply s sᶜ _) - (lsingle_range_le_ker_lapply t tᶜ _)) + (lsingle_range_le_ker_lapply s (sᶜ) _) + (lsingle_range_le_ker_lapply t (tᶜ) _)) _ · apply disjoint_compl_right · apply disjoint_compl_right @@ -1190,6 +1191,16 @@ theorem mem_span_set {m : M} {s : Set M} : exact Finsupp.mem_span_image_iff_total R (v := _root_.id (α := M)) #align mem_span_set mem_span_set +/-- A `R`-submodule spanned by a finite family of `v` with `R` countable is countable. -/ +instance {ι : Type _} [Countable R] [Fintype ι] (v : ι → M) : + Countable (Submodule.span R (Set.range v)) := by + change Countable (Submodule.span R (Set.range v) : Set M) + rw [Set.countable_coe_iff] + refine Set.Countable.mono ?_ (Set.countable_range (fun f : (ι → R) => ∑ i, (f i) • v i)) + intro _ h + obtain ⟨f, rfl⟩ := Finsupp.mem_span_range_iff_exists_finsupp.mp (SetLike.mem_coe.mp h) + exact ⟨↑f, by rw [Finsupp.sum_fintype] ; exact fun _ => zero_smul _ _⟩ + /-- If `Subsingleton R`, then `M ≃ₗ[R] ι →₀ R` for any type `ι`. -/ @[simps] def Module.subsingletonEquiv (R M ι : Type _) [Semiring R] [Subsingleton R] [AddCommMonoid M] diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index e4227695d2345..03f3233e0c7f8 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -171,6 +171,21 @@ theorem integralBasis_apply (i : Free.ChooseBasisIndex ℤ (𝓞 K)) : Basis.localizationLocalization_apply ℚ (nonZeroDivisors ℤ) K (RingOfIntegers.basis K) i #align number_field.integral_basis_apply NumberField.integralBasis_apply +theorem mem_span_integralBasis {x : K} : + x ∈ Submodule.span ℤ (Set.range (integralBasis K)) ↔ x ∈ 𝓞 K := by + suffices Submodule.span ℤ (Set.range (integralBasis K)) = + AddSubgroup.toIntSubmodule (𝓞 K).toSubring.toAddSubgroup by rw [this]; rfl + convert congrArg + (Submodule.map (Submodule.subtype (AddSubgroup.toIntSubmodule (𝓞 K).toSubring.toAddSubgroup))) + (Basis.span_eq (RingOfIntegers.basis K)) + · rw [Submodule.map_span] + congr + ext + simp only [Set.mem_range, integralBasis_apply, Submodule.coeSubtype, Set.mem_image, + exists_exists_eq_and] + rfl + · simp only [Submodule.map_top, Submodule.range_subtype] + theorem RingOfIntegers.rank : FiniteDimensional.finrank ℤ (𝓞 K) = FiniteDimensional.finrank ℚ K := IsIntegralClosure.rank ℤ ℚ K (𝓞 K) #align number_field.ring_of_integers.rank NumberField.RingOfIntegers.rank diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 8cc8355934c9d..baf1d8825ddfb 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -8,15 +8,19 @@ Authors: Xavier Roblot ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ +import Mathlib.Algebra.Module.Zlattice +import Mathlib.MeasureTheory.Group.GeometryOfNumbers +import Mathlib.MeasureTheory.Measure.Haar.NormedSpace import Mathlib.NumberTheory.NumberField.Embeddings +import Mathlib.RingTheory.Discriminant + /-! # Canonical embedding of a number field -The canonical embedding of a number field `K` of signature `(r₁, r₂)` is the ring homomorphism -`K →+* ℝ^r₁ × ℂ^r₂` that sends `x ∈ K` to `(φ_₁(x),...,φ_r₁(x)) × (ψ_₁(x),..., ψ_r₂(x))` where -`φ_₁,...,φ_r₁` are its real embeddings and `ψ_₁,..., ψ_r₂` are its complex embeddings (up to -complex conjugation). +The canonical embedding of a number field `K` of degree `n` is the ring homomorphism +`K →+* ℂ^n` that sends `x ∈ K` to `(φ_₁(x),...,φ_n(x))` where the `φ_i`'s are the complex +embeddings of `K`. ## Main definitions and results @@ -24,164 +28,369 @@ complex conjugation). image of the ring of integers by the canonical embedding and any ball centered at `0` of finite radius is finite. +* `NumberField.mixedEmbedding`: the ring homomorphism `K →+* ℝ^r₁ × ℂ^r₂` that sends `x ∈ K` to +`(φ_₁(x),...,φ_r₁(x)) × (ψ_₁(x),..., ψ_r₂(x)) ` where `(r₁, r₂)` is the signature of `K`, +`φ_₁,...,φ_r₁` are its real embeddings and `ψ_₁,..., ψ_r₂` are its complex embeddings (up to +complex conjugation). + + ## Tags number field, infinite places -/ - -noncomputable section - -open Function FiniteDimensional Finset Fintype NumberField NumberField.InfinitePlace Metric Module - -open scoped Classical NumberField - variable (K : Type _) [Field K] namespace NumberField.canonicalEmbedding --- The ambient space `ℝ^r₁ × ℂ^r₂` with `(r₁, r₂)` the signature of `K`. -set_option quotPrecheck false in -- Porting note: Added. -scoped[CanonicalEmbedding] - notation "E" => - ({ w : InfinitePlace K // IsReal w } → ℝ) × ({ w : InfinitePlace K // IsComplex w } → ℂ) - -open CanonicalEmbedding - -theorem space_rank [NumberField K] : finrank ℝ E = finrank ℚ K := by - haveI : Module.Free ℝ ℂ := inferInstance - rw [finrank_prod, finrank_pi, finrank_pi_fintype, Complex.finrank_real_complex, Finset.sum_const, - Finset.card_univ, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, - ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, - Nat.add_sub_of_le (Fintype.card_subtype_le _)] -#align number_field.canonical_embedding.space_rank NumberField.canonicalEmbedding.space_rank - -theorem nontrivial_space [NumberField K] : Nontrivial E := by - have : Nonempty <| InfinitePlace K := inferInstance - rcases this with ⟨w⟩ - obtain hw | hw := w.isReal_or_isComplex - · haveI : Nonempty { w : InfinitePlace K // IsReal w } := ⟨⟨w, hw⟩⟩ - exact nontrivial_prod_left - · haveI : Nonempty { w : InfinitePlace K // IsComplex w } := ⟨⟨w, hw⟩⟩ - exact nontrivial_prod_right -#align number_field.canonical_embedding.non_trivial_space NumberField.canonicalEmbedding.nontrivial_space +open NumberField -/-- The canonical embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ -def _root_.NumberField.canonicalEmbedding : K →+* E := - RingHom.prod (Pi.ringHom fun w => w.prop.embedding) (Pi.ringHom fun w => w.val.embedding) -#align number_field.canonical_embedding NumberField.canonicalEmbedding +/-- The canonical embedding of a number field `K` of degree `n` into `ℂ^n`. -/ +def _root_.NumberField.canonicalEmbedding : K →+* ((K →+* ℂ) → ℂ) := Pi.ringHom fun φ => φ theorem _root_.NumberField.canonicalEmbedding_injective [NumberField K] : - Injective (NumberField.canonicalEmbedding K) := - @RingHom.injective _ _ _ _ (nontrivial_space K) _ -#align number_field.canonical_embedding_injective NumberField.canonicalEmbedding_injective - -open NumberField + Function.Injective (NumberField.canonicalEmbedding K) := RingHom.injective _ variable {K} @[simp] -theorem apply_at_real_infinitePlace (w : { w : InfinitePlace K // IsReal w }) (x : K) : - (NumberField.canonicalEmbedding K x).1 w = w.prop.embedding x := by - simp only [canonicalEmbedding, RingHom.prod_apply, Pi.ringHom_apply] -#align number_field.canonical_embedding.apply_at_real_infinite_place NumberField.canonicalEmbedding.apply_at_real_infinitePlace +theorem apply_at (φ : K →+* ℂ) (x : K) : + (NumberField.canonicalEmbedding K x) φ = φ x := rfl -@[simp] -theorem apply_at_complex_infinitePlace (w : { w : InfinitePlace K // IsComplex w }) (x : K) : - (NumberField.canonicalEmbedding K x).2 w = embedding w.val x := by - simp only [canonicalEmbedding, RingHom.prod_apply, Pi.ringHom_apply] -#align number_field.canonical_embedding.apply_at_complex_infinite_place NumberField.canonicalEmbedding.apply_at_complex_infinitePlace +open scoped ComplexConjugate + +/-- The image of `canonicalEmbedding` lives in the `ℝ`-submodule of the `x ∈ ((K →+* ℂ) → ℂ)` such +that `conj x_φ = x_(conj φ)` for all `∀ φ : K →+* ℂ`. -/ +theorem conj_apply {x : ((K →+* ℂ) → ℂ)} (φ : K →+* ℂ) + (hx : x ∈ Submodule.span ℝ (Set.range (canonicalEmbedding K))) : + conj (x φ) = x (ComplexEmbedding.conjugate φ) := by + refine Submodule.span_induction hx ?_ ?_ (fun _ _ hx hy => ?_) (fun a _ hx => ?_) + · rintro _ ⟨x, rfl⟩ + rw [apply_at, apply_at, ComplexEmbedding.conjugate_coe_eq] + · rw [Pi.zero_apply, Pi.zero_apply, map_zero] + · rw [Pi.add_apply, Pi.add_apply, map_add, hx, hy] + · rw [Pi.smul_apply, Complex.real_smul, map_mul, Complex.conj_ofReal] + exact congrArg ((a : ℂ) * ·) hx theorem nnnorm_eq [NumberField K] (x : K) : ‖canonicalEmbedding K x‖₊ = - Finset.univ.sup fun w : InfinitePlace K => ⟨w x, map_nonneg w x⟩ := by - rw [Prod.nnnorm_def', Pi.nnnorm_def, Pi.nnnorm_def] - rw [(_ : Finset.univ = - {w : InfinitePlace K | IsReal w}.toFinset ∪ {w : InfinitePlace K | IsComplex w}.toFinset)] - · rw [Finset.sup_union, sup_eq_max] - refine' congr_arg₂ _ _ _ - · convert - (Finset.univ.sup_map (Function.Embedding.subtype fun w : InfinitePlace K => IsReal w) - fun w => (⟨w x, map_nonneg w x⟩ : NNReal)).symm using 2 - ext w - dsimp - simp only [apply_at_real_infinitePlace, coe_nnnorm, Real.norm_eq_abs, - Function.Embedding.coe_subtype, Subtype.coe_mk, IsReal.abs_embedding_apply] - · convert - (Finset.univ.sup_map (Function.Embedding.subtype fun w : InfinitePlace K => IsComplex w) - fun w => (⟨w x, map_nonneg w x⟩ : NNReal)).symm using 2 - ext w - dsimp - simp only [apply_at_complex_infinitePlace, coe_nnnorm, - Complex.norm_eq_abs, Function.Embedding.coe_subtype, Subtype.coe_mk, abs_embedding] - · ext w - simp_rw [Finset.mem_univ, Finset.mem_union, Set.mem_toFinset, Set.mem_setOf_eq, - w.isReal_or_isComplex] -#align number_field.canonical_embedding.nnnorm_eq NumberField.canonicalEmbedding.nnnorm_eq + Finset.univ.sup (fun φ : K →+* ℂ => ‖φ x‖₊) := by + simp_rw [Pi.nnnorm_def, apply_at] theorem norm_le_iff [NumberField K] (x : K) (r : ℝ) : - ‖canonicalEmbedding K x‖ ≤ r ↔ ∀ w : InfinitePlace K, w x ≤ r := by + ‖canonicalEmbedding K x‖ ≤ r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by obtain hr | hr := lt_or_le r 0 - · have : Nonempty <| InfinitePlace K := inferInstance - rcases this with ⟨w⟩ - exact iff_of_false - (hr.trans_le <| norm_nonneg _).not_le fun h => hr.not_le <| (map_nonneg w _).trans <| h _ + · obtain ⟨φ : K →+* ℂ⟩ := (inferInstance : Nonempty _) + refine iff_of_false ?_ ?_ + exact (hr.trans_le (norm_nonneg _)).not_le + exact fun h => hr.not_le (le_trans (norm_nonneg _) (h φ)) · lift r to NNReal using hr simp_rw [← coe_nnnorm, nnnorm_eq, NNReal.coe_le_coe, Finset.sup_le_iff, Finset.mem_univ, - forall_true_left, ← NNReal.coe_le_coe, NNReal.toReal] -#align number_field.canonical_embedding.norm_le_iff NumberField.canonicalEmbedding.norm_le_iff + forall_true_left] variable (K) -/-- The image of `𝓞 K` as a subring of `ℝ^r₁ × ℂ^r₂`. -/ -def integerLattice : Subring E := +/-- The image of `𝓞 K` as a subring of `ℂ^n`. -/ +def integerLattice : Subring ((K →+* ℂ) → ℂ) := (RingHom.range (algebraMap (𝓞 K) K)).map (canonicalEmbedding K) -#align number_field.canonical_embedding.integer_lattice NumberField.canonicalEmbedding.integerLattice - --- Porting note: See https://github.com/leanprover-community/mathlib4/issues/5028 -set_option maxHeartbeats 400000 in -set_option synthInstance.maxHeartbeats 50000 in -/-- The linear equiv between `𝓞 K` and the integer lattice. -/ -def equivIntegerLattice [NumberField K] : 𝓞 K ≃ₗ[ℤ] integerLattice K := - LinearEquiv.ofBijective - { toFun := fun x => (by - refine ⟨canonicalEmbedding K (algebraMap (𝓞 K) K x), ⟨algebraMap (𝓞 K) K x, ⟨?_, rfl⟩⟩⟩ - simp only [Subsemiring.coe_carrier_toSubmonoid, Subring.coe_toSubsemiring, - RingHom.coe_range, Set.mem_range, exists_apply_eq_apply] ) - map_add' := fun x y => (by - apply Subtype.eq - simp [map_add] ) - map_smul' := fun c x => (by - simp only [RingHom.id_apply, zsmul_eq_mul, RingHom.map_mul, map_intCast] - rfl ) } - (by - refine ⟨fun _ _ h => ?_, fun ⟨_, ⟨_, ⟨⟨a, rfl⟩, rfl⟩⟩⟩ => ⟨a, rfl⟩⟩ - dsimp only at h - rw [LinearMap.coe_mk, Subtype.mk_eq_mk] at h - exact IsFractionRing.injective (𝓞 K) K (canonicalEmbedding_injective K h)) -#align number_field.canonical_embedding.equiv_integer_lattice NumberField.canonicalEmbedding.equivIntegerLattice theorem integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : - ((integerLattice K : Set E) ∩ closedBall 0 r).Finite := by + ((integerLattice K : Set ((K →+* ℂ) → ℂ)) ∩ Metric.closedBall 0 r).Finite := by obtain hr | _ := lt_or_le r 0 - · simp [closedBall_eq_empty.2 hr] - have heq : ∀ x, canonicalEmbedding K x ∈ closedBall (0 : E) r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by - simp only [← place_apply, ← InfinitePlace.coe_mk, mem_closedBall_zero_iff, norm_le_iff] - exact fun x => le_iff_le x r - convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) - ext; constructor - · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx2⟩ - exact ⟨x, ⟨SetLike.coe_mem x, (heq x).mp hx2⟩, rfl⟩ - · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ - exact ⟨⟨x, ⟨⟨x, hx1⟩, rfl⟩, rfl⟩, (heq x).mpr hx2⟩ -#align number_field.canonical_embedding.integer_lattice.inter_ball_finite NumberField.canonicalEmbedding.integerLattice.inter_ball_finite - -instance [NumberField K] : Countable (integerLattice K) := by - have : (⋃ n : ℕ, (integerLattice K : Set E) ∩ closedBall 0 n).Countable := - Set.countable_iUnion fun n => (integerLattice.inter_ball_finite K n).countable - refine' (this.mono _).to_subtype - rintro _ ⟨x, hx, rfl⟩ - rw [Set.mem_iUnion] - exact ⟨⌈‖canonicalEmbedding K x‖⌉₊, ⟨x, hx, rfl⟩, mem_closedBall_zero_iff.2 (Nat.le_ceil _)⟩ + · simp [Metric.closedBall_eq_empty.2 hr] + · have heq : ∀ x, canonicalEmbedding K x ∈ Metric.closedBall 0 r ↔ + ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by + intro x ; rw [← norm_le_iff, mem_closedBall_zero_iff] + convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) + ext ; constructor + · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx⟩ + exact ⟨↑x, ⟨SetLike.coe_mem x, fun φ => (heq x).mp hx φ⟩, rfl⟩ + · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ + exact ⟨⟨x, ⟨⟨x, hx1⟩, rfl⟩, rfl⟩, (heq x).mpr hx2⟩ + +open Module Fintype FiniteDimensional + +/-- A `ℂ`-basis of `ℂ^n` that is also a `ℤ`-basis of the `integerLattice`. -/ +noncomputable def latticeBasis [NumberField K] : + Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℂ ((K →+* ℂ) → ℂ) := by + classical + -- Let `B` be the canonical basis of `(K →+* ℂ) → ℂ`. We prove that the determinant of + -- the image by `canonicalEmbedding` of the integral basis of `K` is non-zero. This + -- will imply the result. + let B := Pi.basisFun ℂ (K →+* ℂ) + let e : (K →+* ℂ) ≃ Free.ChooseBasisIndex ℤ (𝓞 K) := + equivOfCardEq ((Embeddings.card K ℂ).trans (finrank_eq_card_basis (integralBasis K))) + let M := B.toMatrix (fun i => canonicalEmbedding K (integralBasis K (e i))) + suffices M.det ≠ 0 by + rw [← isUnit_iff_ne_zero, ← Basis.det_apply, ← is_basis_iff_det] at this + refine basisOfLinearIndependentOfCardEqFinrank + ((linearIndependent_equiv e.symm).mpr this.1) ?_ + rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_fintype_fun_eq_card, + Embeddings.card] + -- In order to prove that the determinant is non-zero, we show that it is equal to the + -- square of the discriminant of the integral basis and thus it is not zero + let N := Algebra.embeddingsMatrixReindex ℚ ℂ (fun i => integralBasis K (e i)) + RingHom.equivRatAlgHom + rw [show M = N.transpose by { ext : 2 ; rfl }] + rw [Matrix.det_transpose, ← @pow_ne_zero_iff ℂ _ _ _ 2 (by norm_num)] + convert (map_ne_zero_iff _ (algebraMap ℚ ℂ).injective).mpr + (Algebra.discr_not_zero_of_basis ℚ (integralBasis K)) + rw [← Algebra.discr_reindex ℚ (integralBasis K) e.symm] + exact (Algebra.discr_eq_det_embeddingsMatrixReindex_pow_two ℚ ℂ + (fun i => integralBasis K (e i)) RingHom.equivRatAlgHom).symm + +@[simp] +theorem latticeBasis_apply [NumberField K] (i : Free.ChooseBasisIndex ℤ (𝓞 K)) : + latticeBasis K i = (canonicalEmbedding K) (integralBasis K i) := by + simp only [latticeBasis, integralBasis_apply, coe_basisOfLinearIndependentOfCardEqFinrank, + Function.comp_apply, Equiv.apply_symm_apply] + +theorem mem_span_latticeBasis [NumberField K] (x : (K →+* ℂ) → ℂ) : + x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ canonicalEmbedding K '' (𝓞 K) := by + rw [show Set.range (latticeBasis K) = + (canonicalEmbedding K).toIntAlgHom.toLinearMap '' (Set.range (integralBasis K)) by + rw [← Set.range_comp] ; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] + rw [← Submodule.map_span, ← SetLike.mem_coe, Submodule.map_coe] + rw [show (Submodule.span ℤ (Set.range (integralBasis K)) : Set K) = 𝓞 K by + ext ; exact mem_span_integralBasis K] + rfl end NumberField.canonicalEmbedding + +namespace NumberField.mixedEmbedding + +open NumberField NumberField.InfinitePlace NumberField.ComplexEmbedding + +/-- The ambient space `ℝ^r₁ × ℂ^r₂` with `(r₁, r₂)` the signature of `K`. -/ +local notation "E" K => + ({ w : InfinitePlace K // IsReal w } → ℝ) × ({ w : InfinitePlace K // IsComplex w } → ℂ) + +/-- The canonical embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ +noncomputable def _root_.NumberField.mixedEmbedding : K →+* (E K) := + RingHom.prod (Pi.ringHom fun w => w.prop.embedding) (Pi.ringHom fun w => w.val.embedding) + +instance [NumberField K] : Nontrivial (E K) := by + obtain ⟨w⟩ := (inferInstance : Nonempty (InfinitePlace K)) + obtain hw | hw := w.isReal_or_isComplex + · have : Nonempty {w : InfinitePlace K // IsReal w} := ⟨⟨w, hw⟩⟩ + exact nontrivial_prod_left + · have : Nonempty {w : InfinitePlace K // IsComplex w} := ⟨⟨w, hw⟩⟩ + exact nontrivial_prod_right + +theorem _root_.NumberField.mixedEmbedding_injective [NumberField K] : + Function.Injective (NumberField.mixedEmbedding K) := by + exact RingHom.injective _ + +section comm_map + +/-- The linear map that makes `canonicalEmbedding` and `mixedEmbedding` commute, see +`comm_map_canonical_eq_mixed`. -/ +noncomputable def comm_map : ((K →+* ℂ) → ℂ) →ₗ[ℝ] (E K) := +{ toFun := fun x => ⟨fun w => (x w.val.embedding).re, fun w => x w.val.embedding⟩ + map_add' := by + simp only [Pi.add_apply, Complex.add_re, Prod.mk_add_mk, Prod.mk.injEq] + exact fun _ _ => ⟨rfl, rfl⟩ + map_smul' := by + simp only [Pi.smul_apply, Complex.real_smul, Complex.mul_re, Complex.ofReal_re, + Complex.ofReal_im, zero_mul, sub_zero, RingHom.id_apply, Prod.smul_mk, Prod.mk.injEq] + exact fun _ _ => ⟨rfl, rfl⟩ } + +theorem comm_map_apply_of_isReal (x : (K →+* ℂ) → ℂ) {w : InfinitePlace K} (hw : IsReal w) : + (comm_map K x).1 ⟨w, hw⟩ = (x w.embedding).re := rfl + +theorem comm_map_apply_of_isComplex (x : (K →+* ℂ) → ℂ) {w : InfinitePlace K} (hw : IsComplex w) : + (comm_map K x).2 ⟨w, hw⟩ = x w.embedding := rfl + +@[simp] +theorem comm_map_canonical_eq_mixed (x : K) : + comm_map K (canonicalEmbedding K x) = mixedEmbedding K x := by + simp only [canonicalEmbedding, comm_map, LinearMap.coe_mk, AddHom.coe_mk, Pi.ringHom_apply, + mixedEmbedding, RingHom.prod_apply, Prod.mk.injEq] + exact ⟨rfl, rfl⟩ + +/-- This is technical result to ensure that the image of the `ℂ`-basis of `ℂ^n`, defined in +`canonicalEmbedding.latticeBasis` is a `ℝ`-basis of `ℝ^r₁ × ℂ^r₂`, +see `mixedEmbedding.latticeBasis`. -/ +theorem disjoint_span_comm_map_ker [NumberField K]: + Disjoint (Submodule.span ℝ (Set.range (canonicalEmbedding.latticeBasis K))) + (LinearMap.ker (comm_map K)) := by + refine LinearMap.disjoint_ker.mpr (fun x h_mem h_zero => ?_) + replace h_mem : x ∈ Submodule.span ℝ (Set.range (canonicalEmbedding K)) := by + refine (Submodule.span_mono ?_) h_mem + rintro _ ⟨i, rfl⟩ + exact ⟨integralBasis K i, (canonicalEmbedding.latticeBasis_apply K i).symm⟩ + ext1 φ + rw [Pi.zero_apply] + by_cases hφ : IsReal φ + · rw [show x φ = (x φ).re by + rw [eq_comm, ← Complex.conj_eq_iff_re, canonicalEmbedding.conj_apply _ h_mem, + ComplexEmbedding.isReal_iff.mp hφ], ← Complex.ofReal_zero] + congr + rw [← ComplexEmbeddings.IsReal.embedding_mk hφ, ← comm_map_apply_of_isReal K x ⟨φ, hφ, rfl⟩] + exact congrFun (congrArg (fun x => x.1) h_zero) ⟨InfinitePlace.mk φ, _⟩ + · have := congrFun (congrArg (fun x => x.2) h_zero) ⟨InfinitePlace.mk φ, ⟨φ, hφ, rfl⟩⟩ + cases ComplexEmbeddings.embedding_mk φ with + | inl h => rwa [← h, ← comm_map_apply_of_isComplex K x ⟨φ, hφ, rfl⟩] + | inr h => + apply RingHom.injective (starRingEnd ℂ) + rwa [canonicalEmbedding.conj_apply _ h_mem, ← h, map_zero, + ← comm_map_apply_of_isComplex K x ⟨φ, hφ, rfl⟩] + +end comm_map + +section integerLattice + +open Module FiniteDimensional + +/-- A `ℝ`-basis of `ℝ^r₁ × ℂ^r₂` that is also a `ℤ`-basis of the image of `𝓞 K`. -/ +noncomputable def latticeBasis [NumberField K] : + Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℝ (E K) := by + classical + -- We construct an `ℝ`-linear independent family from the image of + -- `canonicalEmbedding.lattice_basis` by `comm_map` + have := LinearIndependent.map (LinearIndependent.restrict_scalars + (by { simpa only [Complex.real_smul, mul_one] using Complex.ofReal_injective }) + (canonicalEmbedding.latticeBasis K).linearIndependent) + (disjoint_span_comm_map_ker K) + -- and it's a basis since it has the right cardinality + refine basisOfLinearIndependentOfCardEqFinrank this ?_ + rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_prod, finrank_pi, + finrank_pi_fintype, Complex.finrank_real_complex, Finset.sum_const, Finset.card_univ, + ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, ← card_complex_embeddings, + ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, + Nat.add_sub_of_le (Fintype.card_subtype_le _)] + +@[simp] +theorem latticeBasis_apply [NumberField K] (i : Free.ChooseBasisIndex ℤ (𝓞 K)) : + latticeBasis K i = (mixedEmbedding K) (integralBasis K i) := by + simp only [latticeBasis, coe_basisOfLinearIndependentOfCardEqFinrank, Function.comp_apply, + canonicalEmbedding.latticeBasis_apply, integralBasis_apply, comm_map_canonical_eq_mixed] + +theorem mem_span_latticeBasis [NumberField K] (x : (E K)) : + x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ mixedEmbedding K '' (𝓞 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))] + rw [← Submodule.map_span, ← SetLike.mem_coe, Submodule.map_coe] + rw [show (Submodule.span ℤ (Set.range (integralBasis K)) : Set K) = 𝓞 K by + ext ; exact mem_span_integralBasis K] + rfl + +end integerLattice + +section convex_body + +open Metric ENNReal NNReal + +variable (f : InfinitePlace K → ℝ≥0) + +/-- The convex body defined by `f`: the set of points `x : E` such that `‖x w‖ < f w` for all +infinite place `w`. -/ +def convex_body : Set (E K) := + (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsReal w } => ball 0 (f w))) ×ˢ + (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsComplex w } => ball 0 (f w))) + +theorem convex_body_mem {x : K} : + mixedEmbedding K x ∈ (convex_body K f) ↔ ∀ w : InfinitePlace K, w x < f w := by + simp only [mixedEmbedding, RingHom.prod_apply, convex_body, Set.mem_prod, Set.mem_pi, + Set.mem_univ, Pi.ringHom_apply, mem_ball, dist_zero_right, Real.norm_eq_abs, + IsReal.abs_embedding_apply, forall_true_left, Subtype.forall, Complex.norm_eq_abs, + abs_embedding, ← ball_or_left, ← not_isReal_iff_isComplex, em, forall_true_left] + +theorem convex_body_symmetric (x : E K) (hx : x ∈ (convex_body K f)) : + -x ∈ (convex_body K f) := by + simp only [convex_body, Set.mem_prod, Prod.fst_neg, Set.mem_pi, Set.mem_univ, Pi.neg_apply, + mem_ball, dist_zero_right, norm_neg, Real.norm_eq_abs, forall_true_left, Subtype.forall, + Prod.snd_neg, Complex.norm_eq_abs, hx] at hx ⊢ + exact hx + +theorem convex_body_convex : + Convex ℝ (convex_body K f) := + Convex.prod (convex_pi (fun _ _ => convex_ball _ _)) (convex_pi (fun _ _ => convex_ball _ _)) + +open Classical Fintype MeasureTheory MeasureTheory.Measure BigOperators + +-- See: https://github.com/leanprover/lean4/issues/2220 +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) + +variable [NumberField K] + +/-- The fudge factor that appears in the formula for the volume of `convex_body`. -/ +noncomputable def constant_factor : ℝ≥0∞ := + (2 : ℝ≥0∞) ^ card {w : InfinitePlace K // IsReal w} * + volume (ball (0 : ℂ) 1) ^ card {w : InfinitePlace K // IsComplex w} + +instance : IsAddHaarMeasure (@volume ℂ Complex.measureSpace) := MapLinearEquiv.isAddHaarMeasure _ _ + +theorem constant_factor_pos : 0 < (constant_factor K) := by + refine mul_pos (NeZero.ne _) ?_ + exact ENNReal.pow_ne_zero (ne_of_gt (measure_ball_pos _ _ (by norm_num))) _ + +theorem constant_factor_lt_top : (constant_factor K) < ⊤ := by + refine mul_lt_top ?_ ?_ + · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) + · exact ne_of_lt (pow_lt_top measure_ball_lt_top _) + +set_option maxHeartbeats 400000 in +theorem convex_body_volume : + volume (convex_body K f) = (constant_factor K) * + ∏ w : InfinitePlace K, ite (IsReal w) ↑(f w) ↑(f w ^ 2) := by + rw [volume_eq_prod, convex_body, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi] + conv_lhs => + congr ; congr ; next => skip + ext + rw [Real.volume_ball, ofReal_mul (by norm_num), ofReal_coe_nnreal, mul_comm] + conv_lhs => + congr ; next => skip + congr ; next => skip + ext i + rw [add_haar_ball _ _ (by exact (f i).prop), Complex.finrank_real_complex, ← NNReal.coe_pow, + ofReal_coe_nnreal, mul_comm] + rw [Finset.prod_mul_distrib, Finset.prod_mul_distrib, Finset.prod_const, Finset.prod_const, + Finset.card_univ, Finset.card_univ, mul_assoc, mul_comm, ← mul_assoc, mul_assoc, ofReal_ofNat, + ← constant_factor, Finset.prod_ite] + simp_rw [show (fun w : InfinitePlace K ↦ ¬IsReal w) = (fun w ↦ IsComplex w) + by funext ; rw [not_isReal_iff_isComplex]] + congr 1 ; rw [mul_comm] ; congr 1 + all_goals + · rw [← Finset.prod_subtype_eq_prod_filter] + congr ; ext + exact ⟨fun _ => Finset.mem_subtype.mpr (Finset.mem_univ _), fun _ => Finset.mem_univ _⟩ + +end convex_body + +section minkowski + +open MeasureTheory MeasureTheory.Measure Classical ENNReal FiniteDimensional + +variable [NumberField K] + +/-- The bound that appears in Minkowski theorem, see +`MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure`. -/ +noncomputable def minkowski_bound : ℝ≥0∞ := + volume (Zspan.fundamentalDomain (latticeBasis K)) * 2 ^ (finrank ℝ (E K)) + +theorem minkowski_bound_lt_top : minkowski_bound K < ⊤ := by + refine mul_lt_top ?_ ?_ + · exact ne_of_lt (Zspan.fundamentalDomain_bounded (latticeBasis K)).measure_lt_top + · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) + +theorem exists_ne_zero_mem_ring_of_integers_lt (h : minkowski_bound K < volume (convex_body K f)) : + ∃ (a : 𝓞 K), a ≠ 0 ∧ ∀ w : InfinitePlace K, w a < f w := by + have : @IsAddHaarMeasure (E K) _ _ _ volume := prod.instIsAddHaarMeasure volume volume + have h_fund := Zspan.isAddFundamentalDomain (latticeBasis K) volume + have : Countable (Submodule.span ℤ (Set.range (latticeBasis K)) : Set (E K)) := inferInstance + have : Countable (Submodule.span ℤ (Set.range (latticeBasis K))).toAddSubgroup := by + change Countable (Submodule.span ℤ (Set.range (latticeBasis K)): Set (E K)) + infer_instance + obtain ⟨⟨x, hx⟩, h_nzr, h_mem⟩ := exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure h_fund h + (convex_body_symmetric K f) (convex_body_convex K f) + rw [Submodule.mem_toAddSubgroup, mem_span_latticeBasis] at hx + obtain ⟨a, ha, rfl⟩ := hx + refine ⟨⟨a, ha⟩, ?_, (convex_body_mem K f).mp h_mem⟩ + rw [ne_eq, AddSubgroup.mk_eq_zero_iff, map_eq_zero, ← ne_eq] at h_nzr + exact Subtype.ne_of_val_ne h_nzr + +end minkowski + +end NumberField.mixedEmbedding diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 2807358e515ea..e8d8422cb9051 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -358,6 +358,11 @@ theorem _root_.NumberField.ComplexEmbeddings.IsReal.embedding_mk {φ : K →+* rwa [ComplexEmbedding.isReal_iff.mp h, or_self_iff, eq_comm] at this #align number_field.complex_embeddings.is_real.embedding_mk NumberField.ComplexEmbeddings.IsReal.embedding_mk +@[simp] +theorem _root_.NumberField.ComplexEmbeddings.embedding_mk (φ : K →+* ℂ) : + embedding (mk φ) = φ ∨ embedding (mk φ) = ComplexEmbedding.conjugate φ := by + rw [@eq_comm _ _ φ, @eq_comm _ _ (ComplexEmbedding.conjugate φ), ← mk_eq_iff, mk_embedding] + theorem isReal_iff {w : InfinitePlace K} : IsReal w ↔ ComplexEmbedding.IsReal (embedding w) := by constructor · rintro ⟨φ, ⟨hφ, rfl⟩⟩ @@ -428,10 +433,8 @@ noncomputable def mkComplex : theorem mkComplex_embedding (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) : (mkComplex K φ : InfinitePlace K).embedding = φ ∨ - (mkComplex K φ : InfinitePlace K).embedding = ComplexEmbedding.conjugate φ := by - rw [@eq_comm _ _ φ.val, @eq_comm _ _ (ComplexEmbedding.conjugate φ.val), ← mk_eq_iff, - mk_embedding] - rfl + (mkComplex K φ : InfinitePlace K).embedding = ComplexEmbedding.conjugate φ := + NumberField.ComplexEmbeddings.embedding_mk φ.1 #align number_field.infinite_place.mk_complex_embedding NumberField.InfinitePlace.mkComplex_embedding @[simp] From 0695983d050df90a091279fe1956259ab122eb7b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 1 Jul 2023 16:53:48 +0200 Subject: [PATCH 09/67] Fix line --- Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index baf1d8825ddfb..5d8cd79821a52 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -383,8 +383,8 @@ theorem exists_ne_zero_mem_ring_of_integers_lt (h : minkowski_bound K < volume ( have : Countable (Submodule.span ℤ (Set.range (latticeBasis K))).toAddSubgroup := by change Countable (Submodule.span ℤ (Set.range (latticeBasis K)): Set (E K)) infer_instance - obtain ⟨⟨x, hx⟩, h_nzr, h_mem⟩ := exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure h_fund h - (convex_body_symmetric K f) (convex_body_convex K f) + obtain ⟨⟨x, hx⟩, h_nzr, h_mem⟩ := exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure + h_fund h (convex_body_symmetric K f) (convex_body_convex K f) rw [Submodule.mem_toAddSubgroup, mem_span_latticeBasis] at hx obtain ⟨a, ha, rfl⟩ := hx refine ⟨⟨a, ha⟩, ?_, (convex_body_mem K f).mp h_mem⟩ From c92dd5447ec3347200021cf3533807a9cc8eff5d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 1 Jul 2023 16:56:08 +0200 Subject: [PATCH 10/67] merge --- Mathlib/LinearAlgebra/Finsupp.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index 78d10a2d374aa..27ad90d29110a 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -135,7 +135,7 @@ theorem ker_lsingle (a : α) : ker (lsingle a : M →ₗ[R] α →₀ M) = ⊥ : #align finsupp.ker_lsingle Finsupp.ker_lsingle theorem lsingle_range_le_ker_lapply (s t : Set α) (h : Disjoint s t) : - (⨆ a ∈ s, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M)) ≤ + ⨆ a ∈ s, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M) ≤ ⨅ a ∈ t, ker (lapply a : (α →₀ M) →ₗ[R] M) := by refine' iSup_le fun a₁ => iSup_le fun h₁ => range_le_iff_comap.2 _ simp only [(ker_comp _ _).symm, eq_top_iff, SetLike.le_def, mem_ker, comap_iInf, mem_iInf] @@ -144,12 +144,12 @@ theorem lsingle_range_le_ker_lapply (s t : Set α) (h : Disjoint s t) : exact single_eq_of_ne this #align finsupp.lsingle_range_le_ker_lapply Finsupp.lsingle_range_le_ker_lapply -theorem iInf_ker_lapply_le_bot : (⨅ a, ker (lapply a : (α →₀ M) →ₗ[R] M)) ≤ ⊥ := by +theorem iInf_ker_lapply_le_bot : ⨅ a, ker (lapply a : (α →₀ M) →ₗ[R] M) ≤ ⊥ := by simp only [SetLike.le_def, mem_iInf, mem_ker, mem_bot, lapply_apply] exact fun a h => Finsupp.ext h #align finsupp.infi_ker_lapply_le_bot Finsupp.iInf_ker_lapply_le_bot -theorem iSup_lsingle_range : (⨆ a, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M)) = ⊤ := by +theorem iSup_lsingle_range : ⨆ a, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M) = ⊤ := by refine' eq_top_iff.2 <| SetLike.le_def.2 fun f _ => _ rw [← sum_single f] exact sum_mem fun a _ => Submodule.mem_iSup_of_mem a ⟨_, rfl⟩ @@ -161,8 +161,8 @@ theorem disjoint_lsingle_lsingle (s t : Set α) (hs : Disjoint s t) : -- Porting note: 2 placeholders are added to prevent timeout. refine' (Disjoint.mono - (lsingle_range_le_ker_lapply s (sᶜ) _) - (lsingle_range_le_ker_lapply t (tᶜ) _)) + (lsingle_range_le_ker_lapply s sᶜ _) + (lsingle_range_le_ker_lapply t tᶜ _)) _ · apply disjoint_compl_right · apply disjoint_compl_right From 91e8c7855242bbb7be261126001aa5ae943c31d3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 1 Jul 2023 17:07:42 +0200 Subject: [PATCH 11/67] merge --- Mathlib/LinearAlgebra/Span.lean | 11 ----------- Mathlib/NumberTheory/NumberField/Basic.lean | 4 ++-- Mathlib/NumberTheory/NumberField/Embeddings.lean | 5 +++++ 3 files changed, 7 insertions(+), 13 deletions(-) diff --git a/Mathlib/LinearAlgebra/Span.lean b/Mathlib/LinearAlgebra/Span.lean index f7eb88be796e5..8cceaf425a174 100644 --- a/Mathlib/LinearAlgebra/Span.lean +++ b/Mathlib/LinearAlgebra/Span.lean @@ -583,17 +583,6 @@ theorem apply_mem_span_image_of_mem_span [RingHomSurjective σ₁₂] (f : M → exact Submodule.mem_map_of_mem h #align submodule.apply_mem_span_image_of_mem_span Submodule.apply_mem_span_image_of_mem_span -theorem mem_span_of_apply_mem_span_image [RingHomSurjective σ₁₂] {f : M →ₛₗ[σ₁₂] M₂} {x : M} - {s : Set M} (hf : Function.Injective f) (h : f x ∈ Submodule.span R₂ (f '' s)) : - x ∈ Submodule.span R s := by - rwa [← Submodule.mem_comap, ← Submodule.map_span, Submodule.comap_map_eq_of_injective hf] at h - -theorem apply_mem_span_image_iff_mem_span [RingHomSurjective σ₁₂] {f : M →ₛₗ[σ₁₂] M₂} {x : M} - {s : Set M} (hf : Function.Injective f) : - f x ∈ Submodule.span R₂ (f '' s) ↔ x ∈ Submodule.span R s := - ⟨fun h => Submodule.mem_span_of_apply_mem_span_image hf h, - fun h => Submodule.apply_mem_span_image_of_mem_span _ h⟩ - @[simp] theorem map_subtype_span_singleton {p : Submodule R M} (x : p) : map p.subtype (R ∙ x) = R ∙ (x : M) := by simp [← span_image] diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index b299d5199459f..03f3233e0c7f8 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -171,7 +171,7 @@ theorem integralBasis_apply (i : Free.ChooseBasisIndex ℤ (𝓞 K)) : Basis.localizationLocalization_apply ℚ (nonZeroDivisors ℤ) K (RingOfIntegers.basis K) i #align number_field.integral_basis_apply NumberField.integralBasis_apply -theorem mem_span_integralBasis (x : K) : +theorem mem_span_integralBasis {x : K} : x ∈ Submodule.span ℤ (Set.range (integralBasis K)) ↔ x ∈ 𝓞 K := by suffices Submodule.span ℤ (Set.range (integralBasis K)) = AddSubgroup.toIntSubmodule (𝓞 K).toSubring.toAddSubgroup by rw [this]; rfl @@ -185,7 +185,7 @@ theorem mem_span_integralBasis (x : K) : exists_exists_eq_and] rfl · simp only [Submodule.map_top, Submodule.range_subtype] - + theorem RingOfIntegers.rank : FiniteDimensional.finrank ℤ (𝓞 K) = FiniteDimensional.finrank ℚ K := IsIntegralClosure.rank ℤ ℚ K (𝓞 K) #align number_field.ring_of_integers.rank NumberField.RingOfIntegers.rank diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 2807358e515ea..5f4adecca334c 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -358,6 +358,11 @@ theorem _root_.NumberField.ComplexEmbeddings.IsReal.embedding_mk {φ : K →+* rwa [ComplexEmbedding.isReal_iff.mp h, or_self_iff, eq_comm] at this #align number_field.complex_embeddings.is_real.embedding_mk NumberField.ComplexEmbeddings.IsReal.embedding_mk +@[simp] +theorem _root_.NumberField.ComplexEmbeddings.embedding_mk (φ : K →+* ℂ) : + embedding (mk φ) = φ ∨ embedding (mk φ) = ComplexEmbedding.conjugate φ := by + rw [@eq_comm _ _ φ, @eq_comm _ _ (ComplexEmbedding.conjugate φ), ← mk_eq_iff, mk_embedding] + theorem isReal_iff {w : InfinitePlace K} : IsReal w ↔ ComplexEmbedding.IsReal (embedding w) := by constructor · rintro ⟨φ, ⟨hφ, rfl⟩⟩ From 047b8c9fa949fc9777013326212747fb996a62ba Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 1 Jul 2023 17:20:13 +0200 Subject: [PATCH 12/67] Fix backport --- .../NumberField/CanonicalEmbedding.lean | 23 +++++++++++-------- 1 file changed, 13 insertions(+), 10 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 7966e276977f2..f08d723347b85 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -11,6 +11,7 @@ Authors: Xavier Roblot import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.RingTheory.Discriminant + /-! # Canonical embedding of a number field @@ -50,7 +51,7 @@ theorem apply_at (φ : K →+* ℂ) (x : K) : open scoped ComplexConjugate /-- The image of `canonicalEmbedding` lives in the `ℝ`-submodule of the `x ∈ ((K →+* ℂ) → ℂ)` such -that `conj x_φ = x_(conj φ)` for all `∀ φ : K →+* ℂ`. -/ +that `conj x_φ = x_(conj φ)` for all `∀ φ : K →+* ℂ`. -/ theorem conj_apply {x : ((K →+* ℂ) → ℂ)} (φ : K →+* ℂ) (hx : x ∈ Submodule.span ℝ (Set.range (canonicalEmbedding K))) : conj (x φ) = x (ComplexEmbedding.conjugate φ) := by @@ -90,9 +91,9 @@ theorem integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : · simp [Metric.closedBall_eq_empty.2 hr] · have heq : ∀ x, canonicalEmbedding K x ∈ Metric.closedBall 0 r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by - intro x; rw [← norm_le_iff, mem_closedBall_zero_iff] + intro x ; rw [← norm_le_iff, mem_closedBall_zero_iff] convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) - ext; constructor + ext ; constructor · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx⟩ exact ⟨↑x, ⟨SetLike.coe_mem x, fun φ => (heq x).mp hx φ⟩, rfl⟩ · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ @@ -121,7 +122,7 @@ noncomputable def latticeBasis [NumberField K] : -- square of the discriminant of the integral basis and thus it is not zero let N := Algebra.embeddingsMatrixReindex ℚ ℂ (fun i => integralBasis K (e i)) RingHom.equivRatAlgHom - rw [show M = N.transpose by { ext : 2; rfl }] + rw [show M = N.transpose by { ext : 2 ; rfl }] rw [Matrix.det_transpose, ← @pow_ne_zero_iff ℂ _ _ _ 2 (by norm_num)] convert (map_ne_zero_iff _ (algebraMap ℚ ℂ).injective).mpr (Algebra.discr_not_zero_of_basis ℚ (integralBasis K)) @@ -135,12 +136,14 @@ theorem latticeBasis_apply [NumberField K] (i : Free.ChooseBasisIndex ℤ (𝓞 simp only [latticeBasis, integralBasis_apply, coe_basisOfLinearIndependentOfCardEqFinrank, Function.comp_apply, Equiv.apply_symm_apply] -theorem mem_span_latticeBasis [NumberField K] (x : K) : - canonicalEmbedding K x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ 𝓞 K := by - rw [← mem_span_integralBasis, show Set.range (latticeBasis K) = +theorem mem_span_latticeBasis [NumberField K] (x : (K →+* ℂ) → ℂ) : + x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ canonicalEmbedding K '' (𝓞 K) := by + rw [show Set.range (latticeBasis K) = (canonicalEmbedding K).toIntAlgHom.toLinearMap '' (Set.range (integralBasis K)) by - rw [← Set.range_comp] - exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] - exact Submodule.apply_mem_span_image_iff_mem_span (canonicalEmbedding_injective K) + rw [← Set.range_comp] ; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] + rw [← Submodule.map_span, ← SetLike.mem_coe, Submodule.map_coe] + rw [show (Submodule.span ℤ (Set.range (integralBasis K)) : Set K) = 𝓞 K by + ext ; exact mem_span_integralBasis K] + rfl end NumberField.canonicalEmbedding From 00c63dcc95a8faef0b0d5065797f031c4f343920 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sun, 2 Jul 2023 17:04:52 +0200 Subject: [PATCH 13/67] Move instance --- Mathlib/LinearAlgebra/Finsupp.lean | 13 +------------ .../NumberField/CanonicalEmbedding.lean | 1 - 2 files changed, 1 insertion(+), 13 deletions(-) diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index 27ad90d29110a..e7fd9ab263998 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -8,7 +8,6 @@ Authors: Johannes Hölzl ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathlib.Data.Set.Countable import Mathlib.Data.Finsupp.Defs import Mathlib.LinearAlgebra.Pi import Mathlib.LinearAlgebra.Span @@ -135,7 +134,7 @@ theorem ker_lsingle (a : α) : ker (lsingle a : M →ₗ[R] α →₀ M) = ⊥ : #align finsupp.ker_lsingle Finsupp.ker_lsingle theorem lsingle_range_le_ker_lapply (s t : Set α) (h : Disjoint s t) : - ⨆ a ∈ s, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M) ≤ + ⨆ a ∈ s, LinearMap.range (lsingle a : M →ₗ[R] α →₀ M) ≤ ⨅ a ∈ t, ker (lapply a : (α →₀ M) →ₗ[R] M) := by refine' iSup_le fun a₁ => iSup_le fun h₁ => range_le_iff_comap.2 _ simp only [(ker_comp _ _).symm, eq_top_iff, SetLike.le_def, mem_ker, comap_iInf, mem_iInf] @@ -1191,16 +1190,6 @@ theorem mem_span_set {m : M} {s : Set M} : exact Finsupp.mem_span_image_iff_total R (v := _root_.id (α := M)) #align mem_span_set mem_span_set -/-- A `R`-submodule spanned by a finite family of `v` with `R` countable is countable. -/ -instance {ι : Type _} [Countable R] [Fintype ι] (v : ι → M) : - Countable (Submodule.span R (Set.range v)) := by - change Countable (Submodule.span R (Set.range v) : Set M) - rw [Set.countable_coe_iff] - refine Set.Countable.mono ?_ (Set.countable_range (fun f : (ι → R) => ∑ i, (f i) • v i)) - intro _ h - obtain ⟨f, rfl⟩ := Finsupp.mem_span_range_iff_exists_finsupp.mp (SetLike.mem_coe.mp h) - exact ⟨↑f, by rw [Finsupp.sum_fintype] ; exact fun _ => zero_smul _ _⟩ - /-- If `Subsingleton R`, then `M ≃ₗ[R] ι →₀ R` for any type `ι`. -/ @[simps] def Module.subsingletonEquiv (R M ι : Type _) [Semiring R] [Subsingleton R] [AddCommMonoid M] diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 5d8cd79821a52..f45f3f4c2784d 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -379,7 +379,6 @@ theorem exists_ne_zero_mem_ring_of_integers_lt (h : minkowski_bound K < volume ( ∃ (a : 𝓞 K), a ≠ 0 ∧ ∀ w : InfinitePlace K, w a < f w := by have : @IsAddHaarMeasure (E K) _ _ _ volume := prod.instIsAddHaarMeasure volume volume have h_fund := Zspan.isAddFundamentalDomain (latticeBasis K) volume - have : Countable (Submodule.span ℤ (Set.range (latticeBasis K)) : Set (E K)) := inferInstance have : Countable (Submodule.span ℤ (Set.range (latticeBasis K))).toAddSubgroup := by change Countable (Submodule.span ℤ (Set.range (latticeBasis K)): Set (E K)) infer_instance From 00a90387160be0f0f341e8fc1df6fa22654ac895 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sun, 2 Jul 2023 17:21:23 +0200 Subject: [PATCH 14/67] 1st commit --- Mathlib/Data/Finsupp/Basic.lean | 4 ++++ Mathlib/LinearAlgebra/Finsupp.lean | 8 ++++++++ 2 files changed, 12 insertions(+) diff --git a/Mathlib/Data/Finsupp/Basic.lean b/Mathlib/Data/Finsupp/Basic.lean index 78d4c82147995..37d7c56001929 100644 --- a/Mathlib/Data/Finsupp/Basic.lean +++ b/Mathlib/Data/Finsupp/Basic.lean @@ -13,6 +13,7 @@ import Mathlib.Algebra.Hom.GroupAction import Mathlib.Algebra.Regular.SMul import Mathlib.Data.Finset.Preimage import Mathlib.Data.Rat.BigOperators +import Mathlib.Data.Set.Countable /-! # Miscellaneous definitions, lemmas, and constructions using finsupp @@ -123,6 +124,9 @@ theorem graph_eq_empty {f : α →₀ M} : f.graph = ∅ ↔ f = 0 := (graph_injective α M).eq_iff' graph_zero #align finsupp.graph_eq_empty Finsupp.graph_eq_empty +instance [Countable α] [Countable M] : + Countable (α →₀ M) := Function.Injective.countable (Finsupp.graph_injective α M) + end Graph end Finsupp diff --git a/Mathlib/LinearAlgebra/Finsupp.lean b/Mathlib/LinearAlgebra/Finsupp.lean index 829e8f6155973..de0f45f3dbc71 100644 --- a/Mathlib/LinearAlgebra/Finsupp.lean +++ b/Mathlib/LinearAlgebra/Finsupp.lean @@ -1028,6 +1028,14 @@ theorem finsuppProdLEquiv_symm_apply {α β R M : Type _} [Semiring R] [AddCommM end Prod +/-- If `R` is countable, then any `R`-submodule spanned by a countable family of vectors is +countable. -/ +instance {ι : Type _} [Countable R] [Countable ι] (v : ι → M) : + Countable (Submodule.span R (Set.range v)) := by + refine Set.countable_coe_iff.mpr (Set.Countable.mono ?_ (Set.countable_range + (fun c : (ι →₀ R) => c.sum fun i _ => (c i) • v i))) + exact fun _ h => Finsupp.mem_span_range_iff_exists_finsupp.mp (SetLike.mem_coe.mp h) + end Finsupp section Fintype From 7ef76091f3e76f1a4ea43fa0870a4611ff9eedab Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 5 Jul 2023 11:26:19 +0200 Subject: [PATCH 15/67] 1st commit --- Mathlib/Algebra/Module/Zlattice.lean | 192 +++++++++++++++++++++++++-- 1 file changed, 184 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Module/Zlattice.lean b/Mathlib/Algebra/Module/Zlattice.lean index 7ffddbf964b1a..3c652e0bd6948 100644 --- a/Mathlib/Algebra/Module/Zlattice.lean +++ b/Mathlib/Algebra/Module/Zlattice.lean @@ -8,25 +8,34 @@ Authors: Xavier Roblot ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ +import Mathlib.LinearAlgebra.FreeModule.PID import Mathlib.MeasureTheory.Group.FundamentalDomain +import Mathlib.RingTheory.Localization.Module /-! # ℤ-lattices Let `E` be a finite dimensional vector space over a `NormedLinearOrderedField` `K` with a solid -norm and that is also a `FloorRing`, e.g. `ℚ` or `ℝ`. A (full) ℤ-lattice `L` of `E` is a discrete +norm that is also a `FloorRing`, e.g. `ℚ` or `ℝ`. A (full) `ℤ`-lattice `L` of `E` is a discrete subgroup of `E` such that `L` spans `E` over `K`. -The ℤ-lattice `L` can be defined in two ways: -* For `b` a basis of `E`, then `Submodule.span ℤ (Set.range b)` is a ℤ-lattice of `E`. +A `ℤ`-lattice `L` can be defined in two ways: +* For `b` a basis of `E`, then `L = Submodule.span ℤ (Set.range b)` is a ℤ-lattice of `E` * As an `AddSubgroup E` with the additional properties: * `∀ r : ℝ, (L ∩ Metric.closedBall 0 r).finite`, that is `L` is discrete * `Submodule.span ℝ (L : Set E) = ⊤`, that is `L` spans `E` over `K`. -## Main result +Results about the first point of view are in the `Zspan` namespace and results about the second +point of view are in the `Zlattice` namespace. + +## Main results + * `Zspan.isAddFundamentalDomain`: for a ℤ-lattice `Submodule.span ℤ (Set.range b)`, proves that the set defined by `Zspan.fundamentalDomain` is a fundamental domain. - +* `Zlattice.module_free`: an addsubgroup of `E` that is discrete and spans `E` over `K` is a free +`ℤ`-module +* `Zlattice.rank`: an addsubgroup of `E` that is discrete and spans `E` over `K` is a free +`ℤ`-module has `ℤ`-rank equal to the `K`-rank of `E` -/ @@ -55,8 +64,7 @@ def fundamentalDomain : Set E := {m | ∀ i, b.repr m i ∈ Set.Ico (0 : K) 1} @[simp] theorem mem_fundamentalDomain {m : E} : - m ∈ fundamentalDomain b ↔ ∀ i, b.repr m i ∈ Set.Ico (0 : K) 1 := - Iff.rfl + m ∈ fundamentalDomain b ↔ ∀ i, b.repr m i ∈ Set.Ico (0 : K) 1 := Iff.rfl #align zspan.mem_fundamental_domain Zspan.mem_fundamentalDomain variable [FloorRing K] @@ -110,7 +118,8 @@ theorem ceil_eq_self_of_mem (m : E) (h : m ∈ span ℤ (Set.range b)) : (ceil b #align zspan.ceil_eq_self_of_mem Zspan.ceil_eq_self_of_mem /-- The map that sends a vector `E` to the `fundamentalDomain` of the lattice, -see `Zspan.fract_mem_fundamentalDomain`. -/ +see `Zspan.fract_mem_fundamentalDomain` and `fract_restrict` for the map with the codomain +restricted to `fundamentalDomain`. -/ def fract (m : E) : E := m - floor b m #align zspan.fract Zspan.fract @@ -156,6 +165,15 @@ theorem fract_mem_fundamentalDomain (x : E) : fract b x ∈ fundamentalDomain b fract_eq_self.mp (fract_fract b _) #align zspan.fract_mem_fundamental_domain Zspan.fract_mem_fundamentalDomain +/-- The map `fract` with codomain restricted to `fundamentalDomain`. -/ +def fract_restrict (x : E) : fundamentalDomain b := ⟨fract b x, fract_mem_fundamentalDomain b x⟩ + +theorem fract_restrict_surjective : Function.Surjective (fract_restrict b) := + fun x => ⟨↑x, Subtype.eq (fract_eq_self.mpr (Subtype.mem x))⟩ + +@[simp] +theorem fract_restrict_apply (x : E) : (fract_restrict b x : E) = fract b x := rfl + theorem fract_eq_fract (m n : E) : fract b m = fract b n ↔ -m + n ∈ span ℤ (Set.range b) := by classical rw [eq_comm, Basis.ext_elem_iff b] @@ -225,6 +243,27 @@ theorem exist_unique_vadd_mem_fundamentalDomain [Finite ι] (x : E) : · exact (vadd_mem_fundamentalDomain b y x).mp h #align zspan.exist_unique_vadd_mem_fundamental_domain Zspan.exist_unique_vadd_mem_fundamentalDomain +/-- The map `Zspan.fract` give an equiv map between `E ⧸ span ℤ (Set.range b)` +and `Zspan.fundamentalDomain b`. -/ +def QuotientEquiv [Fintype ι] : + E ⧸ span ℤ (Set.range b) ≃ (fundamentalDomain b) := by + refine Equiv.ofBijective ?_ ⟨fun x y => ?_, fun x => ?_⟩ + · refine fun q => Quotient.liftOn q (fract_restrict b) (fun _ _ h => ?_) + rw [Subtype.mk.injEq, fract_restrict_apply, fract_restrict_apply, fract_eq_fract] + exact QuotientAddGroup.leftRel_apply.mp h + · refine Quotient.inductionOn₂ x y (fun _ _ hxy => ?_) + rw [Quotient.liftOn_mk (s := quotientRel (span ℤ (Set.range b))), fract_restrict, + Quotient.liftOn_mk (s := quotientRel (span ℤ (Set.range b))), fract_restrict, + Subtype.mk.injEq] at hxy + apply Quotient.sound' + rwa [QuotientAddGroup.leftRel_apply, mem_toAddSubgroup, ← fract_eq_fract] + · obtain ⟨a, rfl⟩ := fract_restrict_surjective b x + exact ⟨Quotient.mk'' a, rfl⟩ + +@[simp] +theorem quotientEquiv_apply_mk [Fintype ι] (x : E) : + QuotientEquiv b (Submodule.Quotient.mk x) = fract_restrict b x := rfl + end NormedLatticeField section Real @@ -261,3 +300,140 @@ protected theorem isAddFundamentalDomain [Finite ι] [MeasurableSpace E] [OpensM end Real end Zspan + +section Zlattice + +open Submodule + +variable (K : Type _) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K] +variable {E : Type _} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] +variable {L : AddSubgroup E} +variable (hd : ∀ r : ℝ, ((L : Set E) ∩ (Metric.closedBall 0 r)).Finite) +variable (hs : span K (L : Set E) = ⊤) + +theorem Zlattice.FG : AddSubgroup.FG L := by + suffices (AddSubgroup.toIntSubmodule L).FG by exact (fg_iff_add_subgroup_fg _).mp this + obtain ⟨s, ⟨h_incl, ⟨h_span, h_lind⟩⟩⟩ := exists_linearIndependent K (L : Set E) + -- Let `s` be a maximal `K`-linear independent family of elements of `L`. We show that + -- `L` is finitely generated (as a ℤ-module) because it fits in the exact sequence + -- `0 → span ℤ s → L → L / span ℤ s → 0` with `span ℤ s` and `L / span ℤ s` finitely generated. + refine fg_of_fg_map_of_fg_inf_ker (span ℤ s).mkQ ?_ ?_ + · -- Let `b` be the `K`-basis of `E` formed by the vectors in `s`. The elements of + -- `L / span ℤ s = L /span ℤ b` are in bijection with elements of `L ∩ fundamentalDomain b` + -- so there are finitely many since `fundamentalDomain b` is bounded. + refine fg_def.mpr ⟨map (span ℤ s).mkQ (AddSubgroup.toIntSubmodule L), ?_, span_eq _⟩ + let b := Basis.mk h_lind (by + rw [← hs, ← h_span] + exact span_mono (by simp only [Subtype.range_coe_subtype, Set.setOf_mem_eq, subset_rfl])) + rw [show span ℤ s = span ℤ (Set.range b) by simp [Basis.coe_mk, Subtype.range_coe_subtype]] + have : Fintype s := Set.Finite.fintype h_lind.finite + refine Set.Finite.of_finite_image (f := ((↑) : _ → E) ∘ Zspan.QuotientEquiv b) ?_ + (Function.Injective.injOn (Subtype.coe_injective.comp (Zspan.QuotientEquiv b).injective) _) + obtain ⟨C, hC⟩ := Metric.Bounded.subset_ball (Zspan.fundamentalDomain_bounded b) 0 + refine Set.Finite.subset (hd C) ?_ + rintro _ ⟨_, ⟨⟨x, ⟨h_mem, rfl⟩⟩, rfl⟩⟩ + rw [Function.comp_apply, mkQ_apply, Zspan.quotientEquiv_apply_mk, Zspan.fract_restrict_apply] + refine ⟨?_, hC (Zspan.fract_mem_fundamentalDomain b x)⟩ + rw [Zspan.fract, SetLike.mem_coe, sub_eq_add_neg] + refine AddSubgroup.add_mem _ h_mem + (neg_mem (Set.mem_of_subset_of_mem ?_ (Subtype.mem (Zspan.floor b x)))) + rw [show (L : Set E) = AddSubgroup.toIntSubmodule L by rfl] + rw [SetLike.coe_subset_coe, Basis.coe_mk, Subtype.range_coe_subtype, Set.setOf_mem_eq] + exact span_le.mpr h_incl + · -- `span ℤ s` is finitely generated because `s` is finite + rw [ker_mkQ, inf_of_le_right (span_le.mpr h_incl)] + exact fg_span (LinearIndependent.finite h_lind) + +theorem Zlattice.module_finite : Module.Finite ℤ L := + Module.Finite.iff_addGroup_fg.mpr ((AddGroup.fg_iff_addSubgroup_fg L).mpr (FG K hd hs)) + +theorem Zlattice.module_free : Module.Free ℤ L := by + have : Module.Finite ℤ L := module_finite K hd hs + have : Module ℚ E := Module.compHom E (algebraMap ℚ K) + have : NoZeroSMulDivisors ℤ E := RatModule.noZeroSMulDivisors + have : NoZeroSMulDivisors ℤ L := by + change NoZeroSMulDivisors ℤ (AddSubgroup.toIntSubmodule L) + exact noZeroSMulDivisors _ + exact Module.free_of_finite_type_torsion_free' + +open FiniteDimensional + +theorem Zlattice.rank : finrank ℤ L = finrank K E := by + classical + have : Module.Finite ℤ L := module_finite K hd hs + have : Module.Free ℤ L := module_free K hd hs + have : Module ℚ E := Module.compHom E (algebraMap ℚ K) + let b₀ := Module.Free.chooseBasis ℤ L + -- Let `b` be a `ℤ`-basis of `L` formed of vectors of `E` + let b := Subtype.val ∘ b₀ + have : LinearIndependent ℤ b := + LinearIndependent.map' b₀.linearIndependent (L.toIntSubmodule.subtype) (ker_subtype _) + -- We prove some assertions that will be useful later on + have h_spanL : span ℤ (Set.range b) = AddSubgroup.toIntSubmodule L := by + convert congrArg (map (Submodule.subtype (AddSubgroup.toIntSubmodule L))) b₀.span_eq + · rw [map_span, Set.range_comp] + rfl + · exact (map_subtype_top _).symm + have h_spanE : span K (Set.range b) = ⊤ := by rwa [← span_span_of_tower (R := ℤ), h_spanL] + have h_card : Fintype.card (Module.Free.ChooseBasisIndex ℤ L) = + (Set.range b).toFinset.card := by + rw [Set.toFinset_range, Finset.univ.card_image_of_injective] + rfl + exact Subtype.coe_injective.comp (Basis.injective _) + rw [finrank_eq_card_chooseBasisIndex] + -- We prove that `finrank ℤ L ≤ finrank K E` and `finrank K E ≤ finrank ℤ L` + refine le_antisymm ?_ ?_ + · -- To prove that `finrank ℤ L ≤ finrank K E`, we proceed by contradiction and prove that, in + -- this case, there is a ℤ-relation between the vectors of `b` + obtain ⟨t, ⟨ht_inc, ⟨ht_span, ht_lin⟩⟩⟩ := exists_linearIndependent K (Set.range b) + -- `e` is a `K`-basis of `E` formed of vectors of `b` + let e : Basis t K E := Basis.mk ht_lin (by simp [ht_span, h_spanE]) + have : Fintype t := Set.Finite.fintype ((Set.range b).toFinite.subset ht_inc) + have h : LinearIndependent ℤ (fun x : (Set.range b) => (x : E)) := by + rwa [linearIndependent_subtype_range (Subtype.coe_injective.comp b₀.injective)] + contrapose! h + -- Since `finrank ℤ L ≤ finrank K E`, there exists a vector `v ∈ b` with `v ∉ e` + obtain ⟨v, hv⟩ : (Set.range b \ Set.range e).Nonempty := by + rw [Basis.coe_mk, Subtype.range_coe_subtype, Set.setOf_mem_eq, ← Set.toFinset_nonempty] + contrapose h + rw [Finset.not_nonempty_iff_eq_empty, Set.toFinset_diff, + Finset.sdiff_eq_empty_iff_subset] at h + replace h := Finset.card_le_of_subset h + rwa [not_lt, h_card, ← topEquiv.finrank_eq, ← h_spanE, ← ht_span, + finrank_span_set_eq_card _ ht_lin] + -- Assume that `e ∪ {v}` is not `ℤ`-linear independent then get the contradiction + suffices ¬ LinearIndependent ℤ (fun x : ↥(insert v (Set.range e)) => (x : E)) by + contrapose! this + refine LinearIndependent.mono ?_ this + exact Set.insert_subset (Set.mem_of_mem_diff hv) (by simp [ht_inc]) + -- We prove finally that `e ∪ {v}` is not ℤ-linear independent or, equivalently, + -- not ℚ-linear independent by showing that `v ∈ span ℚ e`. + rw [LinearIndependent.iff_fractionRing ℤ ℚ, + (linearIndependent_insert (Set.not_mem_of_mem_diff hv)), not_and, not_not] + intro _ + -- But that follows from the fact that there exist `n, m : ℕ`, `n ≠ m` + -- such that `(n - m) • v ∈ span ℤ e` because `n ↦ Zspan.fract e (n • v)` takes value + -- into the finite set `fundamentalDomain e ∩ L` + have : Set.MapsTo (fun n : ℤ => Zspan.fract e (n • v)) Set.univ + (L ∩ Metric.closedBall 0 (∑ i, ‖e i‖)) := by + rw [Set.mapsTo_inter, Set.maps_univ_to, Set.maps_univ_to] + refine ⟨fun _ => ?_, fun _ => mem_closedBall_zero_iff.mpr (Zspan.norm_fract_le e _)⟩ + · change _ ∈ AddSubgroup.toIntSubmodule L + rw [← h_spanL] + refine sub_mem ?_ ?_ + · exact zsmul_mem (subset_span (Set.diff_subset _ _ hv)) _ + · exact span_mono (by simp [ht_inc]) (coe_mem _) + obtain ⟨n, -, m, -, h_neq, h_eq⟩ := Set.Infinite.exists_ne_map_eq_of_mapsTo + Set.infinite_univ this (hd (∑ i, ‖e i‖)) + have h_nz : (-n + m : ℚ) ≠ 0 := by + rwa [Ne.def, add_eq_zero_iff_eq_neg.not, neg_inj, Rat.coe_int_inj, ← Ne.def] + apply (smul_mem_iff _ h_nz).mp + refine span_subset_span ℤ ℚ _ ?_ + rwa [add_smul, neg_smul, SetLike.mem_coe, ← Zspan.fract_eq_fract, ← zsmul_eq_smul_cast ℚ, + ← zsmul_eq_smul_cast ℚ] + · -- To prove that `finrank K E ≤ finrank ℤ L`, we use the fact `b` generates `E` over `K` + -- and thus `finrank K E ≤ card b = finrank ℤ L` + rw [← topEquiv.finrank_eq, ← h_spanE] + convert finrank_span_le_card (K := K) (Set.range b) + +end Zlattice From 3a56308f173c6bf0f06c63dbc0685bfa0d5aff78 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 5 Jul 2023 11:29:47 +0200 Subject: [PATCH 16/67] Add Group.fg_iff_Subgroup_fg --- Mathlib/GroupTheory/Finiteness.lean | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/Finiteness.lean b/Mathlib/GroupTheory/Finiteness.lean index d7a849ecec8e1..f5dc0960a4f65 100644 --- a/Mathlib/GroupTheory/Finiteness.lean +++ b/Mathlib/GroupTheory/Finiteness.lean @@ -136,7 +136,7 @@ instance Monoid.fg_of_addMonoid_fg [AddMonoid.FG N] : Monoid.FG (Multiplicative @[to_additive] instance (priority := 100) Monoid.fg_of_finite [Finite M] : Monoid.FG M := by cases nonempty_fintype M - exact ⟨⟨Finset.univ, by rw [Finset.coe_univ]; exact Submonoid.closure_univ⟩⟩ + exact ⟨⟨Finset.univ, by rw [Finset.coe_univ] ; exact Submonoid.closure_univ⟩⟩ #align monoid.fg_of_finite Monoid.fg_of_finite #align add_monoid.fg_of_finite AddMonoid.fg_of_finite @@ -325,6 +325,11 @@ theorem Group.fg_iff_monoid_fg : Group.FG G ↔ Monoid.FG G := #align group.fg_iff_monoid.fg Group.fg_iff_monoid_fg #align add_group.fg_iff_add_monoid.fg AddGroup.fg_iff_addMonoid_fg +@[to_additive (attr := simp)] +theorem Group.fg_iff_subgroup_fg (H : Subgroup G) : Group.FG H ↔ H.FG := + (fg_iff_monoid_fg.trans (Monoid.fg_iff_submonoid_fg _)).trans + (Subgroup.fg_iff_submonoid_fg _).symm + theorem GroupFG.iff_add_fg : Group.FG G ↔ AddGroup.FG (Additive G) := ⟨fun h => ⟨(Subgroup.fg_iff_add_fg ⊤).1 h.out⟩, fun h => ⟨(Subgroup.fg_iff_add_fg ⊤).2 h.out⟩⟩ #align group_fg.iff_add_fg GroupFG.iff_add_fg @@ -345,7 +350,7 @@ instance Group.fg_of_mul_group_fg [AddGroup.FG H] : Group.FG (Multiplicative H) @[to_additive] instance (priority := 100) Group.fg_of_finite [Finite G] : Group.FG G := by cases nonempty_fintype G - exact ⟨⟨Finset.univ, by rw [Finset.coe_univ]; exact Subgroup.closure_univ⟩⟩ + exact ⟨⟨Finset.univ, by rw [Finset.coe_univ] ; exact Subgroup.closure_univ⟩⟩ #align group.fg_of_finite Group.fg_of_finite #align add_group.fg_of_finite AddGroup.fg_of_finite From fa22093206343ad8ffcbca5750626dd8dc6d3b52 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Thu, 6 Jul 2023 18:01:44 +0200 Subject: [PATCH 17/67] semicolon --- Mathlib/GroupTheory/Finiteness.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/Finiteness.lean b/Mathlib/GroupTheory/Finiteness.lean index f5dc0960a4f65..bb846ba6ca267 100644 --- a/Mathlib/GroupTheory/Finiteness.lean +++ b/Mathlib/GroupTheory/Finiteness.lean @@ -136,7 +136,7 @@ instance Monoid.fg_of_addMonoid_fg [AddMonoid.FG N] : Monoid.FG (Multiplicative @[to_additive] instance (priority := 100) Monoid.fg_of_finite [Finite M] : Monoid.FG M := by cases nonempty_fintype M - exact ⟨⟨Finset.univ, by rw [Finset.coe_univ] ; exact Submonoid.closure_univ⟩⟩ + exact ⟨⟨Finset.univ, by rw [Finset.coe_univ]; exact Submonoid.closure_univ⟩⟩ #align monoid.fg_of_finite Monoid.fg_of_finite #align add_monoid.fg_of_finite AddMonoid.fg_of_finite @@ -350,7 +350,7 @@ instance Group.fg_of_mul_group_fg [AddGroup.FG H] : Group.FG (Multiplicative H) @[to_additive] instance (priority := 100) Group.fg_of_finite [Finite G] : Group.FG G := by cases nonempty_fintype G - exact ⟨⟨Finset.univ, by rw [Finset.coe_univ] ; exact Subgroup.closure_univ⟩⟩ + exact ⟨⟨Finset.univ, by rw [Finset.coe_univ]; exact Subgroup.closure_univ⟩⟩ #align group.fg_of_finite Group.fg_of_finite #align add_group.fg_of_finite AddGroup.fg_of_finite From de183ff9862375be655303343aa60b50f9cc96be Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Thu, 6 Jul 2023 22:06:14 +0200 Subject: [PATCH 18/67] 1st commit --- Mathlib/GroupTheory/Torsion.lean | 4 +- Mathlib/NumberTheory/NumberField/Units.lean | 60 +++++++++++++++++++++ 2 files changed, 63 insertions(+), 1 deletion(-) diff --git a/Mathlib/GroupTheory/Torsion.lean b/Mathlib/GroupTheory/Torsion.lean index 49a729a7d3eef..86e6fcdc99b75 100644 --- a/Mathlib/GroupTheory/Torsion.lean +++ b/Mathlib/GroupTheory/Torsion.lean @@ -326,6 +326,9 @@ theorem torsion_eq_torsion_submonoid : CommMonoid.torsion G = (torsion G).toSubm #align comm_group.torsion_eq_torsion_submonoid CommGroup.torsion_eq_torsion_submonoid #align add_comm_group.add_torsion_eq_add_torsion_submonoid AddCommGroup.add_torsion_eq_add_torsion_submonoid +@[to_additive] +theorem mem_torsion (g : G) : g ∈ torsion G ↔ IsOfFinOrder g := Iff.rfl + variable (p : ℕ) [hp : Fact p.Prime] /-- The `p`-primary component is the subgroup of elements with order prime-power of `p`. -/ @@ -441,4 +444,3 @@ theorem IsTorsionFree.quotient_torsion : IsTorsionFree <| G ⧸ torsion G := fun #align add_is_torsion_free.quotient_torsion AddIsTorsionFree.quotient_torsion end CommGroup - diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 32a20af73d41e..55054ea7c7d6a 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -8,7 +8,10 @@ Authors: Xavier Roblot ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ +import Mathlib.GroupTheory.Torsion +import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.NumberTheory.NumberField.Norm +import Mathlib.RingTheory.RootsOfUnity.Basic /-! # Units of a number field @@ -54,3 +57,60 @@ theorem isUnit_iff_norm [NumberField K] (x : 𝓞 K) : #align is_unit_iff_norm isUnit_iff_norm end IsUnit + +namespace NumberField.units + +section coe + +/-- The `MonoidHom` from the group of units `(𝓞 K)ˣ` to the field `K`. -/ +def coe_to_field : (𝓞 K)ˣ →* K := (Units.coeHom K).comp (map (algebraMap (𝓞 K) K)) + +theorem coe_to_field_injective : Function.Injective (coe_to_field K) := + fun _ _ h => Units.eq_iff.mp (SetCoe.ext h) + +/-- There is a natural coercion from `(𝓞 K)ˣ` to `(𝓞 K)` and then from `(𝓞 K)` to `K` but it is +useful to also have a direct one from `(𝓞 K)ˣ` to `K`. -/ +instance : Coe (𝓞 K)ˣ K := ⟨coe_to_field K⟩ + +@[ext] +theorem ext {x y : (𝓞 K)ˣ} (h : (x : K) = y) : x = y := (coe_to_field_injective K).eq_iff.mp h + +end coe + +open NumberField.InfinitePlace + +section torsion + +/-- The torsion subgroup of the group of units. -/ +def torsion : Subgroup (𝓞 K)ˣ := CommGroup.torsion (𝓞 K)ˣ + +theorem mem_torsion {x : (𝓞 K)ˣ} [NumberField K] : + x ∈ torsion K ↔ ∀ w : InfinitePlace K, w x = 1 := by + rw [eq_iff_eq (x : K) 1, torsion, CommGroup.mem_torsion, isOfFinOrder_iff_pow_eq_one] + refine ⟨fun ⟨n, h_pos, h_eq⟩ φ => ?_, fun h => ?_⟩ + · refine norm_map_one_of_pow_eq_one φ.toMonoidHom (k := ⟨n, h_pos⟩) ?_ + rw [PNat.mk_coe, ← map_pow, h_eq, map_one] + · obtain ⟨n, hn, hx⟩ := Embeddings.pow_eq_one_of_norm_eq_one K ℂ x.val.prop h + exact ⟨n, hn, by ext; rwa [map_pow, map_one]⟩ +end torsion + +instance : Nonempty (torsion K) := ⟨1⟩ + +/-- The torsion subgroup is finite. -/ +instance [NumberField K] : Fintype (torsion K) := by + refine @Fintype.ofFinite _ (Set.finite_coe_iff.mpr ?_) + refine Set.Finite.of_finite_image ?_ ((coe_to_field_injective K).injOn _) + refine (Embeddings.finite_of_norm_le K ℂ 1).subset + (fun a ⟨u, ⟨h_tors, h_ua⟩⟩ => ⟨?_, fun φ => ?_⟩) + · rw [← h_ua] + exact u.val.prop + · rw [← h_ua] + exact le_of_eq ((eq_iff_eq _ 1).mp ((mem_torsion K).mp h_tors) φ) + +/-- The torsion subgroup is cylic. -/ +instance [NumberField K] : IsCyclic (torsion K) := subgroup_units_cyclic _ + +/-- The order of the torsion subgroup as positive integer. -/ +def torsion_order [NumberField K] : ℕ+ := ⟨Fintype.card (torsion K), Fintype.card_pos⟩ + +end NumberField.units From 993f95a36e36bf7c2cbd442174432e8665bd60c7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 8 Jul 2023 11:27:18 +0200 Subject: [PATCH 19/67] add results about rootsOfUnity --- Mathlib/NumberTheory/NumberField/Units.lean | 25 +++++++++++++++++++++ 1 file changed, 25 insertions(+) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 55054ea7c7d6a..44b1151684e22 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -113,4 +113,29 @@ instance [NumberField K] : IsCyclic (torsion K) := subgroup_units_cyclic _ /-- The order of the torsion subgroup as positive integer. -/ def torsion_order [NumberField K] : ℕ+ := ⟨Fintype.card (torsion K), Fintype.card_pos⟩ +/-- If `k` does not divide `torsion_order` then there are no nontrivial roots of unity of + order dividing `k`. -/ +theorem rootsOfUnity_eq_one [NumberField K] {k : ℕ+} (hc : Nat.coprime k (torsion_order K)) : + ζ ∈ rootsOfUnity k (𝓞 K) ↔ ζ = 1 := by + rw [mem_rootsOfUnity] + refine ⟨fun h => ?_, fun h => by rw [h, one_pow]⟩ + refine orderOf_eq_one_iff.mp (Nat.eq_one_of_dvd_coprimes hc ?_ ?_) + · exact orderOf_dvd_of_pow_eq_one h + · have hζ : ζ ∈ torsion K := by + rw [torsion, CommGroup.mem_torsion, isOfFinOrder_iff_pow_eq_one] + exact ⟨k, k.prop, h⟩ + rw [orderOf_submonoid (⟨ζ, hζ⟩ : torsion K)] + exact orderOf_dvd_card_univ + +/-- The group of roots of unity of order dividing `torsion_order` is equal to the torsion +group. -/ +theorem rootsOfUnity_eq_torsion [NumberField K] : + rootsOfUnity (torsion_order K) (𝓞 K) = torsion K := by + ext ζ + rw [torsion, mem_rootsOfUnity] + refine ⟨fun h => ?_, fun h => ?_⟩ + · rw [CommGroup.mem_torsion, isOfFinOrder_iff_pow_eq_one] + exact ⟨↑(torsion_order K), (torsion_order K).prop, h⟩ + · exact Subtype.ext_iff.mp (@pow_card_eq_one (torsion K) ⟨ζ, h⟩ _ _) + end NumberField.units From 21d7345e2ef902913f4ed8c55ee40ec2689292da Mon Sep 17 00:00:00 2001 From: Xavier Roblot <46200072+xroblot@users.noreply.github.com> Date: Sat, 8 Jul 2023 15:14:37 +0200 Subject: [PATCH 20/67] Update Zlattice.lean --- Mathlib/Algebra/Module/Zlattice.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Module/Zlattice.lean b/Mathlib/Algebra/Module/Zlattice.lean index 3c652e0bd6948..fd52895413bfd 100644 --- a/Mathlib/Algebra/Module/Zlattice.lean +++ b/Mathlib/Algebra/Module/Zlattice.lean @@ -35,7 +35,7 @@ the set defined by `Zspan.fundamentalDomain` is a fundamental domain. * `Zlattice.module_free`: an addsubgroup of `E` that is discrete and spans `E` over `K` is a free `ℤ`-module * `Zlattice.rank`: an addsubgroup of `E` that is discrete and spans `E` over `K` is a free -`ℤ`-module has `ℤ`-rank equal to the `K`-rank of `E` +`ℤ`-module of `ℤ`-rank equal to the `K`-rank of `E` -/ From eb432f746ec08bd2992702dbaf2cccf427e15499 Mon Sep 17 00:00:00 2001 From: Xavier Roblot <46200072+xroblot@users.noreply.github.com> Date: Sat, 8 Jul 2023 15:16:25 +0200 Subject: [PATCH 21/67] Update Zlattice.lean --- Mathlib/Algebra/Module/Zlattice.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Module/Zlattice.lean b/Mathlib/Algebra/Module/Zlattice.lean index fd52895413bfd..e9733d1fa11dd 100644 --- a/Mathlib/Algebra/Module/Zlattice.lean +++ b/Mathlib/Algebra/Module/Zlattice.lean @@ -118,7 +118,7 @@ theorem ceil_eq_self_of_mem (m : E) (h : m ∈ span ℤ (Set.range b)) : (ceil b #align zspan.ceil_eq_self_of_mem Zspan.ceil_eq_self_of_mem /-- The map that sends a vector `E` to the `fundamentalDomain` of the lattice, -see `Zspan.fract_mem_fundamentalDomain` and `fract_restrict` for the map with the codomain +see `Zspan.fract_mem_fundamentalDomain`, and `fract_restrict` for the map with the codomain restricted to `fundamentalDomain`. -/ def fract (m : E) : E := m - floor b m #align zspan.fract Zspan.fract @@ -243,7 +243,7 @@ theorem exist_unique_vadd_mem_fundamentalDomain [Finite ι] (x : E) : · exact (vadd_mem_fundamentalDomain b y x).mp h #align zspan.exist_unique_vadd_mem_fundamental_domain Zspan.exist_unique_vadd_mem_fundamentalDomain -/-- The map `Zspan.fract` give an equiv map between `E ⧸ span ℤ (Set.range b)` +/-- The map `Zspan.fract_restrict` defines an equiv map between `E ⧸ span ℤ (Set.range b)` and `Zspan.fundamentalDomain b`. -/ def QuotientEquiv [Fintype ι] : E ⧸ span ℤ (Set.range b) ≃ (fundamentalDomain b) := by From 1a266f6ba94a0e2f9dd886b74ddf1f639ca782b6 Mon Sep 17 00:00:00 2001 From: Xavier Roblot <46200072+xroblot@users.noreply.github.com> Date: Sat, 8 Jul 2023 15:17:29 +0200 Subject: [PATCH 22/67] Update Zlattice.lean --- Mathlib/Algebra/Module/Zlattice.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Module/Zlattice.lean b/Mathlib/Algebra/Module/Zlattice.lean index e9733d1fa11dd..2b55dbb3a9939 100644 --- a/Mathlib/Algebra/Module/Zlattice.lean +++ b/Mathlib/Algebra/Module/Zlattice.lean @@ -392,7 +392,7 @@ theorem Zlattice.rank : finrank ℤ L = finrank K E := by have h : LinearIndependent ℤ (fun x : (Set.range b) => (x : E)) := by rwa [linearIndependent_subtype_range (Subtype.coe_injective.comp b₀.injective)] contrapose! h - -- Since `finrank ℤ L ≤ finrank K E`, there exists a vector `v ∈ b` with `v ∉ e` + -- Since `finrank ℤ L > finrank K E`, there exists a vector `v ∈ b` with `v ∉ e` obtain ⟨v, hv⟩ : (Set.range b \ Set.range e).Nonempty := by rw [Basis.coe_mk, Subtype.range_coe_subtype, Set.setOf_mem_eq, ← Set.toFinset_nonempty] contrapose h From 95d2fcd5e11bde4d8b5f7b074c9c9429690bb1aa Mon Sep 17 00:00:00 2001 From: Xavier Roblot <46200072+xroblot@users.noreply.github.com> Date: Sat, 8 Jul 2023 15:20:09 +0200 Subject: [PATCH 23/67] Update Zlattice.lean --- Mathlib/Algebra/Module/Zlattice.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Module/Zlattice.lean b/Mathlib/Algebra/Module/Zlattice.lean index 2b55dbb3a9939..f728e9e14d687 100644 --- a/Mathlib/Algebra/Module/Zlattice.lean +++ b/Mathlib/Algebra/Module/Zlattice.lean @@ -401,7 +401,7 @@ theorem Zlattice.rank : finrank ℤ L = finrank K E := by replace h := Finset.card_le_of_subset h rwa [not_lt, h_card, ← topEquiv.finrank_eq, ← h_spanE, ← ht_span, finrank_span_set_eq_card _ ht_lin] - -- Assume that `e ∪ {v}` is not `ℤ`-linear independent then get the contradiction + -- Assume that `e ∪ {v}` is not `ℤ`-linear independent then we get the contradiction suffices ¬ LinearIndependent ℤ (fun x : ↥(insert v (Set.range e)) => (x : E)) by contrapose! this refine LinearIndependent.mono ?_ this @@ -412,8 +412,8 @@ theorem Zlattice.rank : finrank ℤ L = finrank K E := by (linearIndependent_insert (Set.not_mem_of_mem_diff hv)), not_and, not_not] intro _ -- But that follows from the fact that there exist `n, m : ℕ`, `n ≠ m` - -- such that `(n - m) • v ∈ span ℤ e` because `n ↦ Zspan.fract e (n • v)` takes value - -- into the finite set `fundamentalDomain e ∩ L` + -- such that `(n - m) • v ∈ span ℤ e` which is true since `n ↦ Zspan.fract e (n • v)` + -- takes value into the finite set `fundamentalDomain e ∩ L` have : Set.MapsTo (fun n : ℤ => Zspan.fract e (n • v)) Set.univ (L ∩ Metric.closedBall 0 (∑ i, ‖e i‖)) := by rw [Set.mapsTo_inter, Set.maps_univ_to, Set.maps_univ_to] From f999e2efecbba2025f42b60d45df2a73cafc8bc1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Mon, 17 Jul 2023 17:06:15 +0900 Subject: [PATCH 24/67] 1st commit --- .../LinearAlgebra/Matrix/ToLinearEquiv.lean | 41 +- .../NumberField/CanonicalEmbedding.lean | 490 +++++++++++++----- Mathlib/NumberTheory/NumberField/Units.lean | 368 ++++++++++++- 3 files changed, 766 insertions(+), 133 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean index 694455b6cacac..5be3092e83d7b 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean @@ -39,11 +39,15 @@ matrix, linear_equiv, determinant, inverse namespace Matrix +variable {n : Type _} [Fintype n] + +section LinearEquiv + open LinearMap variable {R M : Type _} [CommRing R] [AddCommGroup M] [Module R M] -variable {n : Type _} [Fintype n] + section ToLinearEquiv' @@ -199,4 +203,39 @@ alias nondegenerate_iff_det_ne_zero ↔ Nondegenerate.det_ne_zero Nondegenerate. end Nondegenerate +end LinearEquiv + +section Determinant + +open BigOperators + +lemma det_ne_zero_of_neg [DecidableEq n] {S : Type _} [LinearOrderedCommRing S] + {A : Matrix n n S} (h1 : ∀ i j, i ≠ j → A i j < 0) (h2 : ∀ j, 0 < ∑ i, A i j) : + A.det ≠ 0 := by + by_cases hn : Nonempty n + · contrapose! h2 + obtain ⟨v, ⟨h_vnz, h_vA⟩⟩ := Matrix.exists_vecMul_eq_zero_iff.mpr h2 + wlog h_sup : 0 < Finset.sup' Finset.univ Finset.univ_nonempty v + · refine this h1 hn h2 (-1 • v) ?_ ?_ ?_ + · exact smul_ne_zero (by norm_num) h_vnz + · rw [Matrix.vecMul_smul, h_vA, smul_zero] + · obtain ⟨i, hi⟩ := Function.ne_iff.mp h_vnz + simp_rw [Finset.lt_sup'_iff, Finset.mem_univ, true_and] at h_sup ⊢ + simp_rw [not_exists, not_lt] at h_sup + refine ⟨i, ?_⟩ + rw [Pi.smul_apply, neg_smul, one_smul, Left.neg_pos_iff] + refine Ne.lt_of_le hi (h_sup i) + · obtain ⟨j₀, -, h_j₀⟩ := Finset.exists_mem_eq_sup' Finset.univ_nonempty v + refine ⟨j₀, ?_⟩ + rw [← mul_le_mul_left (h_j₀ ▸ h_sup), Finset.mul_sum, mul_zero] + rw [show 0 = ∑ i, v i * A i j₀ from (congrFun h_vA j₀).symm] + refine Finset.sum_le_sum (fun i hi => ?_) + by_cases h : i = j₀ + · rw [h] + · exact (mul_le_mul_right_of_neg (h1 i j₀ h)).mpr (h_j₀ ▸ Finset.le_sup' v hi) + · have := (isEmpty_or_nonempty n).resolve_right hn + simp + +end Determinant + end Matrix diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 8cc8355934c9d..51945652bd4d1 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -8,15 +8,19 @@ Authors: Xavier Roblot ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ +import Mathlib.Algebra.Module.Zlattice +import Mathlib.MeasureTheory.Group.GeometryOfNumbers +import Mathlib.MeasureTheory.Measure.Haar.NormedSpace import Mathlib.NumberTheory.NumberField.Embeddings +import Mathlib.RingTheory.Discriminant + /-! # Canonical embedding of a number field -The canonical embedding of a number field `K` of signature `(r₁, r₂)` is the ring homomorphism -`K →+* ℝ^r₁ × ℂ^r₂` that sends `x ∈ K` to `(φ_₁(x),...,φ_r₁(x)) × (ψ_₁(x),..., ψ_r₂(x))` where -`φ_₁,...,φ_r₁` are its real embeddings and `ψ_₁,..., ψ_r₂` are its complex embeddings (up to -complex conjugation). +The canonical embedding of a number field `K` of degree `n` is the ring homomorphism +`K →+* ℂ^n` that sends `x ∈ K` to `(φ_₁(x),...,φ_n(x))` where the `φ_i`'s are the complex +embeddings of `K`. ## Main definitions and results @@ -24,164 +28,388 @@ complex conjugation). image of the ring of integers by the canonical embedding and any ball centered at `0` of finite radius is finite. +* `NumberField.mixedEmbedding`: the ring homomorphism `K →+* ℝ^r₁ × ℂ^r₂` that sends `x ∈ K` to +`(φ_₁(x),...,φ_r₁(x)) × (ψ_₁(x),..., ψ_r₂(x)) ` where `(r₁, r₂)` is the signature of `K`, +`φ_₁,...,φ_r₁` are its real embeddings and `ψ_₁,..., ψ_r₂` are its complex embeddings (up to +complex conjugation). + + ## Tags number field, infinite places -/ - -noncomputable section - -open Function FiniteDimensional Finset Fintype NumberField NumberField.InfinitePlace Metric Module - -open scoped Classical NumberField - variable (K : Type _) [Field K] namespace NumberField.canonicalEmbedding --- The ambient space `ℝ^r₁ × ℂ^r₂` with `(r₁, r₂)` the signature of `K`. -set_option quotPrecheck false in -- Porting note: Added. -scoped[CanonicalEmbedding] - notation "E" => - ({ w : InfinitePlace K // IsReal w } → ℝ) × ({ w : InfinitePlace K // IsComplex w } → ℂ) - -open CanonicalEmbedding - -theorem space_rank [NumberField K] : finrank ℝ E = finrank ℚ K := by - haveI : Module.Free ℝ ℂ := inferInstance - rw [finrank_prod, finrank_pi, finrank_pi_fintype, Complex.finrank_real_complex, Finset.sum_const, - Finset.card_univ, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, - ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, - Nat.add_sub_of_le (Fintype.card_subtype_le _)] -#align number_field.canonical_embedding.space_rank NumberField.canonicalEmbedding.space_rank - -theorem nontrivial_space [NumberField K] : Nontrivial E := by - have : Nonempty <| InfinitePlace K := inferInstance - rcases this with ⟨w⟩ - obtain hw | hw := w.isReal_or_isComplex - · haveI : Nonempty { w : InfinitePlace K // IsReal w } := ⟨⟨w, hw⟩⟩ - exact nontrivial_prod_left - · haveI : Nonempty { w : InfinitePlace K // IsComplex w } := ⟨⟨w, hw⟩⟩ - exact nontrivial_prod_right -#align number_field.canonical_embedding.non_trivial_space NumberField.canonicalEmbedding.nontrivial_space +open NumberField -/-- The canonical embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ -def _root_.NumberField.canonicalEmbedding : K →+* E := - RingHom.prod (Pi.ringHom fun w => w.prop.embedding) (Pi.ringHom fun w => w.val.embedding) -#align number_field.canonical_embedding NumberField.canonicalEmbedding +/-- The canonical embedding of a number field `K` of degree `n` into `ℂ^n`. -/ +def _root_.NumberField.canonicalEmbedding : K →+* ((K →+* ℂ) → ℂ) := Pi.ringHom fun φ => φ theorem _root_.NumberField.canonicalEmbedding_injective [NumberField K] : - Injective (NumberField.canonicalEmbedding K) := - @RingHom.injective _ _ _ _ (nontrivial_space K) _ -#align number_field.canonical_embedding_injective NumberField.canonicalEmbedding_injective - -open NumberField + Function.Injective (NumberField.canonicalEmbedding K) := RingHom.injective _ variable {K} @[simp] -theorem apply_at_real_infinitePlace (w : { w : InfinitePlace K // IsReal w }) (x : K) : - (NumberField.canonicalEmbedding K x).1 w = w.prop.embedding x := by - simp only [canonicalEmbedding, RingHom.prod_apply, Pi.ringHom_apply] -#align number_field.canonical_embedding.apply_at_real_infinite_place NumberField.canonicalEmbedding.apply_at_real_infinitePlace +theorem apply_at (φ : K →+* ℂ) (x : K) : + (NumberField.canonicalEmbedding K x) φ = φ x := rfl -@[simp] -theorem apply_at_complex_infinitePlace (w : { w : InfinitePlace K // IsComplex w }) (x : K) : - (NumberField.canonicalEmbedding K x).2 w = embedding w.val x := by - simp only [canonicalEmbedding, RingHom.prod_apply, Pi.ringHom_apply] -#align number_field.canonical_embedding.apply_at_complex_infinite_place NumberField.canonicalEmbedding.apply_at_complex_infinitePlace +open scoped ComplexConjugate + +/-- The image of `canonicalEmbedding` lives in the `ℝ`-submodule of the `x ∈ ((K →+* ℂ) → ℂ)` such +that `conj x_φ = x_(conj φ)` for all `∀ φ : K →+* ℂ`. -/ +theorem conj_apply {x : ((K →+* ℂ) → ℂ)} (φ : K →+* ℂ) + (hx : x ∈ Submodule.span ℝ (Set.range (canonicalEmbedding K))) : + conj (x φ) = x (ComplexEmbedding.conjugate φ) := by + refine Submodule.span_induction hx ?_ ?_ (fun _ _ hx hy => ?_) (fun a _ hx => ?_) + · rintro _ ⟨x, rfl⟩ + rw [apply_at, apply_at, ComplexEmbedding.conjugate_coe_eq] + · rw [Pi.zero_apply, Pi.zero_apply, map_zero] + · rw [Pi.add_apply, Pi.add_apply, map_add, hx, hy] + · rw [Pi.smul_apply, Complex.real_smul, map_mul, Complex.conj_ofReal] + exact congrArg ((a : ℂ) * ·) hx theorem nnnorm_eq [NumberField K] (x : K) : ‖canonicalEmbedding K x‖₊ = - Finset.univ.sup fun w : InfinitePlace K => ⟨w x, map_nonneg w x⟩ := by - rw [Prod.nnnorm_def', Pi.nnnorm_def, Pi.nnnorm_def] - rw [(_ : Finset.univ = - {w : InfinitePlace K | IsReal w}.toFinset ∪ {w : InfinitePlace K | IsComplex w}.toFinset)] - · rw [Finset.sup_union, sup_eq_max] - refine' congr_arg₂ _ _ _ - · convert - (Finset.univ.sup_map (Function.Embedding.subtype fun w : InfinitePlace K => IsReal w) - fun w => (⟨w x, map_nonneg w x⟩ : NNReal)).symm using 2 - ext w - dsimp - simp only [apply_at_real_infinitePlace, coe_nnnorm, Real.norm_eq_abs, - Function.Embedding.coe_subtype, Subtype.coe_mk, IsReal.abs_embedding_apply] - · convert - (Finset.univ.sup_map (Function.Embedding.subtype fun w : InfinitePlace K => IsComplex w) - fun w => (⟨w x, map_nonneg w x⟩ : NNReal)).symm using 2 - ext w - dsimp - simp only [apply_at_complex_infinitePlace, coe_nnnorm, - Complex.norm_eq_abs, Function.Embedding.coe_subtype, Subtype.coe_mk, abs_embedding] - · ext w - simp_rw [Finset.mem_univ, Finset.mem_union, Set.mem_toFinset, Set.mem_setOf_eq, - w.isReal_or_isComplex] -#align number_field.canonical_embedding.nnnorm_eq NumberField.canonicalEmbedding.nnnorm_eq + Finset.univ.sup (fun φ : K →+* ℂ => ‖φ x‖₊) := by + simp_rw [Pi.nnnorm_def, apply_at] theorem norm_le_iff [NumberField K] (x : K) (r : ℝ) : - ‖canonicalEmbedding K x‖ ≤ r ↔ ∀ w : InfinitePlace K, w x ≤ r := by + ‖canonicalEmbedding K x‖ ≤ r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by obtain hr | hr := lt_or_le r 0 - · have : Nonempty <| InfinitePlace K := inferInstance - rcases this with ⟨w⟩ - exact iff_of_false - (hr.trans_le <| norm_nonneg _).not_le fun h => hr.not_le <| (map_nonneg w _).trans <| h _ + · obtain ⟨φ : K →+* ℂ⟩ := (inferInstance : Nonempty (K →+* ℂ)) + refine iff_of_false ?_ ?_ + exact (hr.trans_le (norm_nonneg _)).not_le + exact fun h => hr.not_le (le_trans (norm_nonneg _) (h φ)) · lift r to NNReal using hr simp_rw [← coe_nnnorm, nnnorm_eq, NNReal.coe_le_coe, Finset.sup_le_iff, Finset.mem_univ, - forall_true_left, ← NNReal.coe_le_coe, NNReal.toReal] -#align number_field.canonical_embedding.norm_le_iff NumberField.canonicalEmbedding.norm_le_iff + forall_true_left] variable (K) -/-- The image of `𝓞 K` as a subring of `ℝ^r₁ × ℂ^r₂`. -/ -def integerLattice : Subring E := +/-- The image of `𝓞 K` as a subring of `ℂ^n`. -/ +def integerLattice : Subring ((K →+* ℂ) → ℂ) := (RingHom.range (algebraMap (𝓞 K) K)).map (canonicalEmbedding K) -#align number_field.canonical_embedding.integer_lattice NumberField.canonicalEmbedding.integerLattice - --- Porting note: See https://github.com/leanprover-community/mathlib4/issues/5028 -set_option maxHeartbeats 400000 in -set_option synthInstance.maxHeartbeats 50000 in -/-- The linear equiv between `𝓞 K` and the integer lattice. -/ -def equivIntegerLattice [NumberField K] : 𝓞 K ≃ₗ[ℤ] integerLattice K := - LinearEquiv.ofBijective - { toFun := fun x => (by - refine ⟨canonicalEmbedding K (algebraMap (𝓞 K) K x), ⟨algebraMap (𝓞 K) K x, ⟨?_, rfl⟩⟩⟩ - simp only [Subsemiring.coe_carrier_toSubmonoid, Subring.coe_toSubsemiring, - RingHom.coe_range, Set.mem_range, exists_apply_eq_apply] ) - map_add' := fun x y => (by - apply Subtype.eq - simp [map_add] ) - map_smul' := fun c x => (by - simp only [RingHom.id_apply, zsmul_eq_mul, RingHom.map_mul, map_intCast] - rfl ) } - (by - refine ⟨fun _ _ h => ?_, fun ⟨_, ⟨_, ⟨⟨a, rfl⟩, rfl⟩⟩⟩ => ⟨a, rfl⟩⟩ - dsimp only at h - rw [LinearMap.coe_mk, Subtype.mk_eq_mk] at h - exact IsFractionRing.injective (𝓞 K) K (canonicalEmbedding_injective K h)) -#align number_field.canonical_embedding.equiv_integer_lattice NumberField.canonicalEmbedding.equivIntegerLattice theorem integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : - ((integerLattice K : Set E) ∩ closedBall 0 r).Finite := by + ((integerLattice K : Set ((K →+* ℂ) → ℂ)) ∩ Metric.closedBall 0 r).Finite := by obtain hr | _ := lt_or_le r 0 - · simp [closedBall_eq_empty.2 hr] - have heq : ∀ x, canonicalEmbedding K x ∈ closedBall (0 : E) r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by - simp only [← place_apply, ← InfinitePlace.coe_mk, mem_closedBall_zero_iff, norm_le_iff] - exact fun x => le_iff_le x r - convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) - ext; constructor - · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx2⟩ - exact ⟨x, ⟨SetLike.coe_mem x, (heq x).mp hx2⟩, rfl⟩ - · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ - exact ⟨⟨x, ⟨⟨x, hx1⟩, rfl⟩, rfl⟩, (heq x).mpr hx2⟩ -#align number_field.canonical_embedding.integer_lattice.inter_ball_finite NumberField.canonicalEmbedding.integerLattice.inter_ball_finite - -instance [NumberField K] : Countable (integerLattice K) := by - have : (⋃ n : ℕ, (integerLattice K : Set E) ∩ closedBall 0 n).Countable := - Set.countable_iUnion fun n => (integerLattice.inter_ball_finite K n).countable - refine' (this.mono _).to_subtype - rintro _ ⟨x, hx, rfl⟩ - rw [Set.mem_iUnion] - exact ⟨⌈‖canonicalEmbedding K x‖⌉₊, ⟨x, hx, rfl⟩, mem_closedBall_zero_iff.2 (Nat.le_ceil _)⟩ + · simp [Metric.closedBall_eq_empty.2 hr] + · have heq : ∀ x, canonicalEmbedding K x ∈ Metric.closedBall 0 r ↔ + ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by + intro x ; rw [← norm_le_iff, mem_closedBall_zero_iff] + convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) + ext ; constructor + · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx⟩ + exact ⟨↑x, ⟨SetLike.coe_mem x, fun φ => (heq x).mp hx φ⟩, rfl⟩ + · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ + exact ⟨⟨x, ⟨⟨x, hx1⟩, rfl⟩, rfl⟩, (heq x).mpr hx2⟩ + +open Module Fintype FiniteDimensional + +/-- A `ℂ`-basis of `ℂ^n` that is also a `ℤ`-basis of the `integerLattice`. -/ +noncomputable def latticeBasis [NumberField K] : + Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℂ ((K →+* ℂ) → ℂ) := by + classical + -- Let `B` be the canonical basis of `(K →+* ℂ) → ℂ`. We prove that the determinant of + -- the image by `canonicalEmbedding` of the integral basis of `K` is non-zero. This + -- will imply the result. + let B := Pi.basisFun ℂ (K →+* ℂ) + let e : (K →+* ℂ) ≃ Free.ChooseBasisIndex ℤ (𝓞 K) := + equivOfCardEq ((Embeddings.card K ℂ).trans (finrank_eq_card_basis (integralBasis K))) + let M := B.toMatrix (fun i => canonicalEmbedding K (integralBasis K (e i))) + suffices M.det ≠ 0 by + rw [← isUnit_iff_ne_zero, ← Basis.det_apply, ← is_basis_iff_det] at this + refine basisOfLinearIndependentOfCardEqFinrank + ((linearIndependent_equiv e.symm).mpr this.1) ?_ + rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_fintype_fun_eq_card, + Embeddings.card] + -- In order to prove that the determinant is non-zero, we show that it is equal to the + -- square of the discriminant of the integral basis and thus it is not zero + let N := Algebra.embeddingsMatrixReindex ℚ ℂ (fun i => integralBasis K (e i)) + RingHom.equivRatAlgHom + rw [show M = N.transpose by { ext : 2 ; rfl }] + rw [Matrix.det_transpose, ← @pow_ne_zero_iff ℂ _ _ _ 2 (by norm_num)] + convert (map_ne_zero_iff _ (algebraMap ℚ ℂ).injective).mpr + (Algebra.discr_not_zero_of_basis ℚ (integralBasis K)) + rw [← Algebra.discr_reindex ℚ (integralBasis K) e.symm] + exact (Algebra.discr_eq_det_embeddingsMatrixReindex_pow_two ℚ ℂ + (fun i => integralBasis K (e i)) RingHom.equivRatAlgHom).symm + +@[simp] +theorem latticeBasis_apply [NumberField K] (i : Free.ChooseBasisIndex ℤ (𝓞 K)) : + latticeBasis K i = (canonicalEmbedding K) (integralBasis K i) := by + simp only [latticeBasis, integralBasis_apply, coe_basisOfLinearIndependentOfCardEqFinrank, + Function.comp_apply, Equiv.apply_symm_apply] + +theorem mem_span_latticeBasis [NumberField K] (x : (K →+* ℂ) → ℂ) : + x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ canonicalEmbedding K '' (𝓞 K) := by + rw [show Set.range (latticeBasis K) = + (canonicalEmbedding K).toIntAlgHom.toLinearMap '' (Set.range (integralBasis K)) by + rw [← Set.range_comp] ; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] + rw [← Submodule.map_span, ← SetLike.mem_coe, Submodule.map_coe] + rw [show (Submodule.span ℤ (Set.range (integralBasis K)) : Set K) = 𝓞 K by + ext ; exact mem_span_integralBasis K] + rfl end NumberField.canonicalEmbedding + +namespace NumberField.mixedEmbedding + +open NumberField NumberField.InfinitePlace NumberField.ComplexEmbedding + +/-- The ambient space `ℝ^r₁ × ℂ^r₂` with `(r₁, r₂)` the signature of `K`. -/ +local notation "E" K => + ({ w : InfinitePlace K // IsReal w } → ℝ) × ({ w : InfinitePlace K // IsComplex w } → ℂ) + +/-- The canonical embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ +noncomputable def _root_.NumberField.mixedEmbedding : K →+* (E K) := + RingHom.prod (Pi.ringHom fun w => w.prop.embedding) (Pi.ringHom fun w => w.val.embedding) + +instance [NumberField K] : Nontrivial (E K) := by + obtain ⟨w⟩ := (inferInstance : Nonempty (InfinitePlace K)) + obtain hw | hw := w.isReal_or_isComplex + · have : Nonempty {w : InfinitePlace K // IsReal w} := ⟨⟨w, hw⟩⟩ + exact nontrivial_prod_left + · have : Nonempty {w : InfinitePlace K // IsComplex w} := ⟨⟨w, hw⟩⟩ + exact nontrivial_prod_right + +theorem _root_.NumberField.mixedEmbedding_injective [NumberField K] : + Function.Injective (NumberField.mixedEmbedding K) := by + exact RingHom.injective _ + +section comm_map + +/-- The linear map that makes `canonicalEmbedding` and `mixedEmbedding` commute, see +`comm_map_canonical_eq_mixed`. -/ +noncomputable def comm_map : ((K →+* ℂ) → ℂ) →ₗ[ℝ] (E K) := +{ toFun := fun x => ⟨fun w => (x w.val.embedding).re, fun w => x w.val.embedding⟩ + map_add' := by + simp only [Pi.add_apply, Complex.add_re, Prod.mk_add_mk, Prod.mk.injEq] + exact fun _ _ => ⟨rfl, rfl⟩ + map_smul' := by + simp only [Pi.smul_apply, Complex.real_smul, Complex.mul_re, Complex.ofReal_re, + Complex.ofReal_im, zero_mul, sub_zero, RingHom.id_apply, Prod.smul_mk, Prod.mk.injEq] + exact fun _ _ => ⟨rfl, rfl⟩ } + +theorem comm_map_apply_of_isReal (x : (K →+* ℂ) → ℂ) {w : InfinitePlace K} (hw : IsReal w) : + (comm_map K x).1 ⟨w, hw⟩ = (x w.embedding).re := rfl + +theorem comm_map_apply_of_isComplex (x : (K →+* ℂ) → ℂ) {w : InfinitePlace K} (hw : IsComplex w) : + (comm_map K x).2 ⟨w, hw⟩ = x w.embedding := rfl + +@[simp] +theorem comm_map_canonical_eq_mixed (x : K) : + comm_map K (canonicalEmbedding K x) = mixedEmbedding K x := by + simp only [canonicalEmbedding, comm_map, LinearMap.coe_mk, AddHom.coe_mk, Pi.ringHom_apply, + mixedEmbedding, RingHom.prod_apply, Prod.mk.injEq] + exact ⟨rfl, rfl⟩ + +/-- This is technical result to ensure that the image of the `ℂ`-basis of `ℂ^n`, defined in +`canonicalEmbedding.latticeBasis` is a `ℝ`-basis of `ℝ^r₁ × ℂ^r₂`, +see `mixedEmbedding.latticeBasis`. -/ +theorem disjoint_span_comm_map_ker [NumberField K]: + Disjoint (Submodule.span ℝ (Set.range (canonicalEmbedding.latticeBasis K))) + (LinearMap.ker (comm_map K)) := by + refine LinearMap.disjoint_ker.mpr (fun x h_mem h_zero => ?_) + replace h_mem : x ∈ Submodule.span ℝ (Set.range (canonicalEmbedding K)) := by + refine (Submodule.span_mono ?_) h_mem + rintro _ ⟨i, rfl⟩ + exact ⟨integralBasis K i, (canonicalEmbedding.latticeBasis_apply K i).symm⟩ + ext1 φ + rw [Pi.zero_apply] + by_cases hφ : IsReal φ + · rw [show x φ = (x φ).re by + rw [eq_comm, ← Complex.conj_eq_iff_re, canonicalEmbedding.conj_apply _ h_mem, + ComplexEmbedding.isReal_iff.mp hφ], ← Complex.ofReal_zero] + congr + rw [← ComplexEmbedding.IsReal.embedding_mk hφ, ← comm_map_apply_of_isReal K x ⟨φ, hφ, rfl⟩] + exact congrFun (congrArg (fun x => x.1) h_zero) ⟨InfinitePlace.mk φ, _⟩ + · have := congrFun (congrArg (fun x => x.2) h_zero) ⟨InfinitePlace.mk φ, ⟨φ, hφ, rfl⟩⟩ + cases ComplexEmbedding.embedding_mk φ with + | inl h => rwa [← h, ← comm_map_apply_of_isComplex K x ⟨φ, hφ, rfl⟩] + | inr h => + apply RingHom.injective (starRingEnd ℂ) + rwa [canonicalEmbedding.conj_apply _ h_mem, ← h, map_zero, + ← comm_map_apply_of_isComplex K x ⟨φ, hφ, rfl⟩] + +end comm_map + +section integerLattice + +open Module FiniteDimensional + +/-- A `ℝ`-basis of `ℝ^r₁ × ℂ^r₂` that is also a `ℤ`-basis of the image of `𝓞 K`. -/ +noncomputable def latticeBasis [NumberField K] : + Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℝ (E K) := by + classical + -- We construct an `ℝ`-linear independent family from the image of + -- `canonicalEmbedding.lattice_basis` by `comm_map` + have := LinearIndependent.map (LinearIndependent.restrict_scalars + (by { simpa only [Complex.real_smul, mul_one] using Complex.ofReal_injective }) + (canonicalEmbedding.latticeBasis K).linearIndependent) + (disjoint_span_comm_map_ker K) + -- and it's a basis since it has the right cardinality + refine basisOfLinearIndependentOfCardEqFinrank this ?_ + rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_prod, finrank_pi, + finrank_pi_fintype, Complex.finrank_real_complex, Finset.sum_const, Finset.card_univ, + ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, ← card_complex_embeddings, + ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, + Nat.add_sub_of_le (Fintype.card_subtype_le _)] + +@[simp] +theorem latticeBasis_apply [NumberField K] (i : Free.ChooseBasisIndex ℤ (𝓞 K)) : + latticeBasis K i = (mixedEmbedding K) (integralBasis K i) := by + simp only [latticeBasis, coe_basisOfLinearIndependentOfCardEqFinrank, Function.comp_apply, + canonicalEmbedding.latticeBasis_apply, integralBasis_apply, comm_map_canonical_eq_mixed] + +theorem mem_span_latticeBasis [NumberField K] (x : (E K)) : + x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ mixedEmbedding K '' (𝓞 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))] + rw [← Submodule.map_span, ← SetLike.mem_coe, Submodule.map_coe] + rw [show (Submodule.span ℤ (Set.range (integralBasis K)) : Set K) = 𝓞 K by + ext ; exact mem_span_integralBasis K] + rfl + +end integerLattice + +section convex_body + +open Metric ENNReal NNReal + +variable (f : InfinitePlace K → ℝ≥0) + +/-- The convex body defined by `f`: the set of points `x : E` such that `‖x w‖ < f w` for all +infinite place `w`. -/ +def convex_body : Set (E K) := + (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsReal w } => ball 0 (f w))) ×ˢ + (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsComplex w } => ball 0 (f w))) + +theorem convex_body_mem {x : K} : + mixedEmbedding K x ∈ (convex_body K f) ↔ ∀ w : InfinitePlace K, w x < f w := by + simp only [mixedEmbedding, RingHom.prod_apply, convex_body, Set.mem_prod, Set.mem_pi, + Set.mem_univ, Pi.ringHom_apply, mem_ball, dist_zero_right, Real.norm_eq_abs, + IsReal.abs_embedding_apply, forall_true_left, Subtype.forall, Complex.norm_eq_abs, + abs_embedding, ← ball_or_left, ← not_isReal_iff_isComplex, em, forall_true_left] + +theorem convex_body_symmetric (x : E K) (hx : x ∈ (convex_body K f)) : + -x ∈ (convex_body K f) := by + simp only [convex_body, Set.mem_prod, Prod.fst_neg, Set.mem_pi, Set.mem_univ, Pi.neg_apply, + mem_ball, dist_zero_right, norm_neg, Real.norm_eq_abs, forall_true_left, Subtype.forall, + Prod.snd_neg, Complex.norm_eq_abs, hx] at hx ⊢ + exact hx + +theorem convex_body_convex : + Convex ℝ (convex_body K f) := + Convex.prod (convex_pi (fun _ _ => convex_ball _ _)) (convex_pi (fun _ _ => convex_ball _ _)) + +open Classical Fintype MeasureTheory MeasureTheory.Measure BigOperators + +-- See: https://github.com/leanprover/lean4/issues/2220 +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) + +variable [NumberField K] + +/-- The fudge factor that appears in the formula for the volume of `convex_body`. -/ +noncomputable def constant_factor : ℝ≥0∞ := + (2 : ℝ≥0∞) ^ card {w : InfinitePlace K // IsReal w} * + volume (ball (0 : ℂ) 1) ^ card {w : InfinitePlace K // IsComplex w} + +instance : IsAddHaarMeasure (@volume ℂ Complex.measureSpace) := MapLinearEquiv.isAddHaarMeasure _ _ + +theorem constant_factor_pos : 0 < (constant_factor K) := by + refine mul_pos (NeZero.ne _) ?_ + exact ENNReal.pow_ne_zero (ne_of_gt (measure_ball_pos _ _ (by norm_num))) _ + +theorem constant_factor_lt_top : (constant_factor K) < ⊤ := by + refine mul_lt_top ?_ ?_ + · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) + · exact ne_of_lt (pow_lt_top measure_ball_lt_top _) + +set_option maxHeartbeats 400000 in +theorem convex_body_volume : + volume (convex_body K f) = (constant_factor K) * ∏ w, (f w) ^ (mult K w) := by + rw [volume_eq_prod, convex_body, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi] + conv_lhs => + congr ; congr ; next => skip + ext + rw [Real.volume_ball, ofReal_mul (by norm_num), ofReal_coe_nnreal, mul_comm] + conv_lhs => + congr ; next => skip + congr ; next => skip + ext i + rw [addHaar_ball _ _ (by exact (f i).prop), Complex.finrank_real_complex, ← NNReal.coe_pow, + ofReal_coe_nnreal, mul_comm] + rw [Finset.prod_mul_distrib, Finset.prod_mul_distrib, Finset.prod_const, Finset.prod_const, + Finset.card_univ, Finset.card_univ, mul_assoc, mul_comm, ← mul_assoc, mul_assoc, ofReal_ofNat, + ← constant_factor] + simp_rw [mult, pow_ite, pow_one] + rw [Finset.prod_ite] + simp_rw [coe_mul, coe_finset_prod] + simp_rw [show (fun w : InfinitePlace K ↦ ¬IsReal w) = (fun w ↦ IsComplex w) + by funext ; rw [not_isReal_iff_isComplex]] + congr 1 ; rw [mul_comm] ; congr 1 + all_goals + · rw [← Finset.prod_subtype_eq_prod_filter] + congr ; ext + exact ⟨fun _ => Finset.mem_subtype.mpr (Finset.mem_univ _), fun _ => Finset.mem_univ _⟩ + +variable {f} + +/-- This is a technical result: quite often, we want to impose conditions at all infinite places +but one and choose the value at the remaining place so that we can apply +`exists_ne_zero_mem_ring_of_integers_lt`. -/ +theorem adjust_f {w₁ : InfinitePlace K} (B : ℝ≥0) (hf : ∀ w, w ≠ w₁→ f w ≠ 0) : + ∃ g : InfinitePlace K → ℝ≥0, (∀ w, w ≠ w₁ → g w = f w) ∧ ∏ w, (g w) ^ mult K w = B := by + let S := ∏ w in Finset.univ.erase w₁, (f w) ^ mult K w + refine ⟨Function.update f w₁ ((B * S⁻¹) ^ (mult K w₁ : ℝ)⁻¹), ?_, ?_⟩ + · exact fun w hw => Function.update_noteq hw _ f + · rw [← Finset.mul_prod_erase Finset.univ _ (Finset.mem_univ w₁), Function.update_same, + Finset.prod_congr rfl fun w hw => by rw [Function.update_noteq (Finset.ne_of_mem_erase hw)], + ← NNReal.rpow_nat_cast, ← NNReal.rpow_mul, inv_mul_cancel, NNReal.rpow_one, mul_assoc, + inv_mul_cancel, mul_one] + · rw [Finset.prod_ne_zero_iff] + exact fun w hw => pow_ne_zero _ (hf w (Finset.ne_of_mem_erase hw)) + · rw [mult]; split_ifs <;> norm_num + +end convex_body + +section minkowski + +open MeasureTheory MeasureTheory.Measure Classical ENNReal FiniteDimensional + +variable [NumberField K] + +/-- The bound that appears in Minkowski theorem, see +`MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure`. -/ +noncomputable def minkowski_bound : ℝ≥0∞ := + volume (Zspan.fundamentalDomain (latticeBasis K)) * 2 ^ (finrank ℝ (E K)) + +theorem minkowski_bound_lt_top : minkowski_bound K < ⊤ := by + refine mul_lt_top ?_ ?_ + · exact ne_of_lt (Zspan.fundamentalDomain_bounded (latticeBasis K)).measure_lt_top + · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) + +theorem exists_ne_zero_mem_ring_of_integers_lt (h : minkowski_bound K < volume (convex_body K f)) : + ∃ (a : 𝓞 K), a ≠ 0 ∧ ∀ w : InfinitePlace K, w a < f w := by + have : @IsAddHaarMeasure (E K) _ _ _ volume := prod.instIsAddHaarMeasure volume volume + have h_fund := Zspan.isAddFundamentalDomain (latticeBasis K) volume + have : Countable (Submodule.span ℤ (Set.range (latticeBasis K))).toAddSubgroup := by + change Countable (Submodule.span ℤ (Set.range (latticeBasis K)): Set (E K)) + infer_instance + obtain ⟨⟨x, hx⟩, h_nzr, h_mem⟩ := exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure + h_fund h (convex_body_symmetric K f) (convex_body_convex K f) + rw [Submodule.mem_toAddSubgroup, mem_span_latticeBasis] at hx + obtain ⟨a, ha, rfl⟩ := hx + refine ⟨⟨a, ha⟩, ?_, (convex_body_mem K f).mp h_mem⟩ + rw [ne_eq, AddSubgroup.mk_eq_zero_iff, map_eq_zero, ← ne_eq] at h_nzr + exact Subtype.ne_of_val_ne h_nzr + +end minkowski + +end NumberField.mixedEmbedding diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 32a20af73d41e..38b62b6f3d2a5 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -8,7 +8,13 @@ Authors: Xavier Roblot ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ +import Mathlib.Algebra.Module.Zlattice +import Mathlib.GroupTheory.Torsion +import Mathlib.LinearAlgebra.Matrix.ToLinearEquiv +import Mathlib.NumberTheory.NumberField.CanonicalEmbedding import Mathlib.NumberTheory.NumberField.Norm +import Mathlib.RingTheory.Ideal.Norm +import Mathlib.RingTheory.RootsOfUnity.Basic /-! # Units of a number field @@ -22,6 +28,8 @@ field `K`. number field, units -/ +-- See: https://github.com/leanprover/lean4/issues/2220 +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) open scoped NumberField @@ -47,10 +55,368 @@ attribute [local instance] NumberField.ringOfIntegersAlgebra variable {K} -theorem isUnit_iff_norm [NumberField K] (x : 𝓞 K) : +theorem isUnit_iff_norm [NumberField K] {x : 𝓞 K} : IsUnit x ↔ |(RingOfIntegers.norm ℚ x : ℚ)| = 1 := by convert (RingOfIntegers.isUnit_norm ℚ (F := K)).symm rw [← abs_one, abs_eq_abs, ← Rat.RingOfIntegers.isUnit_iff] #align is_unit_iff_norm isUnit_iff_norm end IsUnit + +namespace NumberField.Units + +section coe + +/-- The `MonoidHom` from the group of units `(𝓞 K)ˣ` to the field `K`. -/ +def coe_to_field : (𝓞 K)ˣ →* K := (Units.coeHom K).comp (map (algebraMap (𝓞 K) K)) + +theorem coe_to_field_injective : Function.Injective (coe_to_field K) := + fun _ _ h => Units.eq_iff.mp (SetCoe.ext h) + +/-- There is a natural coercion from `(𝓞 K)ˣ` to `(𝓞 K)` and then from `(𝓞 K)` to `K` but it is +useful to also have a direct one from `(𝓞 K)ˣ` to `K`. -/ +instance : Coe (𝓞 K)ˣ K := ⟨coe_to_field K⟩ + +theorem ext {x y : (𝓞 K)ˣ} : x = y ↔ (x : K) = (y : K) := (coe_to_field_injective K).eq_iff.symm + +theorem ne_zero (x : (𝓞 K)ˣ) : (x : K) ≠ 0 := + Subtype.coe_injective.ne_iff.mpr (_root_.Units.ne_zero x) + +end coe + +open NumberField.InfinitePlace + +section torsion + +/-- The torsion subgroup of the group of units. -/ +def torsion : Subgroup (𝓞 K)ˣ := CommGroup.torsion (𝓞 K)ˣ + +theorem mem_torsion {x : (𝓞 K)ˣ} [NumberField K] : + x ∈ torsion K ↔ ∀ w : InfinitePlace K, w x = 1 := by + rw [eq_iff_eq (x : K) 1, torsion, CommGroup.mem_torsion, isOfFinOrder_iff_pow_eq_one] + refine ⟨fun ⟨n, h_pos, h_eq⟩ φ => ?_, fun h => ?_⟩ + · refine norm_map_one_of_pow_eq_one φ.toMonoidHom (k := ⟨n, h_pos⟩) ?_ + rw [PNat.mk_coe, ← map_pow, h_eq, map_one] + · obtain ⟨n, hn, hx⟩ := Embeddings.pow_eq_one_of_norm_eq_one K ℂ x.val.prop h + exact ⟨n, hn, by rwa [ext, map_pow, map_one]⟩ +end torsion + +instance : Nonempty (torsion K) := ⟨1⟩ + +/-- The torsion subgroup is finite. -/ +instance [NumberField K] : Fintype (torsion K) := by + refine @Fintype.ofFinite _ (Set.finite_coe_iff.mpr ?_) + refine Set.Finite.of_finite_image ?_ ((coe_to_field_injective K).injOn _) + refine (Embeddings.finite_of_norm_le K ℂ 1).subset + (fun a ⟨u, ⟨h_tors, h_ua⟩⟩ => ⟨?_, fun φ => ?_⟩) + · rw [← h_ua] + exact u.val.prop + · rw [← h_ua] + exact le_of_eq ((eq_iff_eq _ 1).mp ((mem_torsion K).mp h_tors) φ) + +/-- The torsion subgroup is cylic. -/ +instance [NumberField K] : IsCyclic (torsion K) := subgroup_units_cyclic _ + +/-- The order of the torsion subgroup as positive integer. -/ +def torsion_order [NumberField K] : ℕ+ := ⟨Fintype.card (torsion K), Fintype.card_pos⟩ + +namespace dirichlet +-- This section is devoted to the proof of Dirichlet's unit theorem +-- We define a group morphism from `(𝓞 K)ˣ` to `{w : InfinitePlace K // w ≠ w₀} → ℝ` where `w₀` +-- is a distinguished (arbitrary) infinite place, prove that its kernel is the torsion subgroup +-- (see `log_embedding_eq_zero_iff`) and that its image, called `unit_lattice`, is a full +-- `ℤ`-lattice. It follows that is a free `ℤ`-module (see `unit_lattice_moduleFree `) of +-- rank `card (InfinitePlaces K) - 1` (see `unit_lattice_rank`). +-- To prove that the `unit_lattice` is a full `ℤ`-lattice, we need to prove that it is discrete +-- (see `unit_lattice_inter_ball_finite`) and that it spans the full space over `ℝ` +-- (see ` unit_lattice_span_eq_top`); this is the main part of the proof, see the section +-- `span_top` below for more details. + +open scoped Classical BigOperators + +variable [NumberField K] + +variable {K} + +/-- The distinguished infinite place. -/ +def w₀ : InfinitePlace K := (inferInstance : Nonempty (InfinitePlace K)).some + +variable (K) + +/-- The logarithmic embedding of the units (seen as an `Additive` group). -/ +def log_embedding : Additive ((𝓞 K)ˣ) →+ ({w : InfinitePlace K // w ≠ w₀} → ℝ) := +{ toFun := fun x w => mult K w * Real.log (w.val (Additive.toMul x)) + map_zero' := by simp; rfl + map_add' := by + intro _ _ + simp [ne_eq, toMul_add, map_mul, map_eq_zero, ne_zero, Real.log_mul, mul_add] + rfl } + +@[simp] +theorem log_embedding_component (x : (𝓞 K)ˣ) (w : {w : InfinitePlace K // w ≠ w₀}) : + (log_embedding K x) w = mult K w * Real.log (w.val x) := rfl + +theorem log_embedding_sum_component (x : (𝓞 K)ˣ) : + ∑ w, log_embedding K x w = - mult K w₀ * Real.log (w₀ (x : K)) := by + have h := congrArg Real.log (prod_mult_eq_abs_norm K x) + rw [show |(Algebra.norm ℚ) (x : K)| = 1 from isUnit_iff_norm.mp x.isUnit, Rat.cast_one, + Real.log_one, Real.log_prod] at h + · simp_rw [Real.log_pow] at h + rw [← Finset.insert_erase (Finset.mem_univ w₀), Finset.sum_insert (Finset.not_mem_erase w₀ + Finset.univ), add_comm, add_eq_zero_iff_eq_neg] at h + convert h using 1 + · refine (Finset.sum_subtype _ (fun w => ?_) (fun w => (mult K w) * (Real.log (w ↑x)))).symm + exact ⟨Finset.ne_of_mem_erase, fun h => Finset.mem_erase_of_ne_of_mem h (Finset.mem_univ w)⟩ + · norm_num + · exact fun w _ => pow_ne_zero _ (AbsoluteValue.ne_zero _ (ne_zero K x)) + +theorem mult_log_place_eq_zero {x : (𝓞 K)ˣ} {w : InfinitePlace K} : + mult K w * Real.log (w x) = 0 ↔ w.val x = 1 := by + rw [mul_eq_zero, or_iff_right, Real.log_eq_zero, or_iff_right, or_iff_left] + · have : 0 ≤ w.val x := AbsoluteValue.nonneg _ _ + linarith + · simp only [ne_eq, map_eq_zero, ne_zero K x] + · refine (ne_of_gt ?_) + rw [mult]; split_ifs <;> norm_num + +theorem log_embedding_eq_zero_iff (x : (𝓞 K)ˣ) : + log_embedding K x = 0 ↔ x ∈ torsion K := by + rw [mem_torsion] + refine ⟨fun h w => ?_, fun h => ?_⟩ + · by_cases hw : w = w₀ + · suffices - mult K w₀ * Real.log (w₀ (x : K)) = 0 by + rw [neg_mul, neg_eq_zero, ← hw] at this + exact (mult_log_place_eq_zero K).mp this + rw [← log_embedding_sum_component, Finset.sum_eq_zero] + exact fun w _ => congrFun h w + · exact (mult_log_place_eq_zero K).mp (congrFun h ⟨w, hw⟩) + · ext w + rw [log_embedding_component, h w.val, Real.log_one, mul_zero, Pi.zero_apply] + +/-- The lattice formed by the image of the logarithmic embedding. -/ +noncomputable def unit_lattice : AddSubgroup ({w : InfinitePlace K // w ≠ w₀} → ℝ) := + AddSubgroup.map (log_embedding K) ⊤ + +theorem log_embedding_component_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖log_embedding K x‖ ≤ r) + (w : {w : InfinitePlace K // w ≠ w₀}) : |log_embedding K x w| ≤ r := by + lift r to NNReal using hr + simp_rw [Pi.norm_def, NNReal.coe_le_coe, Finset.sup_le_iff, ← NNReal.coe_le_coe] at h + exact h w (Finset.mem_univ _) + +theorem log_le_of_log_embedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖log_embedding K x‖ ≤ r) + (w : InfinitePlace K) : |Real.log (w x)| ≤ (Fintype.card (InfinitePlace K)) * r := by + have tool : ∀ x : ℝ, 0 ≤ x → x ≤ mult K w * x := fun x hx => by + nth_rw 1 [← one_mul x] + refine mul_le_mul ?_ le_rfl hx ?_ + all_goals { rw [mult]; split_ifs <;> norm_num } + by_cases hw : w = w₀ + · have hyp := congrArg (‖·‖) (log_embedding_sum_component K x).symm + replace hyp := (le_of_eq hyp).trans (norm_sum_le _ _) + simp_rw [norm_mul, norm_neg, Real.norm_eq_abs, Nat.abs_cast] at hyp + refine (le_trans ?_ hyp).trans ?_ + · rw [← hw] + exact tool _ (abs_nonneg _) + · refine (Finset.sum_le_card_nsmul Finset.univ _ _ + (fun w _ => log_embedding_component_le K hr h w)).trans ?_ + rw [nsmul_eq_mul] + refine mul_le_mul ?_ le_rfl hr (Fintype.card (InfinitePlace K)).cast_nonneg + simp [Finset.card_univ] + · have hyp := log_embedding_component_le K hr h ⟨w, hw⟩ + rw [log_embedding_component, abs_mul, Nat.abs_cast] at hyp + refine (le_trans ?_ hyp).trans ?_ + · exact tool _ (abs_nonneg _) + · nth_rw 1 [← one_mul r] + exact mul_le_mul (Nat.one_le_cast.mpr Fintype.card_pos) (le_of_eq rfl) hr (Nat.cast_nonneg _) + +theorem unit_lattice_inter_ball_finite (r : ℝ) : + ((unit_lattice K : Set ({ w : InfinitePlace K // w ≠ w₀} → ℝ)) ∩ + Metric.closedBall 0 r).Finite := by + obtain hr | hr := lt_or_le r 0 + · convert Set.finite_empty + rw [Metric.closedBall_eq_empty.mpr hr] + exact Set.inter_empty _ + · suffices Set.Finite {x : (𝓞 K)ˣ | IsIntegral ℤ (x : K) ∧ + ∀ (φ : K →+* ℂ), ‖φ x‖ ≤ Real.exp ((Fintype.card (InfinitePlace K)) * r)} by + refine (Set.Finite.image (log_embedding K) this).subset ?_ + rintro _ ⟨⟨x, ⟨_, rfl⟩⟩, hx⟩ + refine ⟨x, ⟨x.val.prop, (le_iff_le _ _).mp (fun w => (Real.log_le_iff_le_exp ?_).mp ?_)⟩, rfl⟩ + · exact pos_iff.mpr (ne_zero K x) + · rw [mem_closedBall_zero_iff] at hx + exact (le_abs_self _).trans (log_le_of_log_embedding_le K hr hx w) + refine Set.Finite.of_finite_image ?_ ((coe_to_field_injective K).injOn _) + refine (Embeddings.finite_of_norm_le K ℂ + (Real.exp ((Fintype.card (InfinitePlace K)) * r))).subset ?_ + rintro _ ⟨x, ⟨⟨h_int, h_le⟩, rfl⟩⟩ + exact ⟨h_int, h_le⟩ + +section span_top +-- To prove that the span over `ℝ` of the `unit_lattice` is equal to the full space, we construct +-- for each infinite place `w₁ ≠ w₀` an unit `u_w₁` of `K` such that, for all infinite place +-- `w` such that `w ≠ w₁`, we have `Real.log w (u_w₁) < 0` (and thus `Real.log w₁ (u_w₁) > 0`). +-- It follows then from a determinant computation (using `Matrix.det_ne_zero_of_neg`) that the +-- image by `log_embedding` of these units is a `ℝ`-linearly independent family. +-- The unit `u_w₁` is obtained by construction a sequence `seq n` of nonzero algebraic integers +-- that is strictly decreasing at infinite places distinct from `w₁` and of bounded norms. Since +-- there are finitely many ideals of bounded norms, there exists two terms in the sequence defining +-- the same ideal and their quotient is the unit `u_w₁` (see `exists_unit`). + +open NumberField.mixedEmbedding NNReal + +variable (w₁ : InfinitePlace K) {B : ℕ} (hB : minkowski_bound K < (constant_factor K) * B) + +/-- This result shows that there always exists a next term of the sequence. -/ +theorem seq.next {x : 𝓞 K} (hx : x ≠ 0) : + ∃ y : 𝓞 K, y ≠ 0 ∧ (∀ w, w ≠ w₁ → w y < w x) ∧ |Algebra.norm ℚ (y : K)| ≤ B := by + let f : InfinitePlace K → ℝ≥0 := + fun w => ⟨(w x) / 2, div_nonneg (AbsoluteValue.nonneg _ _) (by norm_num)⟩ + suffices ∀ w, w ≠ w₁ → f w ≠ 0 by + obtain ⟨g, h_geqf, h_gprod⟩ := adjust_f K B this + obtain ⟨y, h_ynz, h_yle⟩ := exists_ne_zero_mem_ring_of_integers_lt (f := g) + (by rw [convex_body_volume]; convert hB; exact congrArg ((↑): NNReal → ENNReal) h_gprod) + refine ⟨y, h_ynz, fun w hw => (h_geqf w hw ▸ h_yle w).trans ?_, ?_⟩ + · rw [← Rat.cast_le (K := ℝ), Rat.cast_coe_nat] + calc + _ = ∏ w : InfinitePlace K, w y ^ mult K w := (prod_mult_eq_abs_norm K y).symm + _ ≤ ∏ w : InfinitePlace K, (g w : ℝ) ^ mult K w := ?_ + _ ≤ (B : ℝ) := ?_ + · refine Finset.prod_le_prod ?_ ?_ + exact fun _ _ => pow_nonneg (by positivity) _ + exact fun w _ => pow_le_pow_of_le_left (by positivity) (le_of_lt (h_yle w)) (mult K w) + · simp_rw [← coe_pow, ← NNReal.coe_prod] + exact le_of_eq (congrArg toReal h_gprod) + · refine div_lt_self ?_ (by norm_num) + simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, hx] + intro _ _ + rw [ne_eq, Nonneg.mk_eq_zero, div_eq_zero_iff, map_eq_zero, not_or, ZeroMemClass.coe_eq_zero] + exact ⟨hx, by norm_num⟩ + +/-- An infinite sequence of nonzero algebraic integers of `K` satisfying the following properties: +• `seq n` is nonzero; +• for `w : InfinitePlace K`, `w ≠ w₁ → w (seq n+1) < w (seq n)`; +• `∣norm (seq n)∣ ≤ B`. -/ +def seq : ℕ → { x : 𝓞 K // x ≠ 0 } + | 0 => ⟨1, by norm_num⟩ + | n + 1 => + ⟨(seq.next K w₁ hB (seq n).prop).choose, (seq.next K w₁ hB (seq n).prop).choose_spec.1⟩ + +/-- The terms of the sequence are nonzero. -/ +theorem seq.ne_zero (n : ℕ) : (seq K w₁ hB n : K) ≠ 0 := by + refine (map_ne_zero_iff (algebraMap (𝓞 K) K) ?_).mpr (seq K w₁ hB n).prop + exact IsFractionRing.injective { x // x ∈ 𝓞 K } K + +/-- The sequence is strictly decreasing at infinite places different from `w₁`. -/ +theorem seq.antitone {n m : ℕ} (h : n < m) : + ∀ w : InfinitePlace K, w ≠ w₁ → w (seq K w₁ hB m) < w (seq K w₁ hB n) := by + induction m with + | zero => + exfalso + exact Nat.not_succ_le_zero n h + | succ m m_ih => + intro w hw + cases eq_or_lt_of_le (Nat.le_of_lt_succ h) with + | inl hr => + rw [hr] + exact (seq.next K w₁ hB (seq K w₁ hB m).prop).choose_spec.2.1 w hw + | inr hr => + refine lt_trans ?_ (m_ih hr w hw) + exact (seq.next K w₁ hB (seq K w₁ hB m).prop).choose_spec.2.1 w hw + +/-- The terms of the sequence have bounded norms. -/ +theorem seq.norm_bdd (n : ℕ) : + 1 ≤ Int.natAbs (Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K)) ∧ + Int.natAbs (Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K)) ≤ B := by + cases n with + | zero => + have : 1 ≤ B := by + contrapose! hB + simp only [Nat.lt_one_iff.mp hB, CharP.cast_eq_zero, mul_zero, zero_le] + simp only [ne_eq, seq, map_one, Int.natAbs_one, le_refl, this, and_self] + | succ n => + refine ⟨Nat.succ_le_iff.mpr (Int.natAbs_pos.mpr ?_), ?_⟩ + · exact Algebra.norm_ne_zero_iff.mpr (seq K w₁ hB n.succ).prop + · rw [← Nat.cast_le (α := ℚ), Int.cast_natAbs, Int.cast_abs] + change |algebraMap ℤ ℚ _| ≤ _ + rw [← Algebra.norm_localization ℤ (Sₘ := K) (nonZeroDivisors ℤ)] + exact (seq.next K w₁ hB (seq K w₁ hB n).prop).choose_spec.2.2 + +/-- Construct an unit associated to the place `w₁`. The family, for `w₁ ≠ w₀`, formed by the +image by the `log_embedding` of these units is `ℝ`-linearly independent, see +`unit_lattice_span_eq_top`. -/ +theorem exists_unit (w₁ : InfinitePlace K ) : + ∃ u : (𝓞 K)ˣ, (∀ w : InfinitePlace K, w ≠ w₁ → Real.log (w u) < 0) := by + obtain ⟨B, hB⟩ : ∃ B : ℕ, minkowski_bound K < (constant_factor K) * B := by + simp_rw [mul_comm] + refine ENNReal.exists_nat_mul_gt ?_ ?_ + exact ne_of_gt (constant_factor_pos K) + exact ne_of_lt (minkowski_bound_lt_top 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 ?_ ?_⟩ + · simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, ne_zero] + · calc + _ = w ((seq K w₁ hB m : K) * (seq K w₁ hB n : K)⁻¹) := ?_ + _ = w (seq K w₁ hB m) * w (seq K w₁ hB n)⁻¹ := map_mul _ _ _ + _ < 1 := ?_ + · rw [← congrArg ((↑) : (𝓞 K) → K) hu.choose_spec, mul_comm, Submonoid.coe_mul, ← mul_assoc, + inv_mul_cancel (seq.ne_zero K w₁ hB n), one_mul] + rfl + · rw [map_inv₀, mul_inv_lt_iff (pos_iff.mpr (seq.ne_zero K w₁ hB n)), mul_one] + exact seq.antitone 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] + exact seq.norm_bdd 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) + +theorem unit_lattice_span_eq_top : + Submodule.span ℝ (unit_lattice K : Set ({w : InfinitePlace K // w ≠ w₀} → ℝ)) = ⊤ := by + refine le_antisymm (le_top) ?_ + -- The standard basis + let B := Pi.basisFun ℝ {w : InfinitePlace K // w ≠ w₀} + -- The family of units constructed above + let v := fun w : { w : InfinitePlace K // w ≠ w₀ } => log_embedding K ((exists_unit K w).choose) + -- To prove the result, it is enough to prove that the family `v` is linearly independent + suffices B.det v ≠ 0 by + rw [← isUnit_iff_ne_zero, ← is_basis_iff_det] at this + rw [← this.2] + exact Submodule.span_monotone (fun _ ⟨w, hw⟩ => + ⟨(exists_unit K w).choose, trivial, by rw [← hw]⟩) + rw [Basis.det_apply] + -- We use a specific lemma to prove that this determinant is nonzero + refine Matrix.det_ne_zero_of_neg (fun i j hij => ?_) (fun j => ?_) + · rw [Basis.coePiBasisFun.toMatrix_eq_transpose, Matrix.transpose_apply] + refine mul_neg_of_pos_of_neg ?_ ((exists_unit K j).choose_spec i ?_) + · rw [mult]; split_ifs <;> norm_num + · exact Subtype.ext_iff_val.not.mp hij + · simp_rw [Basis.coePiBasisFun.toMatrix_eq_transpose, Matrix.transpose_apply, + log_embedding_sum_component] + refine mul_pos_of_neg_of_neg ?_ ?_ + · rw [mult]; split_ifs <;> norm_num + · exact (exists_unit K j).choose_spec _ j.prop.symm + +end span_top + +/-- The unit rank of the number field `K`, it is equal to `card (InfinitePlace K) - 1`. -/ +def _root_.NumberField.Units.rank : ℕ := Fintype.card (InfinitePlace K) - 1 + +open FiniteDimensional + +theorem rank_space : + finrank ℝ ({w : InfinitePlace K // w ≠ w₀} → ℝ) = rank K := by + simp only [finrank_fintype_fun_eq_card, Fintype.card_subtype_compl, + Fintype.card_ofSubsingleton, rank] + +theorem unit_lattice_moduleFree : Module.Free ℤ (unit_lattice K) := +Zlattice.module_free ℝ ((unit_lattice_inter_ball_finite K)) (unit_lattice_span_eq_top K) + +theorem unit_lattice.rank : finrank ℤ (unit_lattice K) = rank K := by + rw [← rank_space] + exact Zlattice.rank ℝ ((unit_lattice_inter_ball_finite K)) (unit_lattice_span_eq_top K) + +end dirichlet + +end NumberField.Units From 07325e1ad53fe8dd3cbf33102496e3356a969e06 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 18 Jul 2023 08:06:03 +0900 Subject: [PATCH 25/67] Update --- Mathlib/NumberTheory/NumberField/Basic.lean | 15 +++ .../NumberTheory/NumberField/Embeddings.lean | 106 ++++++++++++++---- 2 files changed, 100 insertions(+), 21 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Basic.lean b/Mathlib/NumberTheory/NumberField/Basic.lean index e4227695d2345..03f3233e0c7f8 100644 --- a/Mathlib/NumberTheory/NumberField/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/Basic.lean @@ -171,6 +171,21 @@ theorem integralBasis_apply (i : Free.ChooseBasisIndex ℤ (𝓞 K)) : Basis.localizationLocalization_apply ℚ (nonZeroDivisors ℤ) K (RingOfIntegers.basis K) i #align number_field.integral_basis_apply NumberField.integralBasis_apply +theorem mem_span_integralBasis {x : K} : + x ∈ Submodule.span ℤ (Set.range (integralBasis K)) ↔ x ∈ 𝓞 K := by + suffices Submodule.span ℤ (Set.range (integralBasis K)) = + AddSubgroup.toIntSubmodule (𝓞 K).toSubring.toAddSubgroup by rw [this]; rfl + convert congrArg + (Submodule.map (Submodule.subtype (AddSubgroup.toIntSubmodule (𝓞 K).toSubring.toAddSubgroup))) + (Basis.span_eq (RingOfIntegers.basis K)) + · rw [Submodule.map_span] + congr + ext + simp only [Set.mem_range, integralBasis_apply, Submodule.coeSubtype, Set.mem_image, + exists_exists_eq_and] + rfl + · simp only [Submodule.map_top, Submodule.range_subtype] + theorem RingOfIntegers.rank : FiniteDimensional.finrank ℤ (𝓞 K) = FiniteDimensional.finrank ℚ K := IsIntegralClosure.rank ℤ ℚ K (𝓞 K) #align number_field.ring_of_integers.rank NumberField.RingOfIntegers.rank diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 20a94630dcba9..8ce23ea99c3b6 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -182,6 +182,9 @@ def IsReal (φ : K →+* ℂ) : Prop := theorem isReal_iff {φ : K →+* ℂ} : IsReal φ ↔ conjugate φ = φ := isSelfAdjoint_iff #align number_field.complex_embedding.is_real_iff NumberField.ComplexEmbedding.isReal_iff +theorem isReal_iff_isReal_conj {φ : K →+* ℂ} : + IsReal φ ↔ IsReal (conjugate φ) := by rw [IsReal, IsReal, IsSelfAdjoint.star_iff] + /-- A real embedding as a ring homomorphism from `K` to `ℝ` . -/ def IsReal.embedding {φ : K →+* ℂ} (hφ : IsReal φ) : K →+* ℝ where toFun x := (φ x).re @@ -352,16 +355,21 @@ def IsComplex (w : InfinitePlace K) : Prop := #align number_field.infinite_place.is_complex NumberField.InfinitePlace.IsComplex @[simp] -theorem _root_.NumberField.ComplexEmbeddings.IsReal.embedding_mk {φ : K →+* ℂ} +theorem _root_.NumberField.ComplexEmbedding.IsReal.embedding_mk {φ : K →+* ℂ} (h : ComplexEmbedding.IsReal φ) : embedding (mk φ) = φ := by have := mk_eq_iff.mp (mk_embedding (mk φ)).symm rwa [ComplexEmbedding.isReal_iff.mp h, or_self_iff, eq_comm] at this -#align number_field.complex_embeddings.is_real.embedding_mk NumberField.ComplexEmbeddings.IsReal.embedding_mk +#align number_field.complex_embeddings.is_real.embedding_mk NumberField.ComplexEmbedding.IsReal.embedding_mk + +@[simp] +theorem _root_.NumberField.ComplexEmbedding.embedding_mk (φ : K →+* ℂ) : + embedding (mk φ) = φ ∨ embedding (mk φ) = ComplexEmbedding.conjugate φ := by + rw [@eq_comm _ _ φ, @eq_comm _ _ (ComplexEmbedding.conjugate φ), ← mk_eq_iff, mk_embedding] theorem isReal_iff {w : InfinitePlace K} : IsReal w ↔ ComplexEmbedding.IsReal (embedding w) := by constructor · rintro ⟨φ, ⟨hφ, rfl⟩⟩ - rwa [_root_.NumberField.ComplexEmbeddings.IsReal.embedding_mk hφ] + rwa [_root_.NumberField.ComplexEmbedding.IsReal.embedding_mk hφ] · exact fun h => ⟨embedding w, h, mk_embedding w⟩ #align number_field.infinite_place.is_real_iff NumberField.InfinitePlace.isReal_iff @@ -377,11 +385,24 @@ theorem isComplex_iff {w : InfinitePlace K} : · exact fun h => ⟨embedding w, h, mk_embedding w⟩ #align number_field.infinite_place.is_complex_iff NumberField.InfinitePlace.isComplex_iff +theorem _root_.NumberField.ComplexEmbedding.isReal_mk_iff {φ : K →+* ℂ} : + IsReal (mk φ) ↔ ComplexEmbedding.IsReal φ := by + refine ⟨fun h => ?_, fun h => ⟨φ, h, rfl⟩⟩ + · rw [isReal_iff] at h + cases ComplexEmbedding.embedding_mk φ with + | inl hφ => rwa [← hφ] + | inr hφ => rwa [ComplexEmbedding.isReal_iff_isReal_conj, ← hφ] + @[simp] theorem not_isReal_iff_isComplex {w : InfinitePlace K} : ¬IsReal w ↔ IsComplex w := by rw [isComplex_iff, isReal_iff] #align number_field.infinite_place.not_is_real_iff_is_complex NumberField.InfinitePlace.not_isReal_iff_isComplex +theorem _root_.NumberField.ComplexEmbedding.not_isReal_mk_iff {φ : K →+* ℂ} : + ¬ ComplexEmbedding.IsReal φ ↔ IsComplex (mk φ) := by + convert (ComplexEmbedding.isReal_mk_iff (φ := φ)).symm.not + exact not_isReal_iff_isComplex.symm + @[simp] theorem not_isComplex_iff_isReal {w : InfinitePlace K} : ¬IsComplex w ↔ IsReal w := by rw [← not_isReal_iff_isComplex, Classical.not_not] @@ -411,12 +432,18 @@ theorem IsReal.abs_embedding_apply {w : InfinitePlace K} (hw : IsReal w) (x : K) variable (K) +/-- The multiplicity of an infinite place, that is the number of distinct complex embeddings that +defines it, see ` mult_eq`. -/ +noncomputable def mult (w : InfinitePlace K) : ℕ := if (IsReal w) then 1 else 2 + +variable {K} + /-- The map from real embeddings to real infinite places as an equiv -/ noncomputable def mkReal : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } ≃ { w : InfinitePlace K // IsReal w } where toFun := Subtype.map mk fun φ hφ => ⟨φ, hφ, rfl⟩ invFun w := ⟨w.1.embedding, isReal_iff.1 w.2⟩ - left_inv φ := Subtype.ext_iff.2 (NumberField.ComplexEmbeddings.IsReal.embedding_mk φ.2) + left_inv φ := Subtype.ext_iff.2 (NumberField.ComplexEmbedding.IsReal.embedding_mk φ.2) right_inv w := Subtype.ext_iff.2 (mk_embedding w.1) #align number_field.infinite_place.mk_real NumberField.InfinitePlace.mkReal @@ -426,38 +453,38 @@ noncomputable def mkComplex : Subtype.map mk fun φ hφ => ⟨φ, hφ, rfl⟩ #align number_field.infinite_place.mk_complex NumberField.InfinitePlace.mkComplex +variable (K) + theorem mkComplex_embedding (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) : - (mkComplex K φ : InfinitePlace K).embedding = φ ∨ - (mkComplex K φ : InfinitePlace K).embedding = ComplexEmbedding.conjugate φ := by - rw [@eq_comm _ _ φ.val, @eq_comm _ _ (ComplexEmbedding.conjugate φ.val), ← mk_eq_iff, - mk_embedding] - rfl + (mkComplex φ : InfinitePlace K).embedding = φ ∨ + (mkComplex φ : InfinitePlace K).embedding = ComplexEmbedding.conjugate φ := + NumberField.ComplexEmbedding.embedding_mk φ.1 #align number_field.infinite_place.mk_complex_embedding NumberField.InfinitePlace.mkComplex_embedding @[simp] theorem mkReal_coe (φ : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ }) : - (mkReal K φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl + (mkReal φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl #align number_field.infinite_place.mk_real_coe NumberField.InfinitePlace.mkReal_coe @[simp] theorem mkComplex_coe (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) : - (mkComplex K φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl + (mkComplex φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl #align number_field.infinite_place.mk_complex_coe NumberField.InfinitePlace.mkComplex_coe -- @[simp] -- Porting note: not in simpNF theorem mkReal.apply (φ : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ }) (x : K) : - (mkReal K φ).val x = Complex.abs (φ.val x) := rfl + (mkReal φ).val x = Complex.abs (φ.val x) := rfl #align number_field.infinite_place.mk_real.apply NumberField.InfinitePlace.mkReal.apply -- @[simp] -- Porting note: not in simpNF theorem mkComplex.apply (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) (x : K) : - (mkComplex K φ).val x = Complex.abs (φ.val x) := rfl + (mkComplex φ).val x = Complex.abs (φ.val x) := rfl #align number_field.infinite_place.mk_complex.apply NumberField.InfinitePlace.mkComplex.apply variable [NumberField K] theorem mkComplex.filter (w : { w : InfinitePlace K // w.IsComplex }) : - (Finset.univ.filter fun φ => mkComplex K φ = w) = + (Finset.univ.filter fun φ => mkComplex φ = w) = {⟨w.1.embedding, isComplex_iff.1 w.2⟩, ⟨ComplexEmbedding.conjugate w.1.embedding, ComplexEmbedding.isReal_conjugate_iff.not.2 (isComplex_iff.1 w.2)⟩} := by @@ -469,22 +496,52 @@ theorem mkComplex.filter (w : { w : InfinitePlace K // w.IsComplex }) : #align number_field.infinite_place.mk_complex.filter NumberField.InfinitePlace.mkComplex.filter theorem mkComplex.filter_card (w : { w : InfinitePlace K // w.IsComplex }) : - (Finset.univ.filter fun φ => mkComplex K φ = w).card = 2 := by + (Finset.univ.filter fun φ => mkComplex φ = w).card = 2 := by rw [mkComplex.filter] exact Finset.card_doubleton (Subtype.mk_eq_mk.not.2 <| ne_comm.1 <| ComplexEmbedding.isReal_iff.not.1 <| isComplex_iff.1 w.2) #align number_field.infinite_place.mk_complex.filter_card NumberField.InfinitePlace.mkComplex.filter_card +-- TODO. clean up this proof +theorem mult_eq (w : InfinitePlace K) : + mult K w = (Finset.univ.filter fun φ => mk φ = w).card := by + rw [mult] + split_ifs with hw + · rw [eq_comm, Finset.card_eq_one] + refine ⟨w.embedding, ?_⟩ + rw [← Finset.coe_eq_singleton, Finset.coe_filter] + suffices ∀ φ, mk φ = w ↔ φ = w.embedding by simp [this] + refine fun φ => ⟨fun h => ?_, fun h => by rw [h]; exact mk_embedding w⟩ + rw [← h, eq_comm] + apply ComplexEmbedding.IsReal.embedding_mk + rwa [← ComplexEmbedding.isReal_mk_iff, h] + · rw [not_isReal_iff_isComplex] at hw + rw [← mkComplex.filter_card K ⟨w, hw⟩] + refine Finset.card_congr (fun φ _ => φ.val) (fun φ hφ => ?_) (fun _ _ _ _ h => SetCoe.ext h) + (fun φ hφ => ⟨⟨φ, ?_⟩, ⟨?_, rfl⟩⟩) + · simp only [Finset.mem_univ, forall_true_left, Finset.mem_filter, true_and, ← mkComplex_coe] + simp only [Finset.mem_univ, forall_true_left, Subtype.forall, Finset.mem_filter, true_and] + at hφ + exact congrArg Subtype.val hφ + · simp at hφ + rw [← hφ] at hw + rw [← ComplexEmbedding.not_isReal_mk_iff] at hw + exact hw + · simp + simp at hφ + rwa [Subtype.ext_iff, mkComplex_coe] + noncomputable instance NumberField.InfinitePlace.fintype : Fintype (InfinitePlace K) := Set.fintypeRange _ #align number_field.infinite_place.number_field.infinite_place.fintype NumberField.InfinitePlace.NumberField.InfinitePlace.fintype +open scoped BigOperators + /-- The infinite part of the product formula : for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where `‖·‖_w` is the normalized absolute value for `w`. -/ theorem prod_eq_abs_norm (x : K) : - (Finset.univ.prod fun w : InfinitePlace K => ite w.IsReal (w x) (w x ^ 2)) = - abs (Algebra.norm ℚ x) := by + ∏ w : InfinitePlace K, ite w.IsReal (w x) (w x ^ 2) = abs (Algebra.norm ℚ x) := by convert (congr_arg Complex.abs (@Algebra.norm_eq_prod_embeddings ℚ _ _ _ _ ℂ _ _ _ _ _ x)).symm · rw [map_prod, ← Equiv.prod_comp' RingHom.equivRatAlgHom (fun f => Complex.abs (f x)) (fun φ => Complex.abs (φ x)) fun _ => by simp [RingHom.equivRatAlgHom_apply]; rfl] @@ -496,13 +553,13 @@ theorem prod_eq_abs_norm (x : K) : rw [Finset.prod_ite, Finset.prod_ite] refine congr (congr_arg Mul.mul ?_) ?_ · rw [← Finset.prod_subtype_eq_prod_filter, ← Finset.prod_subtype_eq_prod_filter] - convert (Equiv.prod_comp' (mkReal K) (fun φ => Complex.abs (φ.val x)) + convert (Equiv.prod_comp' mkReal (fun φ => Complex.abs (φ.val x)) (fun w => w.val x) _).symm any_goals ext; simp only [Finset.mem_subtype, Finset.mem_univ] exact fun φ => mkReal.apply K φ x · rw [Finset.filter_congr fun (w : InfinitePlace K) _ => @not_isReal_iff_isComplex K _ w, ← Finset.prod_subtype_eq_prod_filter, ← Finset.prod_subtype_eq_prod_filter] - convert Finset.prod_fiberwise Finset.univ (fun φ => mkComplex K φ) + convert Finset.prod_fiberwise Finset.univ (fun φ => mkComplex φ) (fun φ => Complex.abs (φ.val x)) using 2 with w · ext; simp only [Finset.mem_subtype, Finset.mem_univ, not_isReal_iff_isComplex] · rw [@Finset.prod_congr _ _ _ _ _ (fun _ => w.val x) _ (Eq.refl _) fun φ hφ => @@ -512,11 +569,18 @@ theorem prod_eq_abs_norm (x : K) : · rw [eq_ratCast, Rat.cast_abs, ← Complex.abs_ofReal, Complex.ofReal_rat_cast] #align number_field.infinite_place.prod_eq_abs_norm NumberField.InfinitePlace.prod_eq_abs_norm +/-- A version of `prod_eq_abs_norm` that uses `mult`. -/ +theorem prod_mult_eq_abs_norm (x : K) : + ∏ w : InfinitePlace K, w x ^ mult K w = abs (Algebra.norm ℚ x) := by + convert prod_eq_abs_norm K x using 2 + rw [mult] + split_ifs <;> simp + open Fintype theorem card_real_embeddings : card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } = card { w : InfinitePlace K // IsReal w } := - by convert (Fintype.ofEquiv_card (mkReal K)).symm + by convert (Fintype.ofEquiv_card (mkReal (K := K))).symm #align number_field.infinite_place.card_real_embeddings NumberField.InfinitePlace.card_real_embeddings theorem card_complex_embeddings : @@ -528,7 +592,7 @@ theorem card_complex_embeddings : ext x rw [← mkComplex.filter_card K x] simp_rw [Finset.card_eq_sum_ones] - exact (Finset.sum_fiberwise Finset.univ (fun φ => mkComplex K φ) fun _ => 1).symm + exact (Finset.sum_fiberwise Finset.univ (fun φ => mkComplex φ) fun _ => 1).symm #align number_field.infinite_place.card_complex_embeddings NumberField.InfinitePlace.card_complex_embeddings end NumberField.InfinitePlace From 55b67ea1e74c11c96b8818dde20580ff9ff0f530 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 18 Jul 2023 09:16:18 +0900 Subject: [PATCH 26/67] Update --- Mathlib/GroupTheory/Torsion.lean | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/Mathlib/GroupTheory/Torsion.lean b/Mathlib/GroupTheory/Torsion.lean index 49a729a7d3eef..e219e1b680501 100644 --- a/Mathlib/GroupTheory/Torsion.lean +++ b/Mathlib/GroupTheory/Torsion.lean @@ -326,6 +326,9 @@ theorem torsion_eq_torsion_submonoid : CommMonoid.torsion G = (torsion G).toSubm #align comm_group.torsion_eq_torsion_submonoid CommGroup.torsion_eq_torsion_submonoid #align add_comm_group.add_torsion_eq_add_torsion_submonoid AddCommGroup.add_torsion_eq_add_torsion_submonoid +@[to_additive] + theorem mem_torsion (g : G) : g ∈ torsion G ↔ IsOfFinOrder g := Iff.rfl + variable (p : ℕ) [hp : Fact p.Prime] /-- The `p`-primary component is the subgroup of elements with order prime-power of `p`. -/ @@ -441,4 +444,3 @@ theorem IsTorsionFree.quotient_torsion : IsTorsionFree <| G ⧸ torsion G := fun #align add_is_torsion_free.quotient_torsion AddIsTorsionFree.quotient_torsion end CommGroup - From d7158931954bf1b7cd916c12a9e256ad49168979 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 19 Jul 2023 07:56:33 +0900 Subject: [PATCH 27/67] minimize imports --- Mathlib/NumberTheory/NumberField/Units.lean | 4 ---- 1 file changed, 4 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 38b62b6f3d2a5..3317e9dde4f8c 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -8,13 +8,9 @@ Authors: Xavier Roblot ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathlib.Algebra.Module.Zlattice -import Mathlib.GroupTheory.Torsion -import Mathlib.LinearAlgebra.Matrix.ToLinearEquiv import Mathlib.NumberTheory.NumberField.CanonicalEmbedding import Mathlib.NumberTheory.NumberField.Norm import Mathlib.RingTheory.Ideal.Norm -import Mathlib.RingTheory.RootsOfUnity.Basic /-! # Units of a number field From 892e300a9026a7c7d5999690e00bfed3d9dd6462 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 19 Jul 2023 08:08:36 +0900 Subject: [PATCH 28/67] update --- .../NumberField/CanonicalEmbedding.lean | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index f08d723347b85..942a86f223c82 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -71,7 +71,7 @@ theorem nnnorm_eq [NumberField K] (x : K) : theorem norm_le_iff [NumberField K] (x : K) (r : ℝ) : ‖canonicalEmbedding K x‖ ≤ r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by obtain hr | hr := lt_or_le r 0 - · obtain ⟨φ : K →+* ℂ⟩ := (inferInstance : Nonempty _) + · obtain ⟨φ⟩ := (inferInstance : Nonempty (K →+* ℂ)) refine iff_of_false ?_ ?_ exact (hr.trans_le (norm_nonneg _)).not_le exact fun h => hr.not_le (le_trans (norm_nonneg _) (h φ)) @@ -91,9 +91,9 @@ theorem integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : · simp [Metric.closedBall_eq_empty.2 hr] · have heq : ∀ x, canonicalEmbedding K x ∈ Metric.closedBall 0 r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by - intro x ; rw [← norm_le_iff, mem_closedBall_zero_iff] + intro x; rw [← norm_le_iff, mem_closedBall_zero_iff] convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) - ext ; constructor + ext; constructor · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx⟩ exact ⟨↑x, ⟨SetLike.coe_mem x, fun φ => (heq x).mp hx φ⟩, rfl⟩ · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ @@ -106,7 +106,7 @@ noncomputable def latticeBasis [NumberField K] : Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℂ ((K →+* ℂ) → ℂ) := by classical -- Let `B` be the canonical basis of `(K →+* ℂ) → ℂ`. We prove that the determinant of - -- the image by `canonicalEmbedding` of the integral basis of `K` is non-zero. This + -- the image by `canonicalEmbedding` of the integral basis of `K` is nonzero. This -- will imply the result. let B := Pi.basisFun ℂ (K →+* ℂ) let e : (K →+* ℂ) ≃ Free.ChooseBasisIndex ℤ (𝓞 K) := @@ -118,11 +118,11 @@ noncomputable def latticeBasis [NumberField K] : ((linearIndependent_equiv e.symm).mpr this.1) ?_ rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_fintype_fun_eq_card, Embeddings.card] - -- In order to prove that the determinant is non-zero, we show that it is equal to the + -- In order to prove that the determinant is nonzero, we show that it is equal to the -- square of the discriminant of the integral basis and thus it is not zero let N := Algebra.embeddingsMatrixReindex ℚ ℂ (fun i => integralBasis K (e i)) RingHom.equivRatAlgHom - rw [show M = N.transpose by { ext : 2 ; rfl }] + rw [show M = N.transpose by { ext:2; rfl }] rw [Matrix.det_transpose, ← @pow_ne_zero_iff ℂ _ _ _ 2 (by norm_num)] convert (map_ne_zero_iff _ (algebraMap ℚ ℂ).injective).mpr (Algebra.discr_not_zero_of_basis ℚ (integralBasis K)) @@ -140,10 +140,10 @@ theorem mem_span_latticeBasis [NumberField K] (x : (K →+* ℂ) → ℂ) : x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ canonicalEmbedding K '' (𝓞 K) := by rw [show Set.range (latticeBasis K) = (canonicalEmbedding K).toIntAlgHom.toLinearMap '' (Set.range (integralBasis K)) by - rw [← Set.range_comp] ; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] + rw [← Set.range_comp]; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] rw [← Submodule.map_span, ← SetLike.mem_coe, Submodule.map_coe] rw [show (Submodule.span ℤ (Set.range (integralBasis K)) : Set K) = 𝓞 K by - ext ; exact mem_span_integralBasis K] + ext; exact mem_span_integralBasis K] rfl end NumberField.canonicalEmbedding From 4a13f8666738d3aa06671cd26413a6620c04dd02 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 19 Jul 2023 08:45:48 +0900 Subject: [PATCH 29/67] Update --- Mathlib/NumberTheory/NumberField/Units.lean | 15 +++++++++++++-- 1 file changed, 13 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 44b1151684e22..bb7da44c68656 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -65,16 +65,27 @@ section coe /-- The `MonoidHom` from the group of units `(𝓞 K)ˣ` to the field `K`. -/ def coe_to_field : (𝓞 K)ˣ →* K := (Units.coeHom K).comp (map (algebraMap (𝓞 K) K)) +variable {K} + +@[coe] def to_field (x : (𝓞 K)ˣ) : K := coe_to_field K x + +variable (K) + theorem coe_to_field_injective : Function.Injective (coe_to_field K) := fun _ _ h => Units.eq_iff.mp (SetCoe.ext h) /-- There is a natural coercion from `(𝓞 K)ˣ` to `(𝓞 K)` and then from `(𝓞 K)` to `K` but it is useful to also have a direct one from `(𝓞 K)ˣ` to `K`. -/ -instance : Coe (𝓞 K)ˣ K := ⟨coe_to_field K⟩ +instance : Coe (𝓞 K)ˣ K := ⟨to_field⟩ @[ext] theorem ext {x y : (𝓞 K)ˣ} (h : (x : K) = y) : x = y := (coe_to_field_injective K).eq_iff.mp h +theorem map_pow (x : (𝓞 K)ˣ) (n : ℕ) : (x ^ n : K) = (x : K) ^ n := + _root_.map_pow (coe_to_field K) x n + +theorem map_one : ((1 : (𝓞 K)ˣ) : K) = 1 := rfl + end coe open NumberField.InfinitePlace @@ -136,6 +147,6 @@ theorem rootsOfUnity_eq_torsion [NumberField K] : refine ⟨fun h => ?_, fun h => ?_⟩ · rw [CommGroup.mem_torsion, isOfFinOrder_iff_pow_eq_one] exact ⟨↑(torsion_order K), (torsion_order K).prop, h⟩ - · exact Subtype.ext_iff.mp (@pow_card_eq_one (torsion K) ⟨ζ, h⟩ _ _) + · exact Subtype.ext_iff.mp (@pow_card_eq_one (torsion K) _ ⟨ζ, h⟩ _) end NumberField.units From b33d1467e1c085384f4e2022e48f88d298c80746 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 19 Jul 2023 11:36:34 +0900 Subject: [PATCH 30/67] Update --- .../NumberField/CanonicalEmbedding.lean | 41 ++++++++++--------- 1 file changed, 21 insertions(+), 20 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index f45f3f4c2784d..07952b5f6e85a 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -14,7 +14,6 @@ import Mathlib.MeasureTheory.Measure.Haar.NormedSpace import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.RingTheory.Discriminant - /-! # Canonical embedding of a number field @@ -33,6 +32,9 @@ radius is finite. `φ_₁,...,φ_r₁` are its real embeddings and `ψ_₁,..., ψ_r₂` are its complex embeddings (up to complex conjugation). +* `exists_ne_zero_mem_ringOfIntegers_lt`: let `f : InfinitePlace K → ℝ≥0`, if the product +`∏_w f w` is large enough, proves that there exists a nonzero algebraic integer `a` such +that `w a < f w` for all infinite places `w`. ## Tags @@ -60,7 +62,7 @@ theorem apply_at (φ : K →+* ℂ) (x : K) : open scoped ComplexConjugate /-- The image of `canonicalEmbedding` lives in the `ℝ`-submodule of the `x ∈ ((K →+* ℂ) → ℂ)` such -that `conj x_φ = x_(conj φ)` for all `∀ φ : K →+* ℂ`. -/ +that `conj x_φ = x_(conj φ)` for all `φ : K →+* ℂ`. -/ theorem conj_apply {x : ((K →+* ℂ) → ℂ)} (φ : K →+* ℂ) (hx : x ∈ Submodule.span ℝ (Set.range (canonicalEmbedding K))) : conj (x φ) = x (ComplexEmbedding.conjugate φ) := by @@ -73,14 +75,13 @@ theorem conj_apply {x : ((K →+* ℂ) → ℂ)} (φ : K →+* ℂ) exact congrArg ((a : ℂ) * ·) hx theorem nnnorm_eq [NumberField K] (x : K) : - ‖canonicalEmbedding K x‖₊ = - Finset.univ.sup (fun φ : K →+* ℂ => ‖φ x‖₊) := by + ‖canonicalEmbedding K x‖₊ = Finset.univ.sup (fun φ : K →+* ℂ => ‖φ x‖₊) := by simp_rw [Pi.nnnorm_def, apply_at] theorem norm_le_iff [NumberField K] (x : K) (r : ℝ) : ‖canonicalEmbedding K x‖ ≤ r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by obtain hr | hr := lt_or_le r 0 - · obtain ⟨φ : K →+* ℂ⟩ := (inferInstance : Nonempty _) + · obtain ⟨φ⟩ := (inferInstance : Nonempty (K →+* ℂ)) refine iff_of_false ?_ ?_ exact (hr.trans_le (norm_nonneg _)).not_le exact fun h => hr.not_le (le_trans (norm_nonneg _) (h φ)) @@ -100,9 +101,9 @@ theorem integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : · simp [Metric.closedBall_eq_empty.2 hr] · have heq : ∀ x, canonicalEmbedding K x ∈ Metric.closedBall 0 r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by - intro x ; rw [← norm_le_iff, mem_closedBall_zero_iff] + intro x; rw [← norm_le_iff, mem_closedBall_zero_iff] convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) - ext ; constructor + ext; constructor · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx⟩ exact ⟨↑x, ⟨SetLike.coe_mem x, fun φ => (heq x).mp hx φ⟩, rfl⟩ · rintro ⟨x, ⟨hx1, hx2⟩, rfl⟩ @@ -115,7 +116,7 @@ noncomputable def latticeBasis [NumberField K] : Basis (Free.ChooseBasisIndex ℤ (𝓞 K)) ℂ ((K →+* ℂ) → ℂ) := by classical -- Let `B` be the canonical basis of `(K →+* ℂ) → ℂ`. We prove that the determinant of - -- the image by `canonicalEmbedding` of the integral basis of `K` is non-zero. This + -- the image by `canonicalEmbedding` of the integral basis of `K` is nonzero. This -- will imply the result. let B := Pi.basisFun ℂ (K →+* ℂ) let e : (K →+* ℂ) ≃ Free.ChooseBasisIndex ℤ (𝓞 K) := @@ -127,11 +128,11 @@ noncomputable def latticeBasis [NumberField K] : ((linearIndependent_equiv e.symm).mpr this.1) ?_ rw [← finrank_eq_card_chooseBasisIndex, RingOfIntegers.rank, finrank_fintype_fun_eq_card, Embeddings.card] - -- In order to prove that the determinant is non-zero, we show that it is equal to the + -- In order to prove that the determinant is nonzero, we show that it is equal to the -- square of the discriminant of the integral basis and thus it is not zero let N := Algebra.embeddingsMatrixReindex ℚ ℂ (fun i => integralBasis K (e i)) RingHom.equivRatAlgHom - rw [show M = N.transpose by { ext : 2 ; rfl }] + rw [show M = N.transpose by { ext:2; rfl }] rw [Matrix.det_transpose, ← @pow_ne_zero_iff ℂ _ _ _ 2 (by norm_num)] convert (map_ne_zero_iff _ (algebraMap ℚ ℂ).injective).mpr (Algebra.discr_not_zero_of_basis ℚ (integralBasis K)) @@ -208,7 +209,7 @@ theorem comm_map_canonical_eq_mixed (x : K) : mixedEmbedding, RingHom.prod_apply, Prod.mk.injEq] exact ⟨rfl, rfl⟩ -/-- This is technical result to ensure that the image of the `ℂ`-basis of `ℂ^n`, defined in +/-- This is a technical result to ensure that the image of the `ℂ`-basis of `ℂ^n` defined in `canonicalEmbedding.latticeBasis` is a `ℝ`-basis of `ℝ^r₁ × ℂ^r₂`, see `mixedEmbedding.latticeBasis`. -/ theorem disjoint_span_comm_map_ker [NumberField K]: @@ -285,7 +286,7 @@ open Metric ENNReal NNReal variable (f : InfinitePlace K → ℝ≥0) /-- The convex body defined by `f`: the set of points `x : E` such that `‖x w‖ < f w` for all -infinite place `w`. -/ +infinite places `w`. -/ def convex_body : Set (E K) := (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsReal w } => ball 0 (f w))) ×ˢ (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsComplex w } => ball 0 (f w))) @@ -337,24 +338,24 @@ theorem convex_body_volume : ∏ w : InfinitePlace K, ite (IsReal w) ↑(f w) ↑(f w ^ 2) := by rw [volume_eq_prod, convex_body, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi] conv_lhs => - congr ; congr ; next => skip + congr; congr; next => skip ext rw [Real.volume_ball, ofReal_mul (by norm_num), ofReal_coe_nnreal, mul_comm] conv_lhs => - congr ; next => skip - congr ; next => skip + congr; next => skip + congr; next => skip ext i - rw [add_haar_ball _ _ (by exact (f i).prop), Complex.finrank_real_complex, ← NNReal.coe_pow, + rw [addHaar_ball _ _ (by exact (f i).prop), Complex.finrank_real_complex, ← NNReal.coe_pow, ofReal_coe_nnreal, mul_comm] rw [Finset.prod_mul_distrib, Finset.prod_mul_distrib, Finset.prod_const, Finset.prod_const, Finset.card_univ, Finset.card_univ, mul_assoc, mul_comm, ← mul_assoc, mul_assoc, ofReal_ofNat, ← constant_factor, Finset.prod_ite] simp_rw [show (fun w : InfinitePlace K ↦ ¬IsReal w) = (fun w ↦ IsComplex w) - by funext ; rw [not_isReal_iff_isComplex]] - congr 1 ; rw [mul_comm] ; congr 1 + by funext; rw [not_isReal_iff_isComplex]] + congr 1; rw [mul_comm]; congr 1 all_goals · rw [← Finset.prod_subtype_eq_prod_filter] - congr ; ext + congr; ext exact ⟨fun _ => Finset.mem_subtype.mpr (Finset.mem_univ _), fun _ => Finset.mem_univ _⟩ end convex_body @@ -375,7 +376,7 @@ theorem minkowski_bound_lt_top : minkowski_bound K < ⊤ := by · exact ne_of_lt (Zspan.fundamentalDomain_bounded (latticeBasis K)).measure_lt_top · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) -theorem exists_ne_zero_mem_ring_of_integers_lt (h : minkowski_bound K < volume (convex_body K f)) : +theorem exists_ne_zero_mem_ringOfIntegers_lt (h : minkowski_bound K < volume (convex_body K f)) : ∃ (a : 𝓞 K), a ≠ 0 ∧ ∀ w : InfinitePlace K, w a < f w := by have : @IsAddHaarMeasure (E K) _ _ _ volume := prod.instIsAddHaarMeasure volume volume have h_fund := Zspan.isAddFundamentalDomain (latticeBasis K) volume From bc2722060af10a61701576709f6058d954cc78e5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 19 Jul 2023 13:11:01 +0900 Subject: [PATCH 31/67] More semicolons --- Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 07952b5f6e85a..7c8d504b633f4 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -150,10 +150,10 @@ theorem mem_span_latticeBasis [NumberField K] (x : (K →+* ℂ) → ℂ) : x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ canonicalEmbedding K '' (𝓞 K) := by rw [show Set.range (latticeBasis K) = (canonicalEmbedding K).toIntAlgHom.toLinearMap '' (Set.range (integralBasis K)) by - rw [← Set.range_comp] ; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] + rw [← Set.range_comp]; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] rw [← Submodule.map_span, ← SetLike.mem_coe, Submodule.map_coe] rw [show (Submodule.span ℤ (Set.range (integralBasis K)) : Set K) = 𝓞 K by - ext ; exact mem_span_integralBasis K] + ext; exact mem_span_integralBasis K] rfl end NumberField.canonicalEmbedding @@ -271,10 +271,10 @@ theorem mem_span_latticeBasis [NumberField K] (x : (E K)) : x ∈ Submodule.span ℤ (Set.range (latticeBasis K)) ↔ x ∈ mixedEmbedding K '' (𝓞 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))] + rw [← Set.range_comp]; exact congrArg Set.range (funext (fun i => latticeBasis_apply K i))] rw [← Submodule.map_span, ← SetLike.mem_coe, Submodule.map_coe] rw [show (Submodule.span ℤ (Set.range (integralBasis K)) : Set K) = 𝓞 K by - ext ; exact mem_span_integralBasis K] + ext; exact mem_span_integralBasis K] rfl end integerLattice From 18d3d6188b30c944d5b07a6364a9c945c5201194 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 19 Jul 2023 18:59:52 +0900 Subject: [PATCH 32/67] backport --- .../NumberField/CanonicalEmbedding.lean | 30 ++++++-- .../NumberTheory/NumberField/Embeddings.lean | 72 +++++++++++++------ 2 files changed, 77 insertions(+), 25 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 7c8d504b633f4..48667aa8d7372 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -227,10 +227,10 @@ theorem disjoint_span_comm_map_ker [NumberField K]: rw [eq_comm, ← Complex.conj_eq_iff_re, canonicalEmbedding.conj_apply _ h_mem, ComplexEmbedding.isReal_iff.mp hφ], ← Complex.ofReal_zero] congr - rw [← ComplexEmbeddings.IsReal.embedding_mk hφ, ← comm_map_apply_of_isReal K x ⟨φ, hφ, rfl⟩] + rw [← ComplexEmbedding.IsReal.embedding_mk hφ, ← comm_map_apply_of_isReal K x ⟨φ, hφ, rfl⟩] exact congrFun (congrArg (fun x => x.1) h_zero) ⟨InfinitePlace.mk φ, _⟩ · have := congrFun (congrArg (fun x => x.2) h_zero) ⟨InfinitePlace.mk φ, ⟨φ, hφ, rfl⟩⟩ - cases ComplexEmbeddings.embedding_mk φ with + cases ComplexEmbedding.embedding_mk φ with | inl h => rwa [← h, ← comm_map_apply_of_isComplex K x ⟨φ, hφ, rfl⟩] | inr h => apply RingHom.injective (starRingEnd ℂ) @@ -334,8 +334,7 @@ theorem constant_factor_lt_top : (constant_factor K) < ⊤ := by set_option maxHeartbeats 400000 in theorem convex_body_volume : - volume (convex_body K f) = (constant_factor K) * - ∏ w : InfinitePlace K, ite (IsReal w) ↑(f w) ↑(f w ^ 2) := by + volume (convex_body K f) = (constant_factor K) * ∏ w, (f w) ^ (mult K w) := by rw [volume_eq_prod, convex_body, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi] conv_lhs => congr; congr; next => skip @@ -349,7 +348,10 @@ theorem convex_body_volume : ofReal_coe_nnreal, mul_comm] rw [Finset.prod_mul_distrib, Finset.prod_mul_distrib, Finset.prod_const, Finset.prod_const, Finset.card_univ, Finset.card_univ, mul_assoc, mul_comm, ← mul_assoc, mul_assoc, ofReal_ofNat, - ← constant_factor, Finset.prod_ite] + ← constant_factor] + simp_rw [mult, pow_ite, pow_one] + rw [Finset.prod_ite] + simp_rw [coe_mul, coe_finset_prod] simp_rw [show (fun w : InfinitePlace K ↦ ¬IsReal w) = (fun w ↦ IsComplex w) by funext; rw [not_isReal_iff_isComplex]] congr 1; rw [mul_comm]; congr 1 @@ -358,6 +360,24 @@ theorem convex_body_volume : congr; ext exact ⟨fun _ => Finset.mem_subtype.mpr (Finset.mem_univ _), fun _ => Finset.mem_univ _⟩ +variable {f} + +/-- This is a technical result: quite often, we want to impose conditions at all infinite places +but one and choose the value at the remaining place so that we can apply +`exists_ne_zero_mem_ring_of_integers_lt`. -/ +theorem adjust_f {w₁ : InfinitePlace K} (B : ℝ≥0) (hf : ∀ w, w ≠ w₁→ f w ≠ 0) : + ∃ g : InfinitePlace K → ℝ≥0, (∀ w, w ≠ w₁ → g w = f w) ∧ ∏ w, (g w) ^ mult K w = B := by + let S := ∏ w in Finset.univ.erase w₁, (f w) ^ mult K w + refine ⟨Function.update f w₁ ((B * S⁻¹) ^ (mult K w₁ : ℝ)⁻¹), ?_, ?_⟩ + · exact fun w hw => Function.update_noteq hw _ f + · rw [← Finset.mul_prod_erase Finset.univ _ (Finset.mem_univ w₁), Function.update_same, + Finset.prod_congr rfl fun w hw => by rw [Function.update_noteq (Finset.ne_of_mem_erase hw)], + ← NNReal.rpow_nat_cast, ← NNReal.rpow_mul, inv_mul_cancel, NNReal.rpow_one, mul_assoc, + inv_mul_cancel, mul_one] + · rw [Finset.prod_ne_zero_iff] + exact fun w hw => pow_ne_zero _ (hf w (Finset.ne_of_mem_erase hw)) + · rw [mult]; split_ifs <;> norm_num + end convex_body section minkowski diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index 748ab6772a9f9..c90e5b8c0719c 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -182,6 +182,9 @@ def IsReal (φ : K →+* ℂ) : Prop := theorem isReal_iff {φ : K →+* ℂ} : IsReal φ ↔ conjugate φ = φ := isSelfAdjoint_iff #align number_field.complex_embedding.is_real_iff NumberField.ComplexEmbedding.isReal_iff +theorem isReal_iff_isReal_conj {φ : K →+* ℂ} : + IsReal φ ↔ IsReal (conjugate φ) := by rw [IsReal, IsReal, IsSelfAdjoint.star_iff] + /-- A real embedding as a ring homomorphism from `K` to `ℝ` . -/ def IsReal.embedding {φ : K →+* ℂ} (hφ : IsReal φ) : K →+* ℝ where toFun x := (φ x).re @@ -352,21 +355,21 @@ def IsComplex (w : InfinitePlace K) : Prop := #align number_field.infinite_place.is_complex NumberField.InfinitePlace.IsComplex @[simp] -theorem _root_.NumberField.ComplexEmbeddings.IsReal.embedding_mk {φ : K →+* ℂ} +theorem _root_.NumberField.ComplexEmbedding.IsReal.embedding_mk {φ : K →+* ℂ} (h : ComplexEmbedding.IsReal φ) : embedding (mk φ) = φ := by have := mk_eq_iff.mp (mk_embedding (mk φ)).symm rwa [ComplexEmbedding.isReal_iff.mp h, or_self_iff, eq_comm] at this -#align number_field.complex_embeddings.is_real.embedding_mk NumberField.ComplexEmbeddings.IsReal.embedding_mk +#align number_field.complex_embeddings.is_real.embedding_mk NumberField.ComplexEmbedding.IsReal.embedding_mk @[simp] -theorem _root_.NumberField.ComplexEmbeddings.embedding_mk (φ : K →+* ℂ) : +theorem _root_.NumberField.ComplexEmbedding.embedding_mk (φ : K →+* ℂ) : embedding (mk φ) = φ ∨ embedding (mk φ) = ComplexEmbedding.conjugate φ := by rw [@eq_comm _ _ φ, @eq_comm _ _ (ComplexEmbedding.conjugate φ), ← mk_eq_iff, mk_embedding] theorem isReal_iff {w : InfinitePlace K} : IsReal w ↔ ComplexEmbedding.IsReal (embedding w) := by constructor · rintro ⟨φ, ⟨hφ, rfl⟩⟩ - rwa [_root_.NumberField.ComplexEmbeddings.IsReal.embedding_mk hφ] + rwa [_root_.NumberField.ComplexEmbedding.IsReal.embedding_mk hφ] · exact fun h => ⟨embedding w, h, mk_embedding w⟩ #align number_field.infinite_place.is_real_iff NumberField.InfinitePlace.isReal_iff @@ -382,11 +385,24 @@ theorem isComplex_iff {w : InfinitePlace K} : · exact fun h => ⟨embedding w, h, mk_embedding w⟩ #align number_field.infinite_place.is_complex_iff NumberField.InfinitePlace.isComplex_iff +theorem _root_.NumberField.ComplexEmbedding.isReal_mk_iff {φ : K →+* ℂ} : + IsReal (mk φ) ↔ ComplexEmbedding.IsReal φ := by + refine ⟨fun h => ?_, fun h => ⟨φ, h, rfl⟩⟩ + · rw [isReal_iff] at h + cases ComplexEmbedding.embedding_mk φ with + | inl hφ => rwa [← hφ] + | inr hφ => rwa [ComplexEmbedding.isReal_iff_isReal_conj, ← hφ] + @[simp] theorem not_isReal_iff_isComplex {w : InfinitePlace K} : ¬IsReal w ↔ IsComplex w := by rw [isComplex_iff, isReal_iff] #align number_field.infinite_place.not_is_real_iff_is_complex NumberField.InfinitePlace.not_isReal_iff_isComplex +theorem _root_.NumberField.ComplexEmbedding.not_isReal_mk_iff {φ : K →+* ℂ} : + ¬ ComplexEmbedding.IsReal φ ↔ IsComplex (mk φ) := by + convert (ComplexEmbedding.isReal_mk_iff (φ := φ)).symm.not + exact not_isReal_iff_isComplex.symm + @[simp] theorem not_isComplex_iff_isReal {w : InfinitePlace K} : ¬IsComplex w ↔ IsReal w := by rw [← not_isReal_iff_isComplex, Classical.not_not] @@ -416,12 +432,18 @@ theorem IsReal.abs_embedding_apply {w : InfinitePlace K} (hw : IsReal w) (x : K) variable (K) +/-- The multiplicity of an infinite place, that is the number of distinct complex embeddings that +defines it, see ` mult_eq`. -/ +noncomputable def mult (w : InfinitePlace K) : ℕ := if (IsReal w) then 1 else 2 + +variable {K} + /-- The map from real embeddings to real infinite places as an equiv -/ noncomputable def mkReal : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } ≃ { w : InfinitePlace K // IsReal w } where toFun := Subtype.map mk fun φ hφ => ⟨φ, hφ, rfl⟩ invFun w := ⟨w.1.embedding, isReal_iff.1 w.2⟩ - left_inv φ := Subtype.ext_iff.2 (NumberField.ComplexEmbeddings.IsReal.embedding_mk φ.2) + left_inv φ := Subtype.ext_iff.2 (NumberField.ComplexEmbedding.IsReal.embedding_mk φ.2) right_inv w := Subtype.ext_iff.2 (mk_embedding w.1) #align number_field.infinite_place.mk_real NumberField.InfinitePlace.mkReal @@ -431,36 +453,38 @@ noncomputable def mkComplex : Subtype.map mk fun φ hφ => ⟨φ, hφ, rfl⟩ #align number_field.infinite_place.mk_complex NumberField.InfinitePlace.mkComplex +variable (K) + theorem mkComplex_embedding (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) : - (mkComplex K φ : InfinitePlace K).embedding = φ ∨ - (mkComplex K φ : InfinitePlace K).embedding = ComplexEmbedding.conjugate φ := - NumberField.ComplexEmbeddings.embedding_mk φ.1 + (mkComplex φ : InfinitePlace K).embedding = φ ∨ + (mkComplex φ : InfinitePlace K).embedding = ComplexEmbedding.conjugate φ := + NumberField.ComplexEmbedding.embedding_mk φ.1 #align number_field.infinite_place.mk_complex_embedding NumberField.InfinitePlace.mkComplex_embedding @[simp] theorem mkReal_coe (φ : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ }) : - (mkReal K φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl + (mkReal φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl #align number_field.infinite_place.mk_real_coe NumberField.InfinitePlace.mkReal_coe @[simp] theorem mkComplex_coe (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) : - (mkComplex K φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl + (mkComplex φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl #align number_field.infinite_place.mk_complex_coe NumberField.InfinitePlace.mkComplex_coe -- @[simp] -- Porting note: not in simpNF theorem mkReal.apply (φ : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ }) (x : K) : - (mkReal K φ).val x = Complex.abs (φ.val x) := rfl + (mkReal φ).val x = Complex.abs (φ.val x) := rfl #align number_field.infinite_place.mk_real.apply NumberField.InfinitePlace.mkReal.apply -- @[simp] -- Porting note: not in simpNF theorem mkComplex.apply (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) (x : K) : - (mkComplex K φ).val x = Complex.abs (φ.val x) := rfl + (mkComplex φ).val x = Complex.abs (φ.val x) := rfl #align number_field.infinite_place.mk_complex.apply NumberField.InfinitePlace.mkComplex.apply variable [NumberField K] theorem mkComplex.filter (w : { w : InfinitePlace K // w.IsComplex }) : - (Finset.univ.filter fun φ => mkComplex K φ = w) = + (Finset.univ.filter fun φ => mkComplex φ = w) = {⟨w.1.embedding, isComplex_iff.1 w.2⟩, ⟨ComplexEmbedding.conjugate w.1.embedding, ComplexEmbedding.isReal_conjugate_iff.not.2 (isComplex_iff.1 w.2)⟩} := by @@ -472,7 +496,7 @@ theorem mkComplex.filter (w : { w : InfinitePlace K // w.IsComplex }) : #align number_field.infinite_place.mk_complex.filter NumberField.InfinitePlace.mkComplex.filter theorem mkComplex.filter_card (w : { w : InfinitePlace K // w.IsComplex }) : - (Finset.univ.filter fun φ => mkComplex K φ = w).card = 2 := by + (Finset.univ.filter fun φ => mkComplex φ = w).card = 2 := by rw [mkComplex.filter] exact Finset.card_doubleton (Subtype.mk_eq_mk.not.2 <| @@ -483,11 +507,12 @@ noncomputable instance NumberField.InfinitePlace.fintype : Fintype (InfinitePlac Set.fintypeRange _ #align number_field.infinite_place.number_field.infinite_place.fintype NumberField.InfinitePlace.NumberField.InfinitePlace.fintype +open scoped BigOperators + /-- The infinite part of the product formula : for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where `‖·‖_w` is the normalized absolute value for `w`. -/ theorem prod_eq_abs_norm (x : K) : - (Finset.univ.prod fun w : InfinitePlace K => ite w.IsReal (w x) (w x ^ 2)) = - abs (Algebra.norm ℚ x) := by + ∏ w : InfinitePlace K, ite w.IsReal (w x) (w x ^ 2) = abs (Algebra.norm ℚ x) := by convert (congr_arg Complex.abs (@Algebra.norm_eq_prod_embeddings ℚ _ _ _ _ ℂ _ _ _ _ _ x)).symm · rw [map_prod, ← Equiv.prod_comp' RingHom.equivRatAlgHom (fun f => Complex.abs (f x)) (fun φ => Complex.abs (φ x)) fun _ => by simp [RingHom.equivRatAlgHom_apply]; rfl] @@ -499,13 +524,13 @@ theorem prod_eq_abs_norm (x : K) : rw [Finset.prod_ite, Finset.prod_ite] refine congr (congr_arg Mul.mul ?_) ?_ · rw [← Finset.prod_subtype_eq_prod_filter, ← Finset.prod_subtype_eq_prod_filter] - convert (Equiv.prod_comp' (mkReal K) (fun φ => Complex.abs (φ.val x)) + convert (Equiv.prod_comp' mkReal (fun φ => Complex.abs (φ.val x)) (fun w => w.val x) _).symm any_goals ext; simp only [Finset.mem_subtype, Finset.mem_univ] exact fun φ => mkReal.apply K φ x · rw [Finset.filter_congr fun (w : InfinitePlace K) _ => @not_isReal_iff_isComplex K _ w, ← Finset.prod_subtype_eq_prod_filter, ← Finset.prod_subtype_eq_prod_filter] - convert Finset.prod_fiberwise Finset.univ (fun φ => mkComplex K φ) + convert Finset.prod_fiberwise Finset.univ (fun φ => mkComplex φ) (fun φ => Complex.abs (φ.val x)) using 2 with w · ext; simp only [Finset.mem_subtype, Finset.mem_univ, not_isReal_iff_isComplex] · rw [@Finset.prod_congr _ _ _ _ _ (fun _ => w.val x) _ (Eq.refl _) fun φ hφ => @@ -515,11 +540,18 @@ theorem prod_eq_abs_norm (x : K) : · rw [eq_ratCast, Rat.cast_abs, ← Complex.abs_ofReal, Complex.ofReal_rat_cast] #align number_field.infinite_place.prod_eq_abs_norm NumberField.InfinitePlace.prod_eq_abs_norm +/-- A version of `prod_eq_abs_norm` that uses `mult`. -/ +theorem prod_mult_eq_abs_norm (x : K) : + ∏ w : InfinitePlace K, w x ^ mult K w = abs (Algebra.norm ℚ x) := by + convert prod_eq_abs_norm K x using 2 + rw [mult] + split_ifs <;> simp + open Fintype theorem card_real_embeddings : card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } = card { w : InfinitePlace K // IsReal w } := - by convert (Fintype.ofEquiv_card (mkReal K)).symm + by convert (Fintype.ofEquiv_card (mkReal (K := K))).symm #align number_field.infinite_place.card_real_embeddings NumberField.InfinitePlace.card_real_embeddings theorem card_complex_embeddings : @@ -531,7 +563,7 @@ theorem card_complex_embeddings : ext x rw [← mkComplex.filter_card K x] simp_rw [Finset.card_eq_sum_ones] - exact (Finset.sum_fiberwise Finset.univ (fun φ => mkComplex K φ) fun _ => 1).symm + exact (Finset.sum_fiberwise Finset.univ (fun φ => mkComplex φ) fun _ => 1).symm #align number_field.infinite_place.card_complex_embeddings NumberField.InfinitePlace.card_complex_embeddings end NumberField.InfinitePlace From 8f6fc4fe0f3966907092ccf434acc0bc6c5c7286 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Thu, 27 Jul 2023 07:23:08 +0900 Subject: [PATCH 33/67] 1st commit --- .../NumberField/CanonicalEmbedding.lean | 16 +- .../NumberTheory/NumberField/Embeddings.lean | 335 ++++++++---------- 2 files changed, 153 insertions(+), 198 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 5cedd30627d96..58d9c2ade9630 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -65,7 +65,8 @@ theorem nontrivial_space [NumberField K] : Nontrivial E := by /-- The canonical embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ def _root_.NumberField.canonicalEmbedding : K →+* E := - RingHom.prod (Pi.ringHom fun w => w.prop.embedding) (Pi.ringHom fun w => w.val.embedding) + RingHom.prod (Pi.ringHom fun w => embedding_of_isReal w.prop) + (Pi.ringHom fun w => w.val.embedding) #align number_field.canonical_embedding NumberField.canonicalEmbedding theorem _root_.NumberField.canonicalEmbedding_injective [NumberField K] : @@ -79,7 +80,7 @@ variable {K} @[simp] theorem apply_at_real_infinitePlace (w : { w : InfinitePlace K // IsReal w }) (x : K) : - (NumberField.canonicalEmbedding K x).1 w = w.prop.embedding x := by + (NumberField.canonicalEmbedding K x).1 w = embedding_of_isReal w.prop x := by simp only [canonicalEmbedding, RingHom.prod_apply, Pi.ringHom_apply] #align number_field.canonical_embedding.apply_at_real_infinite_place NumberField.canonicalEmbedding.apply_at_real_infinitePlace @@ -102,15 +103,14 @@ theorem nnnorm_eq [NumberField K] (x : K) : fun w => (⟨w x, map_nonneg w x⟩ : NNReal)).symm using 2 ext w dsimp - simp only [apply_at_real_infinitePlace, coe_nnnorm, Real.norm_eq_abs, - Function.Embedding.coe_subtype, Subtype.coe_mk, IsReal.abs_embedding_apply] + rw [apply_at_real_infinitePlace, ← Complex.abs_ofReal, embedding_of_isReal_apply, + ← Complex.norm_eq_abs, norm_embedding_eq] · convert (Finset.univ.sup_map (Function.Embedding.subtype fun w : InfinitePlace K => IsComplex w) fun w => (⟨w x, map_nonneg w x⟩ : NNReal)).symm using 2 ext w dsimp - simp only [apply_at_complex_infinitePlace, coe_nnnorm, - Complex.norm_eq_abs, Function.Embedding.coe_subtype, Subtype.coe_mk, abs_embedding] + rw [apply_at_complex_infinitePlace, ← Complex.norm_eq_abs, norm_embedding_eq] · ext w simp_rw [Finset.mem_univ, Finset.mem_union, Set.mem_toFinset, Set.mem_setOf_eq, w.isReal_or_isComplex] @@ -163,8 +163,8 @@ theorem integerLattice.inter_ball_finite [NumberField K] (r : ℝ) : obtain hr | _ := lt_or_le r 0 · simp [closedBall_eq_empty.2 hr] have heq : ∀ x, canonicalEmbedding K x ∈ closedBall (0 : E) r ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := by - simp only [← place_apply, ← InfinitePlace.coe_mk, mem_closedBall_zero_iff, norm_le_iff] - exact fun x => le_iff_le x r + simp_rw [← place_apply, mem_closedBall_zero_iff, norm_le_iff, le_iff_le, place_apply, + implies_true] convert (Embeddings.finite_of_norm_le K ℂ r).image (canonicalEmbedding K) ext; constructor · rintro ⟨⟨_, ⟨x, rfl⟩, rfl⟩, hx2⟩ diff --git a/Mathlib/NumberTheory/NumberField/Embeddings.lean b/Mathlib/NumberTheory/NumberField/Embeddings.lean index b5a2742268a88..51cf070ccf156 100644 --- a/Mathlib/NumberTheory/NumberField/Embeddings.lean +++ b/Mathlib/NumberTheory/NumberField/Embeddings.lean @@ -15,12 +15,18 @@ import Mathlib.Topology.Instances.Complex # Embeddings of number fields This file defines the embeddings of a number field into an algebraic closed field. -## Main Results +## Main Definitions and Results * `NumberField.Embeddings.range_eval_eq_rootSet_minpoly`: let `x ∈ K` with `K` number field and let `A` be an algebraic closed field of char. 0, then the images of `x` by the embeddings of `K` in `A` are exactly the roots in `A` of the minimal polynomial of `x` over `ℚ`. * `NumberField.Embeddings.pow_eq_one_of_norm_eq_one`: an algebraic integer whose conjugates are all of norm one is a root of unity. +* `NumberField.InfinitePlace`: the type of infinite places of a number field `K`. +* `NumberField.InfinitePlace.mk_eq_iff`: two complex embeddings define the same infinite place iff +they are equal or complex conjugates. +* `NumberField.InfinitePlace.prod_eq_abs_norm`: the infinite part of the product formula, that is +for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where the product is over the infinite place `w` and +`‖·‖_w` is the normalized absolute value for `w`. ## Tags number field, embeddings, places, infinite places @@ -158,27 +164,25 @@ variable {K : Type _} [Field K] /-- The conjugate of a complex embedding as a complex embedding. -/ @[reducible] -def conjugate (φ : K →+* ℂ) : K →+* ℂ := - star φ +def conjugate (φ : K →+* ℂ) : K →+* ℂ := star φ #align number_field.complex_embedding.conjugate NumberField.ComplexEmbedding.conjugate @[simp] theorem conjugate_coe_eq (φ : K →+* ℂ) (x : K) : (conjugate φ) x = conj (φ x) := rfl #align number_field.complex_embedding.conjugate_coe_eq NumberField.ComplexEmbedding.conjugate_coe_eq -theorem place_conjugate (φ : K →+* ℂ) : place (conjugate φ) = place φ := by - ext; simp only [place_apply, norm_eq_abs, abs_conj, conjugate_coe_eq] -#align number_field.complex_embedding.place_conjugate NumberField.ComplexEmbedding.place_conjugate - /-- An embedding into `ℂ` is real if it is fixed by complex conjugation. -/ @[reducible] -def IsReal (φ : K →+* ℂ) : Prop := - IsSelfAdjoint φ +def IsReal (φ : K →+* ℂ) : Prop := IsSelfAdjoint φ #align number_field.complex_embedding.is_real NumberField.ComplexEmbedding.IsReal theorem isReal_iff {φ : K →+* ℂ} : IsReal φ ↔ conjugate φ = φ := isSelfAdjoint_iff #align number_field.complex_embedding.is_real_iff NumberField.ComplexEmbedding.isReal_iff +theorem isReal_conjugate_iff {φ : K →+* ℂ} : IsReal (conjugate φ) ↔ IsReal φ := + IsSelfAdjoint.star_iff +#align number_field.complex_embedding.is_real_conjugate_iff NumberField.ComplexEmbedding.isReal_conjugate_iff + /-- A real embedding as a ring homomorphism from `K` to `ℝ` . -/ def IsReal.embedding {φ : K →+* ℂ} (hφ : IsReal φ) : K →+* ℝ where toFun x := (φ x).re @@ -199,15 +203,6 @@ theorem IsReal.coe_embedding_apply {φ : K →+* ℂ} (hφ : IsReal φ) (x : K) exact RingHom.congr_fun hφ x #align number_field.complex_embedding.is_real.coe_embedding_apply NumberField.ComplexEmbedding.IsReal.coe_embedding_apply -theorem IsReal.place_embedding {φ : K →+* ℂ} (hφ : IsReal φ) : place hφ.embedding = place φ := by - ext x - simp only [place_apply, Real.norm_eq_abs, ← abs_ofReal, norm_eq_abs, hφ.coe_embedding_apply x] -#align number_field.complex_embedding.is_real.place_embedding NumberField.ComplexEmbedding.IsReal.place_embedding - -theorem isReal_conjugate_iff {φ : K →+* ℂ} : IsReal (conjugate φ) ↔ IsReal φ := - IsSelfAdjoint.star_iff -#align number_field.complex_embedding.is_real_conjugate_iff NumberField.ComplexEmbedding.isReal_conjugate_iff - end NumberField.ComplexEmbedding section InfinitePlace @@ -217,12 +212,10 @@ open NumberField variable (K : Type _) [Field K] /-- An infinite place of a number field `K` is a place associated to a complex embedding. -/ -def NumberField.InfinitePlace := - { w : AbsoluteValue K ℝ // ∃ φ : K →+* ℂ, place φ = w } +def NumberField.InfinitePlace := { w : AbsoluteValue K ℝ // ∃ φ : K →+* ℂ, place φ = w } #align number_field.infinite_place NumberField.InfinitePlace -instance [NumberField K] : Nonempty (NumberField.InfinitePlace K) := - Set.instNonemptyElemRange _ +instance [NumberField K] : Nonempty (NumberField.InfinitePlace K) := Set.instNonemptyElemRange _ variable {K} @@ -235,7 +228,9 @@ namespace NumberField.InfinitePlace open NumberField -instance : CoeFun (InfinitePlace K) fun _ => K → ℝ where coe w := w.1 +instance {K : Type _} [Field K] : FunLike (InfinitePlace K) K (fun _ => ℝ) := +{ coe := fun w x => w.1 x + coe_injective' := fun _ _ h => Subtype.eq (AbsoluteValue.ext fun x => congr_fun h x)} instance : MonoidWithZeroHomClass (InfinitePlace K) K ℝ where coe w x := w.1 x @@ -249,54 +244,40 @@ instance : NonnegHomClass (InfinitePlace K) K ℝ where coe_injective' _ _ h := Subtype.eq (AbsoluteValue.ext fun x => congr_fun h x) map_nonneg w _ := w.1.nonneg _ --- Porting note: new -@[ext] -theorem ext {w z : InfinitePlace K} (h : ∀ x, w x = z x) : w = z := FunLike.ext _ _ h - -theorem coe_mk (φ : K →+* ℂ) : ⇑(mk φ) = place φ := rfl -#align number_field.infinite_place.coe_mk NumberField.InfinitePlace.coe_mk - +@[simp] theorem apply (φ : K →+* ℂ) (x : K) : (mk φ) x = Complex.abs (φ x) := rfl #align number_field.infinite_place.apply NumberField.InfinitePlace.apply /-- For an infinite place `w`, return an embedding `φ` such that `w = infinite_place φ` . -/ -noncomputable def embedding (w : InfinitePlace K) : K →+* ℂ := - w.2.choose +noncomputable def embedding (w : InfinitePlace K) : K →+* ℂ := w.2.choose #align number_field.infinite_place.embedding NumberField.InfinitePlace.embedding @[simp] -theorem mk_embedding (w : InfinitePlace K) : mk (embedding w) = w := - Subtype.ext w.2.choose_spec +theorem mk_embedding (w : InfinitePlace K) : mk (embedding w) = w := Subtype.ext w.2.choose_spec #align number_field.infinite_place.mk_embedding NumberField.InfinitePlace.mk_embedding @[simp] -theorem abs_embedding (w : InfinitePlace K) (x : K) : Complex.abs (embedding w x) = w x := - congr_fun (congr_arg (↑) w.2.choose_spec) x -#align number_field.infinite_place.abs_embedding NumberField.InfinitePlace.abs_embedding +theorem mk_conjugate_eq (φ : K →+* ℂ) : mk (ComplexEmbedding.conjugate φ) = mk φ := by + refine FunLike.ext _ _ (fun x => ?_) + rw [apply, apply, ComplexEmbedding.conjugate_coe_eq, Complex.abs_conj] +#align number_field.infinite_place.mk_conjugate_eq NumberField.InfinitePlace.mk_conjugate_eq + +theorem norm_embedding_eq (w : InfinitePlace K) (x : K) : + ‖(embedding w) x‖ = w x := by + nth_rewrite 2 [← mk_embedding w] + rfl theorem eq_iff_eq (x : K) (r : ℝ) : (∀ w : InfinitePlace K, w x = r) ↔ ∀ φ : K →+* ℂ, ‖φ x‖ = r := --- Porting note: original proof was: ⟨fun hw φ => hw (mk φ), fun hφ ⟨w, ⟨φ, rfl⟩⟩ => hφ φ⟩ --- but `rfl` does not work here anymore ⟨fun hw φ => hw (mk φ), by rintro hφ ⟨w, ⟨φ, rfl⟩⟩; exact hφ φ⟩ - #align number_field.infinite_place.eq_iff_eq NumberField.InfinitePlace.eq_iff_eq theorem le_iff_le (x : K) (r : ℝ) : (∀ w : InfinitePlace K, w x ≤ r) ↔ ∀ φ : K →+* ℂ, ‖φ x‖ ≤ r := --- Porting note: original proof was: ⟨fun hw φ => hw (mk φ), fun hφ ⟨w, ⟨φ, rfl⟩⟩ => hφ φ⟩ --- but `rfl` does not work here anymore ⟨fun hw φ => hw (mk φ), by rintro hφ ⟨w, ⟨φ, rfl⟩⟩; exact hφ φ⟩ #align number_field.infinite_place.le_iff_le NumberField.InfinitePlace.le_iff_le -theorem pos_iff {w : InfinitePlace K} {x : K} : 0 < w x ↔ x ≠ 0 := - AbsoluteValue.pos_iff w.1 +theorem pos_iff {w : InfinitePlace K} {x : K} : 0 < w x ↔ x ≠ 0 := AbsoluteValue.pos_iff w.1 #align number_field.infinite_place.pos_iff NumberField.InfinitePlace.pos_iff -@[simp] -theorem mk_conjugate_eq (φ : K →+* ℂ) : mk (ComplexEmbedding.conjugate φ) = mk φ := by - ext x - exact congr_fun (congr_arg (↑) (ComplexEmbedding.place_conjugate φ)) x -#align number_field.infinite_place.mk_conjugate_eq NumberField.InfinitePlace.mk_conjugate_eq - @[simp] theorem mk_eq_iff {φ ψ : K →+* ℂ} : mk φ = mk ψ ↔ φ = ψ ∨ ComplexEmbedding.conjugate φ = ψ := by constructor @@ -315,23 +296,17 @@ theorem mk_eq_iff {φ ψ : K →+* ℂ} : mk φ = mk ψ ↔ φ = ψ ∨ ComplexE suffices ‖φ (ι.symm (x - y))‖ = ‖ψ (ι.symm (x - y))‖ by rw [← this, ← RingEquiv.ofLeftInverse_apply hiφ _, RingEquiv.apply_symm_apply ι _] rfl - exact congr_fun (congr_arg (↑) h₀) _ - cases' - Complex.uniformContinuous_ringHom_eq_id_or_conj φ.fieldRange hlip.uniformContinuous with h h - · left; ext1 x - -- Porting note: original proof was just - -- convert (congr_fun h (ι x)).symm - have := (congr_fun h (ι x)).symm - dsimp at this - convert this - exact (RingEquiv.apply_symm_apply ι.symm x).symm - · right; ext1 x - -- Porting note: original proof was just - -- convert (congr_fun h (ι x)).symm - have := (congr_fun h (ι x)).symm - dsimp at this - convert this - exact (RingEquiv.apply_symm_apply ι.symm x).symm + exact congrFun (congrArg (↑) h₀) _ + cases + Complex.uniformContinuous_ringHom_eq_id_or_conj φ.fieldRange hlip.uniformContinuous with + | inl h => + left; ext1 x + conv_rhs => rw [← hiφ x] + exact (congrFun h (ι x)).symm + | inr h => + right; ext1 x + conv_rhs => rw [← hiφ x] + exact (congrFun h (ι x)).symm · rintro (⟨h⟩ | ⟨h⟩) · exact congr_arg mk h · rw [← mk_conjugate_eq] @@ -339,39 +314,39 @@ theorem mk_eq_iff {φ ψ : K →+* ℂ} : mk φ = mk ψ ↔ φ = ψ ∨ ComplexE #align number_field.infinite_place.mk_eq_iff NumberField.InfinitePlace.mk_eq_iff /-- An infinite place is real if it is defined by a real embedding. -/ -def IsReal (w : InfinitePlace K) : Prop := - ∃ φ : K →+* ℂ, ComplexEmbedding.IsReal φ ∧ mk φ = w +def IsReal (w : InfinitePlace K) : Prop := ∃ φ : K →+* ℂ, ComplexEmbedding.IsReal φ ∧ mk φ = w #align number_field.infinite_place.is_real NumberField.InfinitePlace.IsReal /-- An infinite place is complex if it is defined by a complex (ie. not real) embedding. -/ -def IsComplex (w : InfinitePlace K) : Prop := - ∃ φ : K →+* ℂ, ¬ComplexEmbedding.IsReal φ ∧ mk φ = w +def IsComplex (w : InfinitePlace K) : Prop := ∃ φ : K →+* ℂ, ¬ComplexEmbedding.IsReal φ ∧ mk φ = w #align number_field.infinite_place.is_complex NumberField.InfinitePlace.IsComplex -@[simp] -theorem _root_.NumberField.ComplexEmbeddings.IsReal.embedding_mk {φ : K →+* ℂ} - (h : ComplexEmbedding.IsReal φ) : embedding (mk φ) = φ := by - have := mk_eq_iff.mp (mk_embedding (mk φ)).symm - rwa [ComplexEmbedding.isReal_iff.mp h, or_self_iff, eq_comm] at this -#align number_field.complex_embeddings.is_real.embedding_mk NumberField.ComplexEmbeddings.IsReal.embedding_mk +theorem embedding_mk_eq (φ : K →+* ℂ) : + embedding (mk φ) = φ ∨ embedding (mk φ) = ComplexEmbedding.conjugate φ := by + rw [@eq_comm _ _ φ, @eq_comm _ _ (ComplexEmbedding.conjugate φ), ← mk_eq_iff, mk_embedding] -theorem isReal_iff {w : InfinitePlace K} : IsReal w ↔ ComplexEmbedding.IsReal (embedding w) := by - constructor - · rintro ⟨φ, ⟨hφ, rfl⟩⟩ - rwa [_root_.NumberField.ComplexEmbeddings.IsReal.embedding_mk hφ] - · exact fun h => ⟨embedding w, h, mk_embedding w⟩ +@[simp] +theorem embedding_mk_eq_of_isReal {φ : K →+* ℂ} (h : ComplexEmbedding.IsReal φ) : + embedding (mk φ) = φ := by + have := embedding_mk_eq φ + rwa [ComplexEmbedding.isReal_iff.mp h, or_self] at this +#align number_field.complex_embeddings.is_real.embedding_mk NumberField.InfinitePlace.embedding_mk_eq_of_isReal + +theorem isReal_iff {w : InfinitePlace K} : + IsReal w ↔ ComplexEmbedding.IsReal (embedding w) := by + refine ⟨?_, fun h => ⟨embedding w, h, mk_embedding w⟩⟩ + rintro ⟨φ, ⟨hφ, rfl⟩⟩ + rwa [embedding_mk_eq_of_isReal hφ] #align number_field.infinite_place.is_real_iff NumberField.InfinitePlace.isReal_iff theorem isComplex_iff {w : InfinitePlace K} : IsComplex w ↔ ¬ComplexEmbedding.IsReal (embedding w) := by - constructor - · rintro ⟨φ, ⟨hφ, rfl⟩⟩ - contrapose! hφ - cases mk_eq_iff.mp (mk_embedding (mk φ)) with - | inl h => rwa [← h] - | inr h => rw [← ComplexEmbedding.isReal_conjugate_iff] at hφ - rwa [← h] - · exact fun h => ⟨embedding w, h, mk_embedding w⟩ + refine ⟨?_, fun h => ⟨embedding w, h, mk_embedding w⟩⟩ + rintro ⟨φ, ⟨hφ, rfl⟩⟩ + contrapose! hφ + cases mk_eq_iff.mp (mk_embedding (mk φ)) with + | inl h => rwa [h] at hφ + | inr h => rwa [← ComplexEmbedding.isReal_conjugate_iff, h] at hφ #align number_field.infinite_place.is_complex_iff NumberField.InfinitePlace.isComplex_iff @[simp] @@ -379,43 +354,59 @@ theorem not_isReal_iff_isComplex {w : InfinitePlace K} : ¬IsReal w ↔ IsComple rw [isComplex_iff, isReal_iff] #align number_field.infinite_place.not_is_real_iff_is_complex NumberField.InfinitePlace.not_isReal_iff_isComplex -@[simp] -theorem not_isComplex_iff_isReal {w : InfinitePlace K} : ¬IsComplex w ↔ IsReal w := by - rw [← not_isReal_iff_isComplex, Classical.not_not] -#align number_field.infinite_place.not_is_complex_iff_is_real NumberField.InfinitePlace.not_isComplex_iff_isReal - theorem isReal_or_isComplex (w : InfinitePlace K) : IsReal w ∨ IsComplex w := by rw [← not_isReal_iff_isComplex]; exact em _ #align number_field.infinite_place.is_real_or_is_complex NumberField.InfinitePlace.isReal_or_isComplex -/-- For `w` a real infinite place, return the corresponding embedding as a morphism `K →+* ℝ`. -/ -noncomputable def IsReal.embedding {w : InfinitePlace K} (hw : IsReal w) : K →+* ℝ := - (isReal_iff.mp hw).embedding -#align number_field.infinite_place.is_real.embedding NumberField.InfinitePlace.IsReal.embedding +/-- The real embedding associated to a real infinite place. -/ +noncomputable def embedding_of_isReal {w : InfinitePlace K} (hw : IsReal w) : K →+* ℝ := + ComplexEmbedding.IsReal.embedding (isReal_iff.mp hw) -@[simp, nolint simpNF] -- Porting note: simp cannot prove this -theorem IsReal.place_embedding_apply {w : InfinitePlace K} (hw : IsReal w) (x : K) : - place (IsReal.embedding hw) x = w x := by - rw [IsReal.embedding, ComplexEmbedding.IsReal.place_embedding, ← coe_mk] - exact congr_fun (congr_arg (↑) (mk_embedding w)) x -#align number_field.infinite_place.is_real.place_embedding_apply NumberField.InfinitePlace.IsReal.place_embedding_apply +@[simp] +theorem embedding_of_isReal_apply {w : InfinitePlace K} (hw : IsReal w) (x : K) : + ((embedding_of_isReal hw) x : ℂ) = (embedding w) x := + ComplexEmbedding.IsReal.coe_embedding_apply (isReal_iff.mp hw) x @[simp] -theorem IsReal.abs_embedding_apply {w : InfinitePlace K} (hw : IsReal w) (x : K) : - |IsReal.embedding hw x| = w x := by - rw [← IsReal.place_embedding_apply hw x]; congr 1 -#align number_field.infinite_place.is_real.abs_embedding_apply NumberField.InfinitePlace.IsReal.abs_embedding_apply +theorem isReal_of_mk_isReal {φ : K →+* ℂ} (h : IsReal (mk φ)) : + ComplexEmbedding.IsReal φ := by + contrapose! h + rw [not_isReal_iff_isComplex] + exact ⟨φ, h, rfl⟩ -variable (K) +@[simp] +theorem not_isReal_of_mk_isComplex {φ : K →+* ℂ} (h : IsComplex (mk φ)) : + ¬ ComplexEmbedding.IsReal φ := by + contrapose! h + rw [← not_isReal_iff_isComplex.not, not_not] + exact ⟨φ, h, rfl⟩ + +/-- The multiplicity of an infinite place, that is the number of distinct complex embeddings that +define it, see `card_filter_mk_eq`. -/ +noncomputable def mult (w : InfinitePlace K) : ℕ := if (IsReal w) then 1 else 2 + +theorem card_filter_mk_eq [NumberField K] (w : InfinitePlace K) : + (Finset.univ.filter fun φ => mk φ = w).card = mult w := by + conv_lhs => + congr; congr; ext + rw [← mk_embedding w, mk_eq_iff, ComplexEmbedding.conjugate, star_involutive.eq_iff] + simp_rw [Finset.filter_or, Finset.filter_eq' _ (embedding w), + Finset.filter_eq' _ (ComplexEmbedding.conjugate (embedding w)), + Finset.mem_univ, ite_true, mult] + split_ifs with hw + · rw [ComplexEmbedding.isReal_iff.mp (isReal_iff.mp hw), Finset.union_idempotent, + Finset.card_singleton] + · refine Finset.card_doubleton ?_ + rwa [Ne.def, eq_comm, ← ComplexEmbedding.isReal_iff, ← isReal_iff] /-- The map from real embeddings to real infinite places as an equiv -/ noncomputable def mkReal : - { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } ≃ { w : InfinitePlace K // IsReal w } where - toFun := Subtype.map mk fun φ hφ => ⟨φ, hφ, rfl⟩ - invFun w := ⟨w.1.embedding, isReal_iff.1 w.2⟩ - left_inv φ := Subtype.ext_iff.2 (NumberField.ComplexEmbeddings.IsReal.embedding_mk φ.2) - right_inv w := Subtype.ext_iff.2 (mk_embedding w.1) -#align number_field.infinite_place.mk_real NumberField.InfinitePlace.mkReal + { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } ≃ { w : InfinitePlace K // IsReal w } := by + refine (Equiv.ofBijective (fun φ => ⟨mk φ, ?_⟩) ⟨fun φ ψ h => ?_, fun w => ?_⟩) + · exact ⟨φ, φ.prop, rfl⟩ + · rwa [Subtype.mk.injEq, mk_eq_iff, ComplexEmbedding.isReal_iff.mp φ.prop, or_self, + ← Subtype.ext_iff] at h + · exact ⟨⟨embedding w, isReal_iff.mp w.prop⟩, by simp⟩ /-- The map from nonreal embeddings to complex infinite places -/ noncomputable def mkComplex : @@ -423,111 +414,75 @@ noncomputable def mkComplex : Subtype.map mk fun φ hφ => ⟨φ, hφ, rfl⟩ #align number_field.infinite_place.mk_complex NumberField.InfinitePlace.mkComplex -theorem mkComplex_embedding (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) : - (mkComplex K φ : InfinitePlace K).embedding = φ ∨ - (mkComplex K φ : InfinitePlace K).embedding = ComplexEmbedding.conjugate φ := by - rw [@eq_comm _ _ φ.val, @eq_comm _ _ (ComplexEmbedding.conjugate φ.val), ← mk_eq_iff, - mk_embedding] - rfl -#align number_field.infinite_place.mk_complex_embedding NumberField.InfinitePlace.mkComplex_embedding - @[simp] theorem mkReal_coe (φ : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ }) : - (mkReal K φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl + (mkReal φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl #align number_field.infinite_place.mk_real_coe NumberField.InfinitePlace.mkReal_coe @[simp] theorem mkComplex_coe (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) : - (mkComplex K φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl + (mkComplex φ : InfinitePlace K) = mk (φ : K →+* ℂ) := rfl #align number_field.infinite_place.mk_complex_coe NumberField.InfinitePlace.mkComplex_coe --- @[simp] -- Porting note: not in simpNF -theorem mkReal.apply (φ : { φ : K →+* ℂ // ComplexEmbedding.IsReal φ }) (x : K) : - (mkReal K φ).val x = Complex.abs (φ.val x) := rfl -#align number_field.infinite_place.mk_real.apply NumberField.InfinitePlace.mkReal.apply - --- @[simp] -- Porting note: not in simpNF -theorem mkComplex.apply (φ : { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ }) (x : K) : - (mkComplex K φ).val x = Complex.abs (φ.val x) := rfl -#align number_field.infinite_place.mk_complex.apply NumberField.InfinitePlace.mkComplex.apply - variable [NumberField K] -theorem mkComplex.filter (w : { w : InfinitePlace K // w.IsComplex }) : - (Finset.univ.filter fun φ => mkComplex K φ = w) = - {⟨w.1.embedding, isComplex_iff.1 w.2⟩, - ⟨ComplexEmbedding.conjugate w.1.embedding, - ComplexEmbedding.isReal_conjugate_iff.not.2 (isComplex_iff.1 w.2)⟩} := by - ext φ - simp_rw [Finset.mem_filter, Finset.mem_insert, Finset.mem_singleton, - @Subtype.ext_iff_val (InfinitePlace K), @Subtype.ext_iff_val (K →+* ℂ), @eq_comm _ φ.val, ← - mk_eq_iff, mk_embedding, @eq_comm _ _ w.val] - simp only [Finset.mem_univ, mkComplex_coe, true_and] -#align number_field.infinite_place.mk_complex.filter NumberField.InfinitePlace.mkComplex.filter - -theorem mkComplex.filter_card (w : { w : InfinitePlace K // w.IsComplex }) : - (Finset.univ.filter fun φ => mkComplex K φ = w).card = 2 := by - rw [mkComplex.filter] - exact Finset.card_doubleton - (Subtype.mk_eq_mk.not.2 <| - ne_comm.1 <| ComplexEmbedding.isReal_iff.not.1 <| isComplex_iff.1 w.2) -#align number_field.infinite_place.mk_complex.filter_card NumberField.InfinitePlace.mkComplex.filter_card - noncomputable instance NumberField.InfinitePlace.fintype : Fintype (InfinitePlace K) := Set.fintypeRange _ #align number_field.infinite_place.number_field.infinite_place.fintype NumberField.InfinitePlace.NumberField.InfinitePlace.fintype +open scoped BigOperators + /-- The infinite part of the product formula : for `x ∈ K`, we have `Π_w ‖x‖_w = |norm(x)|` where `‖·‖_w` is the normalized absolute value for `w`. -/ theorem prod_eq_abs_norm (x : K) : - (Finset.univ.prod fun w : InfinitePlace K => ite w.IsReal (w x) (w x ^ 2)) = - abs (Algebra.norm ℚ x) := by + ∏ w : InfinitePlace K, w x ^ mult w = abs (Algebra.norm ℚ x) := by convert (congr_arg Complex.abs (@Algebra.norm_eq_prod_embeddings ℚ _ _ _ _ ℂ _ _ _ _ _ x)).symm · rw [map_prod, ← Equiv.prod_comp' RingHom.equivRatAlgHom (fun f => Complex.abs (f x)) - (fun φ => Complex.abs (φ x)) fun _ => by simp [RingHom.equivRatAlgHom_apply]; rfl] - conv => - rhs; congr; next => skip - ext f - rw [show Complex.abs (f x) = - ite (ComplexEmbedding.IsReal f) (Complex.abs (f x)) (Complex.abs (f x)) by simp] - rw [Finset.prod_ite, Finset.prod_ite] - refine congr (congr_arg Mul.mul ?_) ?_ - · rw [← Finset.prod_subtype_eq_prod_filter, ← Finset.prod_subtype_eq_prod_filter] - convert (Equiv.prod_comp' (mkReal K) (fun φ => Complex.abs (φ.val x)) - (fun w => w.val x) _).symm - any_goals ext; simp only [Finset.mem_subtype, Finset.mem_univ] - exact fun φ => mkReal.apply K φ x - · rw [Finset.filter_congr fun (w : InfinitePlace K) _ => @not_isReal_iff_isComplex K _ w, - ← Finset.prod_subtype_eq_prod_filter, ← Finset.prod_subtype_eq_prod_filter] - convert Finset.prod_fiberwise Finset.univ (fun φ => mkComplex K φ) - (fun φ => Complex.abs (φ.val x)) using 2 with w - · ext; simp only [Finset.mem_subtype, Finset.mem_univ, not_isReal_iff_isComplex] - · rw [@Finset.prod_congr _ _ _ _ _ (fun _ => w.val x) _ (Eq.refl _) fun φ hφ => - (mkComplex.apply K φ x).symm.trans ?_, Finset.prod_const, mkComplex.filter_card K w] - rw [congr_arg Subtype.val (Finset.mem_filter.1 hφ).2] - · ext; simp only [Finset.mem_subtype, Finset.mem_univ] + (fun φ => Complex.abs (φ x)) fun _ => by simp [RingHom.equivRatAlgHom_apply]; rfl] + rw [← Finset.prod_fiberwise Finset.univ (fun φ => mk φ) (fun φ => Complex.abs (φ x))] + have : ∀ w : InfinitePlace K, ∀ φ ∈ Finset.filter (fun a ↦ mk a = w) Finset.univ, + Complex.abs (φ x) = w x := by + intro _ _ hφ + rw [← (Finset.mem_filter.mp hφ).2] + rfl + simp_rw [Finset.prod_congr rfl (this _), Finset.mem_univ, forall_true_left, Finset.prod_const, + card_filter_mk_eq] · rw [eq_ratCast, Rat.cast_abs, ← Complex.abs_ofReal, Complex.ofReal_rat_cast] #align number_field.infinite_place.prod_eq_abs_norm NumberField.InfinitePlace.prod_eq_abs_norm -open Fintype +open Fintype FiniteDimensional theorem card_real_embeddings : - card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } = card { w : InfinitePlace K // IsReal w } := - by convert (Fintype.ofEquiv_card (mkReal K)).symm + card { φ : K →+* ℂ // ComplexEmbedding.IsReal φ } + = card { w : InfinitePlace K // IsReal w } := Fintype.card_congr mkReal #align number_field.infinite_place.card_real_embeddings NumberField.InfinitePlace.card_real_embeddings theorem card_complex_embeddings : card { φ : K →+* ℂ // ¬ComplexEmbedding.IsReal φ } = 2 * card { w : InfinitePlace K // IsComplex w } := by - rw [Fintype.card, Fintype.card, mul_comm, ← Algebra.id.smul_eq_mul, ← Finset.sum_const] - conv => - rhs; congr; next => skip - ext x - rw [← mkComplex.filter_card K x] - simp_rw [Finset.card_eq_sum_ones] - exact (Finset.sum_fiberwise Finset.univ (fun φ => mkComplex K φ) fun _ => 1).symm + suffices ∀ w : { w : InfinitePlace K // IsComplex w }, (Finset.univ.filter + fun φ : { φ // ¬ ComplexEmbedding.IsReal φ } => mkComplex φ = w).card = 2 by + rw [Fintype.card, Finset.card_eq_sum_ones, ← Finset.sum_fiberwise _ (fun φ => mkComplex φ)] + simp_rw [Finset.sum_const, this, smul_eq_mul, mul_one, Fintype.card, Finset.card_eq_sum_ones, + Finset.mul_sum] + rintro ⟨w, hw⟩ + convert card_filter_mk_eq w + · rw [← Fintype.card_subtype, ← Fintype.card_subtype] + refine Fintype.card_congr (Equiv.ofBijective ?_ ⟨fun _ _ h => ?_, fun ⟨φ, hφ⟩ => ?_⟩) + · exact fun ⟨φ, hφ⟩ => ⟨φ.val, by rwa [Subtype.ext_iff] at hφ⟩ + · rwa [Subtype.mk_eq_mk, ← Subtype.ext_iff, ← Subtype.ext_iff] at h + · refine ⟨⟨⟨φ, not_isReal_of_mk_isComplex (hφ.symm ▸ hw)⟩, ?_⟩, rfl⟩ + rwa [Subtype.ext_iff, mkComplex_coe] + · simp_rw [mult, not_isReal_iff_isComplex.mpr hw] #align number_field.infinite_place.card_complex_embeddings NumberField.InfinitePlace.card_complex_embeddings +theorem card_add_two_mul_card_eq_rank : + card { w : InfinitePlace K // IsReal w } + 2 * card { w : InfinitePlace K // IsComplex w } = + finrank ℚ K := by + rw [← card_real_embeddings, ← card_complex_embeddings] + rw [Fintype.card_subtype_compl, ← Embeddings.card K ℂ, Nat.add_sub_of_le] + exact Fintype.card_subtype_le _ + end NumberField.InfinitePlace end InfinitePlace From fd620f4c94ecdd51aac11e1038e9b2404f2e88f6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 4 Aug 2023 08:54:22 +0900 Subject: [PATCH 34/67] Fix align_import --- Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean | 6 +----- 1 file changed, 1 insertion(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 7635c3888d686..813cca031d070 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -2,15 +2,11 @@ Copyright (c) 2022 Xavier Roblot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Xavier Roblot - -! This file was ported from Lean 3 source module number_theory.number_field.canonical_embedding -! leanprover-community/mathlib commit 60da01b41bbe4206f05d34fd70c8dd7498717a30 -! Please do not edit these lines, except to modify the commit id -! if you have ported upstream changes. -/ import Mathlib.NumberTheory.NumberField.Embeddings import Mathlib.RingTheory.Discriminant +#align_import number_theory.number_field.canonical_embedding from "leanprover-community/mathlib"@"60da01b41bbe4206f05d34fd70c8dd7498717a30" /-! # Canonical embedding of a number field From b3030fc794a15cd0ee405f196a604f793b062e27 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 4 Aug 2023 18:00:56 +0900 Subject: [PATCH 35/67] add Finite_bounded_inter_isClosed --- Mathlib/Topology/MetricSpace/Basic.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Topology/MetricSpace/Basic.lean b/Mathlib/Topology/MetricSpace/Basic.lean index e0cb436fe261d..a79d453136111 100644 --- a/Mathlib/Topology/MetricSpace/Basic.lean +++ b/Mathlib/Topology/MetricSpace/Basic.lean @@ -3027,6 +3027,12 @@ theorem uniformEmbedding_bot_of_pairwise_le_dist {β : Type _} {ε : ℝ} (hε : uniformEmbedding_of_spaced_out (dist_mem_uniformity hε) <| by simpa using hf #align metric.uniform_embedding_bot_of_pairwise_le_dist Metric.uniformEmbedding_bot_of_pairwise_le_dist +theorem Finite_bounded_inter_isClosed [ProperSpace α] {K s : Set α} [DiscreteTopology s] + (hK : Metric.Bounded K) (hs : IsClosed s) : Set.Finite (K ∩ s) := by + refine Set.Finite.subset (IsCompact.finite ?_ ?_) (Set.inter_subset_inter_left s subset_closure) + · exact IsCompact.inter_right (Metric.Bounded.isCompact_closure hK) hs + · exact DiscreteTopology.of_subset inferInstance (Set.inter_subset_right _ s) + end Metric /-- Build a new metric space from an old one where the bundled uniform structure is provably From 9bc4c5cc3bf1e870ab8997e30ed9e872df156f64 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 4 Aug 2023 18:01:18 +0900 Subject: [PATCH 36/67] Switch to DiscreteTopology --- Mathlib/Algebra/Module/Zlattice.lean | 47 +++++++++++++++------------- 1 file changed, 25 insertions(+), 22 deletions(-) diff --git a/Mathlib/Algebra/Module/Zlattice.lean b/Mathlib/Algebra/Module/Zlattice.lean index 46f85b3489979..3601c7b0bceac 100644 --- a/Mathlib/Algebra/Module/Zlattice.lean +++ b/Mathlib/Algebra/Module/Zlattice.lean @@ -19,7 +19,7 @@ subgroup of `E` such that `L` spans `E` over `K`. A `ℤ`-lattice `L` can be defined in two ways: * For `b` a basis of `E`, then `L = Submodule.span ℤ (Set.range b)` is a ℤ-lattice of `E` * As an `AddSubgroup E` with the additional properties: - * `∀ r : ℝ, (L ∩ Metric.closedBall 0 r).finite`, that is `L` is discrete + * `DiscreteTopology L`, that is `L` is discrete * `Submodule.span ℝ (L : Set E) = ⊤`, that is `L` spans `E` over `K`. Results about the first point of view are in the `Zspan` namespace and results about the second @@ -303,9 +303,8 @@ section Zlattice open Submodule variable (K : Type _) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K] -variable {E : Type _} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] -variable {L : AddSubgroup E} -variable (hd : ∀ r : ℝ, ((L : Set E) ∩ (Metric.closedBall 0 r)).Finite) +variable {E : Type _} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] [ProperSpace E] +variable {L : AddSubgroup E} [DiscreteTopology L] variable (hs : span K (L : Set E) = ⊤) theorem Zlattice.FG : AddSubgroup.FG L := by @@ -326,26 +325,28 @@ theorem Zlattice.FG : AddSubgroup.FG L := by have : Fintype s := Set.Finite.fintype h_lind.finite refine Set.Finite.of_finite_image (f := ((↑) : _ → E) ∘ Zspan.QuotientEquiv b) ?_ (Function.Injective.injOn (Subtype.coe_injective.comp (Zspan.QuotientEquiv b).injective) _) - obtain ⟨C, hC⟩ := Metric.Bounded.subset_ball (Zspan.fundamentalDomain_bounded b) 0 - refine Set.Finite.subset (hd C) ?_ + have : Set.Finite ((Zspan.fundamentalDomain b) ∩ L) := + Metric.Finite_bounded_inter_isClosed (Zspan.fundamentalDomain_bounded b) inferInstance + refine Set.Finite.subset this ?_ rintro _ ⟨_, ⟨⟨x, ⟨h_mem, rfl⟩⟩, rfl⟩⟩ rw [Function.comp_apply, mkQ_apply, Zspan.quotientEquiv_apply_mk, Zspan.fract_restrict_apply] - refine ⟨?_, hC (Zspan.fract_mem_fundamentalDomain b x)⟩ - rw [Zspan.fract, SetLike.mem_coe, sub_eq_add_neg] - refine AddSubgroup.add_mem _ h_mem - (neg_mem (Set.mem_of_subset_of_mem ?_ (Subtype.mem (Zspan.floor b x)))) - rw [show (L : Set E) = AddSubgroup.toIntSubmodule L by rfl] - rw [SetLike.coe_subset_coe, Basis.coe_mk, Subtype.range_coe_subtype, Set.setOf_mem_eq] - exact span_le.mpr h_incl + refine ⟨?_, ?_⟩ + · exact Zspan.fract_mem_fundamentalDomain b x + · rw [Zspan.fract, SetLike.mem_coe, sub_eq_add_neg] + refine AddSubgroup.add_mem _ h_mem + (neg_mem (Set.mem_of_subset_of_mem ?_ (Subtype.mem (Zspan.floor b x)))) + rw [show (L : Set E) = AddSubgroup.toIntSubmodule L by rfl] + rw [SetLike.coe_subset_coe, Basis.coe_mk, Subtype.range_coe_subtype, Set.setOf_mem_eq] + exact span_le.mpr h_incl · -- `span ℤ s` is finitely generated because `s` is finite rw [ker_mkQ, inf_of_le_right (span_le.mpr h_incl)] exact fg_span (LinearIndependent.finite h_lind) theorem Zlattice.module_finite : Module.Finite ℤ L := - Module.Finite.iff_addGroup_fg.mpr ((AddGroup.fg_iff_addSubgroup_fg L).mpr (FG K hd hs)) + Module.Finite.iff_addGroup_fg.mpr ((AddGroup.fg_iff_addSubgroup_fg L).mpr (FG K hs)) theorem Zlattice.module_free : Module.Free ℤ L := by - have : Module.Finite ℤ L := module_finite K hd hs + have : Module.Finite ℤ L := module_finite K hs have : Module ℚ E := Module.compHom E (algebraMap ℚ K) have : NoZeroSMulDivisors ℤ E := RatModule.noZeroSMulDivisors have : NoZeroSMulDivisors ℤ L := by @@ -357,8 +358,8 @@ open FiniteDimensional theorem Zlattice.rank : finrank ℤ L = finrank K E := by classical - have : Module.Finite ℤ L := module_finite K hd hs - have : Module.Free ℤ L := module_free K hd hs + have : Module.Finite ℤ L := module_finite K hs + have : Module.Free ℤ L := module_free K hs have : Module ℚ E := Module.compHom E (algebraMap ℚ K) let b₀ := Module.Free.chooseBasis ℤ L -- Let `b` be a `ℤ`-basis of `L` formed of vectors of `E` @@ -409,19 +410,21 @@ theorem Zlattice.rank : finrank ℤ L = finrank K E := by (linearIndependent_insert (Set.not_mem_of_mem_diff hv)), not_and, not_not] intro _ -- But that follows from the fact that there exist `n, m : ℕ`, `n ≠ m` - -- such that `(n - m) • v ∈ span ℤ e` which is true since `n ↦ Zspan.fract e (n • v)` + -- such that `(n - m) • v ∈ span ℤ e` which is true since `n ↦ Zspan.fract e (n • v)` -- takes value into the finite set `fundamentalDomain e ∩ L` - have : Set.MapsTo (fun n : ℤ => Zspan.fract e (n • v)) Set.univ - (L ∩ Metric.closedBall 0 (∑ i, ‖e i‖)) := by + have h_mapsto : Set.MapsTo (fun n : ℤ => Zspan.fract e (n • v)) Set.univ + (Metric.closedBall 0 (∑ i, ‖e i‖) ∩ (L : Set E)) := by rw [Set.mapsTo_inter, Set.maps_univ_to, Set.maps_univ_to] - refine ⟨fun _ => ?_, fun _ => mem_closedBall_zero_iff.mpr (Zspan.norm_fract_le e _)⟩ + refine ⟨fun _ => mem_closedBall_zero_iff.mpr (Zspan.norm_fract_le e _), fun _ => ?_⟩ · change _ ∈ AddSubgroup.toIntSubmodule L rw [← h_spanL] refine sub_mem ?_ ?_ · exact zsmul_mem (subset_span (Set.diff_subset _ _ hv)) _ · exact span_mono (by simp [ht_inc]) (coe_mem _) + have h_finite : Set.Finite (Metric.closedBall 0 (∑ i, ‖e i‖) ∩ (L : Set E)) := + Metric.Finite_bounded_inter_isClosed Metric.bounded_closedBall inferInstance obtain ⟨n, -, m, -, h_neq, h_eq⟩ := Set.Infinite.exists_ne_map_eq_of_mapsTo - Set.infinite_univ this (hd (∑ i, ‖e i‖)) + Set.infinite_univ h_mapsto h_finite have h_nz : (-n + m : ℚ) ≠ 0 := by rwa [Ne.def, add_eq_zero_iff_eq_neg.not, neg_inj, Rat.coe_int_inj, ← Ne.def] apply (smul_mem_iff _ h_nz).mp From bfd868e33a1c062503bf8ef554048bf0f3749452 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 4 Aug 2023 18:10:40 +0900 Subject: [PATCH 37/67] Line too long --- Mathlib/Algebra/Module/Zlattice.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Module/Zlattice.lean b/Mathlib/Algebra/Module/Zlattice.lean index 3601c7b0bceac..bff50260929a7 100644 --- a/Mathlib/Algebra/Module/Zlattice.lean +++ b/Mathlib/Algebra/Module/Zlattice.lean @@ -303,8 +303,8 @@ section Zlattice open Submodule variable (K : Type _) [NormedLinearOrderedField K] [HasSolidNorm K] [FloorRing K] -variable {E : Type _} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] [ProperSpace E] -variable {L : AddSubgroup E} [DiscreteTopology L] +variable {E : Type _} [NormedAddCommGroup E] [NormedSpace K E] [FiniteDimensional K E] +variable [ProperSpace E] {L : AddSubgroup E} [DiscreteTopology L] variable (hs : span K (L : Set E) = ⊤) theorem Zlattice.FG : AddSubgroup.FG L := by From 923f9c9c5c6a9d1b26067d4682742caf4f1d68ae Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 4 Aug 2023 18:49:27 +0900 Subject: [PATCH 38/67] merge --- Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean | 9 ++++++++- 1 file changed, 8 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 9e41625a3ac27..e7961fbe914e5 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -30,7 +30,7 @@ radius is finite. complex conjugation). * `exists_ne_zero_mem_ringOfIntegers_lt`: let `f : InfinitePlace K → ℝ≥0`, if the product -`∏_w f w` is large enough, proves that there exists a nonzero algebraic integer `a` such +`∏ w, f w` is large enough, proves that there exists a nonzero algebraic integer `a` such that `w a < f w` for all infinite places `w`. ## Tags @@ -330,6 +330,8 @@ theorem constant_factor_lt_top : (constant_factor K) < ⊤ := by · exact ne_of_lt (pow_lt_top measure_ball_lt_top _) set_option maxHeartbeats 400000 in +/-- The volume of `(convex_body K f)` where `convex_body K f` is the set of points `x` such that +`‖x w‖ < f w` for all infinite places `w`. -/ theorem convex_body_volume : volume (convex_body K f) = (constant_factor K) * ∏ w, (f w) ^ (mult w) := by rw [volume_eq_prod, convex_body, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi] @@ -393,6 +395,11 @@ theorem minkowski_bound_lt_top : minkowski_bound K < ⊤ := by · exact ne_of_lt (Zspan.fundamentalDomain_bounded (latticeBasis K)).measure_lt_top · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) +/-- Assume that `f : InfinitePlace K → ℝ≥0` is such that +`minkowski_bound K < volume (convex_body K f)` where `convex_body K f` is the set of points `x` +such that `‖x w‖ < f w` for all infinite places `w` (see `convex_body_volume` for the computation +of this volume), then there exists a nonzero algebraic integer `a` in `𝓞 K` such that +`w a < f w` for all infinite places `w`. -/ theorem exists_ne_zero_mem_ringOfIntegers_lt (h : minkowski_bound K < volume (convex_body K f)) : ∃ (a : 𝓞 K), a ≠ 0 ∧ ∀ w : InfinitePlace K, w a < f w := by have : @IsAddHaarMeasure (E K) _ _ _ volume := prod.instIsAddHaarMeasure volume volume From 3dc3436e6b42ca726b0d976d70eee8078abf9787 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 8 Aug 2023 11:11:19 +0900 Subject: [PATCH 39/67] add main statements --- Mathlib/NumberTheory/NumberField/Units.lean | 140 ++++++++++++++------ 1 file changed, 99 insertions(+), 41 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 1aafe1454a976..c7250abd2cf75 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -24,13 +24,13 @@ number field, units -/ -- See: https://github.com/leanprover/lean4/issues/2220 -local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) +local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) open scoped NumberField noncomputable section -open NumberField Units +open NumberField Units BigOperators section Rat @@ -147,7 +147,6 @@ theorem rootsOfUnity_eq_torsion [NumberField K] : end torsion namespace dirichlet - -- This section is devoted to the proof of Dirichlet's unit theorem -- We define a group morphism from `(𝓞 K)ˣ` to `{w : InfinitePlace K // w ≠ w₀} → ℝ` where `w₀` -- is a distinguished (arbitrary) infinite place, prove that its kernel is the torsion subgroup @@ -159,7 +158,7 @@ namespace dirichlet -- (see ` unit_lattice_span_eq_top`); this is the main part of the proof, see the section -- `span_top` below for more details. -open scoped Classical BigOperators +open scoped Classical variable [NumberField K] @@ -168,23 +167,25 @@ variable {K} /-- The distinguished infinite place. -/ def w₀ : InfinitePlace K := (inferInstance : Nonempty (InfinitePlace K)).some +variable (K) + /-- The logarithmic embedding of the units (seen as an `Additive` group). -/ def log_embedding : Additive ((𝓞 K)ˣ) →+ ({w : InfinitePlace K // w ≠ w₀} → ℝ) := { toFun := fun x w => mult w.val * Real.log (w.val (Additive.toMul x)) map_zero' := by simp; rfl map_add' := by intro _ _ - simp only [ne_eq, toMul_add, val_mul, Submonoid.coe_mul, Subsemiring.coe_toSubmonoid, - Subalgebra.coe_toSubsemiring, map_mul, map_eq_zero, ZeroMemClass.coe_eq_zero, ne_zero, - not_false_eq_true, Real.log_mul, mul_add] + simp [ne_eq, toMul_add, map_mul, map_eq_zero, coe_ne_zero, Real.log_mul, mul_add] rfl } +variable {K} + @[simp] theorem log_embedding_component (x : (𝓞 K)ˣ) (w : {w : InfinitePlace K // w ≠ w₀}) : - (log_embedding x) w = mult w.val * Real.log (w.val x) := rfl + (log_embedding K x) w = mult w.val * Real.log (w.val x) := rfl theorem log_embedding_sum_component (x : (𝓞 K)ˣ) : - ∑ w, log_embedding x w = - mult (w₀ : InfinitePlace K) * Real.log (w₀ (x : K)) := by + ∑ w, log_embedding K x w = - mult (w₀ : InfinitePlace K) * Real.log (w₀ (x : K)) := by have h := congrArg Real.log (prod_eq_abs_norm (x : K)) rw [show |(Algebra.norm ℚ) (x : K)| = 1 from isUnit_iff_norm.mp x.isUnit, Rat.cast_one, Real.log_one, Real.log_prod] at h @@ -207,7 +208,7 @@ theorem mult_log_place_eq_zero {x : (𝓞 K)ˣ} {w : InfinitePlace K} : rw [mult]; split_ifs <;> norm_num theorem log_embedding_eq_zero_iff (x : (𝓞 K)ˣ) : - log_embedding x = 0 ↔ x ∈ torsion K := by + log_embedding K x = 0 ↔ x ∈ torsion K := by rw [mem_torsion] refine ⟨fun h w => ?_, fun h => ?_⟩ · by_cases hw : w = w₀ @@ -220,13 +221,13 @@ theorem log_embedding_eq_zero_iff (x : (𝓞 K)ˣ) : · ext w rw [log_embedding_component, h w.val, Real.log_one, mul_zero, Pi.zero_apply] -theorem log_embedding_component_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖log_embedding x‖ ≤ r) - (w : {w : InfinitePlace K // w ≠ w₀}) : |log_embedding x w| ≤ r := by +theorem log_embedding_component_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖log_embedding K x‖ ≤ r) + (w : {w : InfinitePlace K // w ≠ w₀}) : |log_embedding K x w| ≤ r := by lift r to NNReal using hr simp_rw [Pi.norm_def, NNReal.coe_le_coe, Finset.sup_le_iff, ← NNReal.coe_le_coe] at h exact h w (Finset.mem_univ _) -theorem log_le_of_log_embedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖log_embedding x‖ ≤ r) +theorem log_le_of_log_embedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖log_embedding K x‖ ≤ r) (w : InfinitePlace K) : |Real.log (w x)| ≤ (Fintype.card (InfinitePlace K)) * r := by have tool : ∀ x : ℝ, 0 ≤ x → x ≤ mult w * x := fun x hx => by nth_rw 1 [← one_mul x] @@ -255,7 +256,7 @@ variable (K) /-- The lattice formed by the image of the logarithmic embedding. -/ noncomputable def unit_lattice : AddSubgroup ({w : InfinitePlace K // w ≠ w₀} → ℝ) := - AddSubgroup.map log_embedding ⊤ + AddSubgroup.map (log_embedding K) ⊤ theorem unit_lattice_inter_ball_finite (r : ℝ) : ((unit_lattice K : Set ({ w : InfinitePlace K // w ≠ w₀} → ℝ)) ∩ @@ -266,7 +267,7 @@ theorem unit_lattice_inter_ball_finite (r : ℝ) : exact Set.inter_empty _ · suffices Set.Finite {x : (𝓞 K)ˣ | IsIntegral ℤ (x : K) ∧ ∀ (φ : K →+* ℂ), ‖φ x‖ ≤ Real.exp ((Fintype.card (InfinitePlace K)) * r)} by - refine (Set.Finite.image log_embedding this).subset ?_ + refine (Set.Finite.image (log_embedding K) this).subset ?_ rintro _ ⟨⟨x, ⟨_, rfl⟩⟩, hx⟩ refine ⟨x, ⟨x.val.prop, (le_iff_le _ _).mp (fun w => (Real.log_le_iff_le_exp ?_).mp ?_)⟩, rfl⟩ · exact pos_iff.mpr (coe_ne_zero x) @@ -279,22 +280,21 @@ theorem unit_lattice_inter_ball_finite (r : ℝ) : exact ⟨h_int, h_le⟩ section span_top - -- To prove that the span over `ℝ` of the `unit_lattice` is equal to the full space, we construct --- for each infinite place `w₁ ≠ w₀` an unit `u_w₁` of `K` such that, for all infinite place +-- for each infinite place `w₁ ≠ w₀` an unit `u_w₁` of `K` such that, for all infinite places -- `w` such that `w ≠ w₁`, we have `Real.log w (u_w₁) < 0` (and thus `Real.log w₁ (u_w₁) > 0`). -- It follows then from a determinant computation (using `Matrix.det_ne_zero_of_neg`) that the -- image by `log_embedding` of these units is a `ℝ`-linearly independent family. -- The unit `u_w₁` is obtained by construction a sequence `seq n` of nonzero algebraic integers --- that is strictly decreasing at infinite places distinct from `w₁` and of bounded norms. Since --- there are finitely many ideals of bounded norms, there exists two terms in the sequence defining --- the same ideal and their quotient is the unit `u_w₁` (see `exists_unit`). +-- that is strictly decreasing at infinite places distinct from `w₁` and of norm `≤ B`. Since +-- there are finitely many ideals of norm `≤ B`, there exists two terms in the sequence defining +-- the same ideal and their quotient is the desired unit `u_w₁` (see `exists_unit`). open NumberField.mixedEmbedding NNReal variable (w₁ : InfinitePlace K) {B : ℕ} (hB : minkowski_bound K < (constant_factor K) * B) -/-- This result shows that there always exists a next term of the sequence. -/ +/-- This result shows that there always exists a next term in the sequence. -/ theorem seq.next {x : 𝓞 K} (hx : x ≠ 0) : ∃ y : 𝓞 K, y ≠ 0 ∧ (∀ w, w ≠ w₁ → w y < w x) ∧ |Algebra.norm ℚ (y : K)| ≤ B := by let f : InfinitePlace K → ℝ≥0 := @@ -351,7 +351,7 @@ theorem seq.antitone {n m : ℕ} (h : n < m) : refine lt_trans ?_ (m_ih hr w hw) exact (seq.next K w₁ hB (seq K w₁ hB m).prop).choose_spec.2.1 w hw -/-- The terms of the sequence have bounded norms. -/ +/-- The terms of the sequence have norm bounded by `B`. -/ theorem seq.norm_bdd (n : ℕ) : 1 ≤ Int.natAbs (Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K)) ∧ Int.natAbs (Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K)) ≤ B := by @@ -407,7 +407,7 @@ theorem unit_lattice_span_eq_top : -- The standard basis let B := Pi.basisFun ℝ {w : InfinitePlace K // w ≠ w₀} -- The family of units constructed above - let v := fun w : { w : InfinitePlace K // w ≠ w₀ } => log_embedding ((exists_unit K w).choose) + let v := fun w : { w : InfinitePlace K // w ≠ w₀ } => log_embedding K ((exists_unit K w).choose) -- To prove the result, it is enough to prove that the family `v` is linearly independent suffices B.det v ≠ 0 by rw [← isUnit_iff_ne_zero, ← is_basis_iff_det] at this @@ -434,31 +434,89 @@ def _root_.NumberField.Units.rank : ℕ := Fintype.card (InfinitePlace K) - 1 open FiniteDimensional -theorem unit_lattice_discrete : DiscreteTopology (unit_lattice K) := by - refine discreteTopology_of_open_singleton_zero ?_ - refine isOpen_singleton_of_finite_mem_nhds 0 (s := Metric.closedBall 0 1) ?_ ?_ - exact Metric.closedBall_mem_nhds _ (by norm_num) - refine Set.Finite.of_finite_image ?_ (Set.injOn_of_injective Subtype.val_injective _) - convert unit_lattice_inter_ball_finite K 1 - ext x - refine ⟨?_, fun ⟨hx1, hx2⟩ => ⟨⟨x, hx1⟩, hx2, rfl⟩⟩ - rintro ⟨x, hx, rfl⟩ - exact ⟨Subtype.mem x, hx⟩ - theorem rank_space : - finrank ℝ ({w : InfinitePlace K // w ≠ w₀} → ℝ) = Units.rank K := by + finrank ℝ ({w : InfinitePlace K // w ≠ w₀} → ℝ) = rank K := by simp only [finrank_fintype_fun_eq_card, Fintype.card_subtype_compl, Fintype.card_ofSubsingleton, rank] -theorem unit_lattice_moduleFree : Module.Free ℤ (unit_lattice K) := by - have := unit_lattice_discrete K - exact Zlattice.module_free ℝ (unit_lattice_span_eq_top K) +theorem unit_lattice_moduleFree : Module.Free ℤ (unit_lattice K) := +Zlattice.module_free ℝ ((unit_lattice_inter_ball_finite K)) (unit_lattice_span_eq_top K) + +theorem unit_lattice_moduleFinite : Module.Finite ℤ (unit_lattice K) := +Zlattice.module_finite ℝ ((unit_lattice_inter_ball_finite K)) (unit_lattice_span_eq_top K) -theorem unit_lattice.rank : finrank ℤ (unit_lattice K) = Units.rank K := by - have := unit_lattice_discrete K +theorem unit_lattice_rank : finrank ℤ (unit_lattice K) = rank K := by rw [← rank_space] - exact Zlattice.rank ℝ (unit_lattice_span_eq_top K) + exact Zlattice.rank ℝ ((unit_lattice_inter_ball_finite K)) (unit_lattice_span_eq_top K) end dirichlet +variable [NumberField K] + +def basis_mod_torsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by + let f : (dirichlet.unit_lattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (torsion K)) := by + refine AddEquiv.toIntLinearEquiv ?_ + rw [dirichlet.unit_lattice, ← AddMonoidHom.range_eq_map (dirichlet.log_embedding K)] + refine (QuotientAddGroup.quotientKerEquivRange (dirichlet.log_embedding K)).symm.trans ?_ + refine (QuotientAddGroup.quotientAddEquivOfEq ?_).trans + (QuotientAddGroup.quotientKerEquivOfSurjective + (MonoidHom.toAdditive (QuotientGroup.mk' (torsion K))) (fun x => ?_)) + · ext + rw [AddMonoidHom.mem_ker, AddMonoidHom.mem_ker, dirichlet.log_embedding_eq_zero_iff, + MonoidHom.toAdditive_apply_apply, ofMul_eq_zero, QuotientGroup.mk'_apply, + QuotientGroup.eq_one_iff] + rfl + · refine ⟨Additive.ofMul x.out', ?_⟩ + simp only [MonoidHom.toAdditive_apply_apply, toMul_ofMul, QuotientGroup.mk'_apply, + QuotientGroup.out_eq'] + rfl + have : Module.Free ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := + (dirichlet.unit_lattice_moduleFree K).of_equiv' f + have : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by + have := dirichlet.unit_lattice_moduleFinite K + exact Module.Finite.equiv f + have : FiniteDimensional.finrank ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) = rank K := by + rw [← LinearEquiv.finrank_eq f, dirichlet.unit_lattice_rank] + refine Basis.reindex (Module.Free.chooseBasis ℤ _) (Fintype.equivOfCardEq ?_) + rw [← FiniteDimensional.finrank_eq_card_chooseBasisIndex, this, Fintype.card_fin] + +/-- A fundamental system of units of `K`. -/ +def fund_system : Fin (rank K) → (𝓞 K)ˣ := + fun i => Quotient.out' (Additive.toMul (basis_mod_torsion K i)) + +theorem fun_eq_repr {x ζ : (𝓞 K)ˣ} {f : Fin (rank K) → ℤ} (hζ : ζ ∈ torsion K) + (h : x = ζ * ∏ i, (fund_system K i) ^ (f i)) : + f = (basis_mod_torsion K).repr (Additive.ofMul ↑x) := by + suffices Additive.ofMul ↑x = ∑ i, (f i) • (basis_mod_torsion K i) by + rw [← (basis_mod_torsion K).repr_sum_self f, ← this] + calc + Additive.ofMul ↑x = ∑ i, (f i) • Additive.ofMul ↑(fund_system K i) := by + rw [h, QuotientGroup.mk_mul, (QuotientGroup.eq_one_iff _).mpr hζ, one_mul, + QuotientGroup.mk_prod, ofMul_prod]; rfl + _ = ∑ i, (f i) • (basis_mod_torsion K i) := by + simp_rw [fund_system, QuotientGroup.out_eq', ofMul_toMul] + +set_option maxHeartbeats 300000 in +/-- Any unit `x` of `𝓞 K` can be written uniquely as a root of unity times the product of powers +of the units of the fundamental system. -/ +theorem exist_unique_eq_mul_prod (x : (𝓞 K)ˣ) : ∃! (ζ : torsion K) (e : Fin (rank K) → ℤ), + x = ζ * ∏ i, (fund_system K i) ^ (e i) := by + let ζ := x * (∏ i, (fund_system K i) ^ ((basis_mod_torsion K).repr (Additive.ofMul ↑x) i))⁻¹ + have h_tors : ζ ∈ torsion K := by + rw [← QuotientGroup.eq_one_iff, QuotientGroup.mk_mul, QuotientGroup.mk_inv, ← ofMul_eq_zero, + ofMul_mul, ofMul_inv, QuotientGroup.mk_prod, ofMul_prod] + simp_rw [QuotientGroup.mk_zpow, ofMul_zpow, fund_system, QuotientGroup.out_eq', ofMul_toMul] + rw [add_eq_zero_iff_eq_neg, neg_neg] + exact ((basis_mod_torsion K).sum_repr (Additive.ofMul ↑x)).symm + refine ⟨⟨ζ, h_tors⟩, ?_, ?_⟩ + · refine ⟨((basis_mod_torsion K).repr (Additive.ofMul ↑x) : Fin (rank K) → ℤ), ?_, ?_⟩ + · simp only [_root_.inv_mul_cancel_right] + · exact fun _ hf => fun_eq_repr K h_tors hf + · rintro η ⟨_, hf, _⟩ + have f_eq := fun_eq_repr K η.prop hf + simp_rw [f_eq] at hf + ext1; dsimp only + nth_rewrite 1 [hf] + rw [_root_.mul_inv_cancel_right] + end NumberField.Units From 01e244ba7f85d80b8b2fb50a293aa94157c27c0b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 8 Aug 2023 11:16:01 +0900 Subject: [PATCH 40/67] 1st commit --- .../LinearAlgebra/Matrix/ToLinearEquiv.lean | 41 ++++++++++++++++++- 1 file changed, 40 insertions(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean index 79018b59a513d..a4f27c17d4465 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean @@ -36,11 +36,15 @@ matrix, linear_equiv, determinant, inverse namespace Matrix +variable {n : Type _} [Fintype n] + +section LinearEquiv + open LinearMap variable {R M : Type _} [CommRing R] [AddCommGroup M] [Module R M] -variable {n : Type _} [Fintype n] + section ToLinearEquiv' @@ -196,4 +200,39 @@ alias nondegenerate_iff_det_ne_zero ↔ Nondegenerate.det_ne_zero Nondegenerate. end Nondegenerate +end LinearEquiv + +section Determinant + +open BigOperators + +lemma det_ne_zero_of_neg [DecidableEq n] {S : Type _} [LinearOrderedCommRing S] + {A : Matrix n n S} (h1 : ∀ i j, i ≠ j → A i j < 0) (h2 : ∀ j, 0 < ∑ i, A i j) : + A.det ≠ 0 := by + by_cases hn : Nonempty n + · contrapose! h2 + obtain ⟨v, ⟨h_vnz, h_vA⟩⟩ := Matrix.exists_vecMul_eq_zero_iff.mpr h2 + wlog h_sup : 0 < Finset.sup' Finset.univ Finset.univ_nonempty v + · refine this h1 hn h2 (-1 • v) ?_ ?_ ?_ + · exact smul_ne_zero (by norm_num) h_vnz + · rw [Matrix.vecMul_smul, h_vA, smul_zero] + · obtain ⟨i, hi⟩ := Function.ne_iff.mp h_vnz + simp_rw [Finset.lt_sup'_iff, Finset.mem_univ, true_and] at h_sup ⊢ + simp_rw [not_exists, not_lt] at h_sup + refine ⟨i, ?_⟩ + rw [Pi.smul_apply, neg_smul, one_smul, Left.neg_pos_iff] + refine Ne.lt_of_le hi (h_sup i) + · obtain ⟨j₀, -, h_j₀⟩ := Finset.exists_mem_eq_sup' Finset.univ_nonempty v + refine ⟨j₀, ?_⟩ + rw [← mul_le_mul_left (h_j₀ ▸ h_sup), Finset.mul_sum, mul_zero] + rw [show 0 = ∑ i, v i * A i j₀ from (congrFun h_vA j₀).symm] + refine Finset.sum_le_sum (fun i hi => ?_) + by_cases h : i = j₀ + · rw [h] + · exact (mul_le_mul_right_of_neg (h1 i j₀ h)).mpr (h_j₀ ▸ Finset.le_sup' v hi) + · have := (isEmpty_or_nonempty n).resolve_right hn + simp + +end Determinant + end Matrix From 6324f7148a4f999b9809af217ec3264ffa6334ed Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 8 Aug 2023 11:33:31 +0900 Subject: [PATCH 41/67] Add docs --- Mathlib/NumberTheory/NumberField/Units.lean | 16 +++++++++++++++- 1 file changed, 15 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index c7250abd2cf75..8dc8cc2857675 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -15,9 +15,22 @@ import Mathlib.RingTheory.RootsOfUnity.Basic We prove results about the group `(𝓞 K)ˣ` of units of the ring of integers `𝓞 K` of a number field `K`. +## Main definitions + +* `Units.rank`: the unit rank of the number field `K` + +* `Units.fund_system`: a fundamental system of units of `K` + +* `Units.basis_mod_torsion`: a `ℤ`-basis of `(𝓞 K)ˣ ⧸ (torsion K)` (as an additive `ℤ`-module) + ## Main results + * `isUnit_iff_norm`: an algebraic integer `x : 𝓞 K` is a unit if and only if `|norm ℚ x| = 1`. -* `mem_torsion`: a unit `x : (𝓞 K)ˣ` is torsion iff `w x = 1` for all infinite places of `K`. + +* `Units.mem_torsion`: an unit `x : (𝓞 K)ˣ` is torsion iff `w x = 1` for all infinite places of `K`. + +* `Units.exist_unique_eq_mul_prod`: an unit `x : (𝓞 K)ˣ` can be written in a unique way as the +product of a root of unity times the product of powers of units of the fundamenal system. ## Tags number field, units @@ -55,6 +68,7 @@ theorem isUnit_iff_norm [NumberField K] {x : 𝓞 K} : #align is_unit_iff_norm isUnit_iff_norm end IsUnit + namespace NumberField.Units section coe From 6f2f39be5b525873a2bcdcee140e85c410e36981 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 8 Aug 2023 12:07:36 +0900 Subject: [PATCH 42/67] Add QuotientGroup.mk_prod --- Mathlib/GroupTheory/QuotientGroup.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/GroupTheory/QuotientGroup.lean b/Mathlib/GroupTheory/QuotientGroup.lean index d4dd389ac5e67..97406d69c6a52 100644 --- a/Mathlib/GroupTheory/QuotientGroup.lean +++ b/Mathlib/GroupTheory/QuotientGroup.lean @@ -190,6 +190,11 @@ theorem mk_zpow (a : G) (n : ℤ) : ((a ^ n : G) : Q ) = (a : Q) ^ n := #align quotient_group.coe_zpow QuotientGroup.mk_zpow #align quotient_add_group.coe_zsmul QuotientAddGroup.mk_zsmul +@[to_additive (attr := simp)] +theorem mk_prod {G ι : Type _} [CommGroup G] (N : Subgroup G) [Fintype ι] {f : ι → G} : + ((Finset.univ.prod f : G) : G ⧸ N) = Finset.univ.prod (fun i => (f i : G ⧸ N)) := + map_prod (QuotientGroup.mk' N) _ _ + /-- A group homomorphism `φ : G →* H` with `N ⊆ ker(φ)` descends (i.e. `lift`s) to a group homomorphism `G/N →* H`. -/ @[to_additive "An `AddGroup` homomorphism `φ : G →+ H` with `N ⊆ ker(φ)` descends (i.e. `lift`s) From bcc1b6e51b153db5036b8191b6f4c20a13fb8330 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 8 Aug 2023 13:36:16 +0900 Subject: [PATCH 43/67] Fix statement --- Mathlib/GroupTheory/QuotientGroup.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/QuotientGroup.lean b/Mathlib/GroupTheory/QuotientGroup.lean index 97406d69c6a52..11965dcbe9ff9 100644 --- a/Mathlib/GroupTheory/QuotientGroup.lean +++ b/Mathlib/GroupTheory/QuotientGroup.lean @@ -191,8 +191,8 @@ theorem mk_zpow (a : G) (n : ℤ) : ((a ^ n : G) : Q ) = (a : Q) ^ n := #align quotient_add_group.coe_zsmul QuotientAddGroup.mk_zsmul @[to_additive (attr := simp)] -theorem mk_prod {G ι : Type _} [CommGroup G] (N : Subgroup G) [Fintype ι] {f : ι → G} : - ((Finset.univ.prod f : G) : G ⧸ N) = Finset.univ.prod (fun i => (f i : G ⧸ N)) := +theorem mk_prod {G ι : Type _} [CommGroup G] (N : Subgroup G) (s : Finset ι) {f : ι → G} : + ((Finset.prod s f : G) : G ⧸ N) = Finset.prod s (fun i => (f i : G ⧸ N)) := map_prod (QuotientGroup.mk' N) _ _ /-- A group homomorphism `φ : G →* H` with `N ⊆ ker(φ)` descends (i.e. `lift`s) to a From cdf26b24e3d7263e48dc1fd815eaa66b57392559 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 8 Aug 2023 13:36:41 +0900 Subject: [PATCH 44/67] Use DiscreteTopology --- Mathlib/NumberTheory/NumberField/Units.lean | 26 ++++++++++++++------- 1 file changed, 18 insertions(+), 8 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 8dc8cc2857675..2ceb2e20634b0 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -448,20 +448,30 @@ def _root_.NumberField.Units.rank : ℕ := Fintype.card (InfinitePlace K) - 1 open FiniteDimensional +theorem unit_lattice_discrete : DiscreteTopology (unit_lattice K) := by + refine discreteTopology_of_open_singleton_zero ?_ + refine isOpen_singleton_of_finite_mem_nhds 0 (s := Metric.closedBall 0 1) ?_ ?_ + · exact Metric.closedBall_mem_nhds _ (by norm_num) + · refine Set.Finite.of_finite_image ?_ (Set.injOn_of_injective Subtype.val_injective _) + convert unit_lattice_inter_ball_finite K 1 + ext x + refine ⟨?_, fun ⟨hx1, hx2⟩ => ⟨⟨x, hx1⟩, hx2, rfl⟩⟩ + rintro ⟨x, hx, rfl⟩ + exact ⟨Subtype.mem x, hx⟩ + theorem rank_space : - finrank ℝ ({w : InfinitePlace K // w ≠ w₀} → ℝ) = rank K := by + finrank ℝ ({w : InfinitePlace K // w ≠ w₀} → ℝ) = Units.rank K := by simp only [finrank_fintype_fun_eq_card, Fintype.card_subtype_compl, Fintype.card_ofSubsingleton, rank] -theorem unit_lattice_moduleFree : Module.Free ℤ (unit_lattice K) := -Zlattice.module_free ℝ ((unit_lattice_inter_ball_finite K)) (unit_lattice_span_eq_top K) - -theorem unit_lattice_moduleFinite : Module.Finite ℤ (unit_lattice K) := -Zlattice.module_finite ℝ ((unit_lattice_inter_ball_finite K)) (unit_lattice_span_eq_top K) +theorem unit_lattice_moduleFree : Module.Free ℤ (unit_lattice K) := by + have := unit_lattice_discrete K + exact Zlattice.module_free ℝ (unit_lattice_span_eq_top K) -theorem unit_lattice_rank : finrank ℤ (unit_lattice K) = rank K := by +theorem unit_lattice.rank : finrank ℤ (unit_lattice K) = Units.rank K := by + have := unit_lattice_discrete K rw [← rank_space] - exact Zlattice.rank ℝ ((unit_lattice_inter_ball_finite K)) (unit_lattice_span_eq_top K) + exact Zlattice.rank ℝ (unit_lattice_span_eq_top K) end dirichlet From 4114da1c78c89528709b5a7329d44d888ceaa69a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 8 Aug 2023 13:53:19 +0900 Subject: [PATCH 45/67] Fix docs --- Mathlib/NumberTheory/NumberField/Units.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 2ceb2e20634b0..424b36f94c294 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -17,11 +17,11 @@ field `K`. ## Main definitions -* `Units.rank`: the unit rank of the number field `K` +* `Units.rank`: the unit rank of the number field `K`. -* `Units.fund_system`: a fundamental system of units of `K` +* `Units.fund_system`: a fundamental system of units of `K`. -* `Units.basis_mod_torsion`: a `ℤ`-basis of `(𝓞 K)ˣ ⧸ (torsion K)` (as an additive `ℤ`-module) +* `Units.basis_mod_torsion`: a `ℤ`-basis of `(𝓞 K)ˣ ⧸ (torsion K)` (as an additive `ℤ`-module). ## Main results @@ -30,7 +30,7 @@ field `K`. * `Units.mem_torsion`: an unit `x : (𝓞 K)ˣ` is torsion iff `w x = 1` for all infinite places of `K`. * `Units.exist_unique_eq_mul_prod`: an unit `x : (𝓞 K)ˣ` can be written in a unique way as the -product of a root of unity times the product of powers of units of the fundamenal system. +product of a root of unity times the product of powers of units of the fundamental system. ## Tags number field, units From dfd08d12cdd3cfc6c0d4860ffe6f373a1b0a3b6a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Tue, 8 Aug 2023 22:49:21 +0900 Subject: [PATCH 46/67] More fixes --- Mathlib/NumberTheory/NumberField/Units.lean | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 424b36f94c294..659f5508c5bda 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -468,7 +468,11 @@ theorem unit_lattice_moduleFree : Module.Free ℤ (unit_lattice K) := by have := unit_lattice_discrete K exact Zlattice.module_free ℝ (unit_lattice_span_eq_top K) -theorem unit_lattice.rank : finrank ℤ (unit_lattice K) = Units.rank K := by +theorem unit_lattice_moduleFinite : Module.Finite ℤ (unit_lattice K) := by + have := unit_lattice_discrete K + exact Zlattice.module_finite ℝ (unit_lattice_span_eq_top K) + +theorem unit_lattice_rank : finrank ℤ (unit_lattice K) = Units.rank K := by have := unit_lattice_discrete K rw [← rank_space] exact Zlattice.rank ℝ (unit_lattice_span_eq_top K) From 44a095d4734d18740d3c088c6cab2cb8b8fffa25 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 9 Aug 2023 17:48:35 +0900 Subject: [PATCH 47/67] Add the proof of Gershgorin --- Mathlib.lean | 1 + Mathlib/LinearAlgebra/Matrix/Gershgorin.lean | 71 +++++++++++++++++++ .../LinearAlgebra/Matrix/ToLinearEquiv.lean | 41 +---------- 3 files changed, 73 insertions(+), 40 deletions(-) create mode 100644 Mathlib/LinearAlgebra/Matrix/Gershgorin.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3352d57fad5e2..8cebbd36677dc 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2203,6 +2203,7 @@ import Mathlib.LinearAlgebra.Matrix.DotProduct import Mathlib.LinearAlgebra.Matrix.Dual import Mathlib.LinearAlgebra.Matrix.FiniteDimensional import Mathlib.LinearAlgebra.Matrix.GeneralLinearGroup +import Mathlib.LinearAlgebra.Matrix.Gershgorin import Mathlib.LinearAlgebra.Matrix.Hermitian import Mathlib.LinearAlgebra.Matrix.InvariantBasisNumber import Mathlib.LinearAlgebra.Matrix.IsDiag diff --git a/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean new file mode 100644 index 0000000000000..23b8a81d69896 --- /dev/null +++ b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean @@ -0,0 +1,71 @@ +/- +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.Normed.Field.Basic +import Mathlib.LinearAlgebra.Eigenspace.Basic +import Mathlib.LinearAlgebra.Determinant + +/-! +# Gershgorin's circle theorem + +This file gives the proof of Gershgorin's circle theorem `Matrix.toLinearEquiv` on the eigenvalue +of matrices and some applications + +## Reference + +* https://en.wikipedia.org/wiki/Gershgorin_circle_theorem +-/ + +open BigOperators + +variable {K n : Type _} [NormedField K] [Fintype n] [DecidableEq n] {A : Matrix n n K} + +theorem eigenvalue_mem_ball {μ : K} (hμ : Module.End.HasEigenvalue (Matrix.toLin' A) μ) : + ∃ k, μ ∈ Metric.closedBall (A k k) (∑ j in Finset.univ.erase k, ‖A k j‖) := by + by_cases hn : Nonempty n + · obtain ⟨v, h_eg, h_nz⟩ := hμ.exists_hasEigenvector + obtain ⟨i, -, h_i⟩ := Finset.exists_mem_eq_sup' Finset.univ_nonempty (fun i => ‖v i‖) + have h_nz : v i ≠ 0 := by + contrapose! h_nz + ext j + rw [Pi.zero_apply, ← norm_le_zero_iff] + refine (h_i ▸ Finset.le_sup' (fun i => ‖v i‖) (Finset.mem_univ j)).trans ?_ + exact norm_le_zero_iff.mpr h_nz + have h_le : ∀ j, ‖v j * (v i)⁻¹‖ ≤ 1 := fun j => by + rw [norm_mul, norm_inv, mul_inv_le_iff' (norm_pos_iff.mpr h_nz), one_mul] + exact h_i ▸ Finset.le_sup' (fun i => ‖v i‖) (Finset.mem_univ j) + simp_rw [mem_closedBall_iff_norm'] + refine ⟨i, ?_⟩ + calc + _ = ‖(A i i * v i - μ * v i) * (v i)⁻¹‖ := by congr; field_simp [h_nz]; ring + _ = ‖(A i i * v i - ∑ j, A i j * v j) * (v i)⁻¹‖ := by + rw [show μ * v i = ∑ x : n, A i x * v x by + rw [← Matrix.dotProduct, ← Matrix.mulVec] + exact (congrFun (Module.End.mem_eigenspace_iff.mp h_eg) i).symm] + _ = ‖(∑ j in Finset.univ.erase i, A i j * v j) * (v i)⁻¹‖ := by + rw [Finset.sum_erase_eq_sub (Finset.mem_univ i), ← neg_sub, neg_mul, norm_neg] + _ ≤ ∑ j in Finset.univ.erase i, ‖A i j‖ * ‖v j * (v i)⁻¹‖ := by + rw [Finset.sum_mul] + exact (norm_sum_le _ _).trans (le_of_eq (by simp_rw [mul_assoc, norm_mul])) + _ ≤ ∑ j in Finset.univ.erase i, ‖A i j‖ := + (Finset.sum_le_sum fun j _ => mul_le_of_le_one_right (norm_nonneg _) (h_le j)) + · rw [not_nonempty_iff] at hn + exfalso + exact hμ (Submodule.eq_bot_of_subsingleton _) + +theorem det_ne_zero_of_sum_row_lt_diag (h : ∀ k, ∑ j in Finset.univ.erase k, ‖A k j‖ < ‖A k k‖) : + A.det ≠ 0 := by + contrapose! h + suffices ∃ k, 0 ∈ Metric.closedBall (A k k) (∑ j in Finset.univ.erase k, ‖A k j‖) by + simp_rw [mem_closedBall_iff_norm', sub_zero] at this + exact this + refine eigenvalue_mem_ball ?_ + rw [Module.End.HasEigenvalue, Module.End.eigenspace_zero, ne_comm] + exact ne_of_lt (LinearMap.bot_lt_ker_of_det_eq_zero (by rwa [LinearMap.det_toLin'])) + +theorem det_ne_zero_of_sum_col_lt_diag (h : ∀ k, ∑ i in Finset.univ.erase k, ‖A i k‖ < ‖A k k‖) : + A.det ≠ 0 := by + rw [← Matrix.det_transpose] + exact det_ne_zero_of_sum_row_lt_diag (by simp_rw [Matrix.transpose_apply]; exact h) diff --git a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean index a4f27c17d4465..79018b59a513d 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean @@ -36,15 +36,11 @@ matrix, linear_equiv, determinant, inverse namespace Matrix -variable {n : Type _} [Fintype n] - -section LinearEquiv - open LinearMap variable {R M : Type _} [CommRing R] [AddCommGroup M] [Module R M] - +variable {n : Type _} [Fintype n] section ToLinearEquiv' @@ -200,39 +196,4 @@ alias nondegenerate_iff_det_ne_zero ↔ Nondegenerate.det_ne_zero Nondegenerate. end Nondegenerate -end LinearEquiv - -section Determinant - -open BigOperators - -lemma det_ne_zero_of_neg [DecidableEq n] {S : Type _} [LinearOrderedCommRing S] - {A : Matrix n n S} (h1 : ∀ i j, i ≠ j → A i j < 0) (h2 : ∀ j, 0 < ∑ i, A i j) : - A.det ≠ 0 := by - by_cases hn : Nonempty n - · contrapose! h2 - obtain ⟨v, ⟨h_vnz, h_vA⟩⟩ := Matrix.exists_vecMul_eq_zero_iff.mpr h2 - wlog h_sup : 0 < Finset.sup' Finset.univ Finset.univ_nonempty v - · refine this h1 hn h2 (-1 • v) ?_ ?_ ?_ - · exact smul_ne_zero (by norm_num) h_vnz - · rw [Matrix.vecMul_smul, h_vA, smul_zero] - · obtain ⟨i, hi⟩ := Function.ne_iff.mp h_vnz - simp_rw [Finset.lt_sup'_iff, Finset.mem_univ, true_and] at h_sup ⊢ - simp_rw [not_exists, not_lt] at h_sup - refine ⟨i, ?_⟩ - rw [Pi.smul_apply, neg_smul, one_smul, Left.neg_pos_iff] - refine Ne.lt_of_le hi (h_sup i) - · obtain ⟨j₀, -, h_j₀⟩ := Finset.exists_mem_eq_sup' Finset.univ_nonempty v - refine ⟨j₀, ?_⟩ - rw [← mul_le_mul_left (h_j₀ ▸ h_sup), Finset.mul_sum, mul_zero] - rw [show 0 = ∑ i, v i * A i j₀ from (congrFun h_vA j₀).symm] - refine Finset.sum_le_sum (fun i hi => ?_) - by_cases h : i = j₀ - · rw [h] - · exact (mul_le_mul_right_of_neg (h1 i j₀ h)).mpr (h_j₀ ▸ Finset.le_sup' v hi) - · have := (isEmpty_or_nonempty n).resolve_right hn - simp - -end Determinant - end Matrix From b20a2601bf5c578e818f4c8e6a441d23ee2f78b0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Thu, 10 Aug 2023 07:10:09 +0900 Subject: [PATCH 48/67] review --- Mathlib/LinearAlgebra/Matrix/Gershgorin.lean | 7 +++---- 1 file changed, 3 insertions(+), 4 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean index 23b8a81d69896..19c6de4e4faff 100644 --- a/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean +++ b/Mathlib/LinearAlgebra/Matrix/Gershgorin.lean @@ -24,7 +24,9 @@ variable {K n : Type _} [NormedField K] [Fintype n] [DecidableEq n] {A : Matrix theorem eigenvalue_mem_ball {μ : K} (hμ : Module.End.HasEigenvalue (Matrix.toLin' A) μ) : ∃ k, μ ∈ Metric.closedBall (A k k) (∑ j in Finset.univ.erase k, ‖A k j‖) := by - by_cases hn : Nonempty n + cases isEmpty_or_nonempty n + · exfalso + exact hμ (Submodule.eq_bot_of_subsingleton _) · obtain ⟨v, h_eg, h_nz⟩ := hμ.exists_hasEigenvector obtain ⟨i, -, h_i⟩ := Finset.exists_mem_eq_sup' Finset.univ_nonempty (fun i => ‖v i‖) have h_nz : v i ≠ 0 := by @@ -51,9 +53,6 @@ theorem eigenvalue_mem_ball {μ : K} (hμ : Module.End.HasEigenvalue (Matrix.toL exact (norm_sum_le _ _).trans (le_of_eq (by simp_rw [mul_assoc, norm_mul])) _ ≤ ∑ j in Finset.univ.erase i, ‖A i j‖ := (Finset.sum_le_sum fun j _ => mul_le_of_le_one_right (norm_nonneg _) (h_le j)) - · rw [not_nonempty_iff] at hn - exfalso - exact hμ (Submodule.eq_bot_of_subsingleton _) theorem det_ne_zero_of_sum_row_lt_diag (h : ∀ k, ∑ j in Finset.univ.erase k, ‖A k j‖ < ‖A k k‖) : A.det ≠ 0 := by From 2169056a40a4bae2dc09969420dae983ee779518 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Thu, 10 Aug 2023 09:10:14 +0900 Subject: [PATCH 49/67] Merge --- .../LinearAlgebra/Matrix/ToLinearEquiv.lean | 41 +------------------ Mathlib/NumberTheory/NumberField/Units.lean | 26 +++++++----- 2 files changed, 16 insertions(+), 51 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean index a4f27c17d4465..79018b59a513d 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean @@ -36,15 +36,11 @@ matrix, linear_equiv, determinant, inverse namespace Matrix -variable {n : Type _} [Fintype n] - -section LinearEquiv - open LinearMap variable {R M : Type _} [CommRing R] [AddCommGroup M] [Module R M] - +variable {n : Type _} [Fintype n] section ToLinearEquiv' @@ -200,39 +196,4 @@ alias nondegenerate_iff_det_ne_zero ↔ Nondegenerate.det_ne_zero Nondegenerate. end Nondegenerate -end LinearEquiv - -section Determinant - -open BigOperators - -lemma det_ne_zero_of_neg [DecidableEq n] {S : Type _} [LinearOrderedCommRing S] - {A : Matrix n n S} (h1 : ∀ i j, i ≠ j → A i j < 0) (h2 : ∀ j, 0 < ∑ i, A i j) : - A.det ≠ 0 := by - by_cases hn : Nonempty n - · contrapose! h2 - obtain ⟨v, ⟨h_vnz, h_vA⟩⟩ := Matrix.exists_vecMul_eq_zero_iff.mpr h2 - wlog h_sup : 0 < Finset.sup' Finset.univ Finset.univ_nonempty v - · refine this h1 hn h2 (-1 • v) ?_ ?_ ?_ - · exact smul_ne_zero (by norm_num) h_vnz - · rw [Matrix.vecMul_smul, h_vA, smul_zero] - · obtain ⟨i, hi⟩ := Function.ne_iff.mp h_vnz - simp_rw [Finset.lt_sup'_iff, Finset.mem_univ, true_and] at h_sup ⊢ - simp_rw [not_exists, not_lt] at h_sup - refine ⟨i, ?_⟩ - rw [Pi.smul_apply, neg_smul, one_smul, Left.neg_pos_iff] - refine Ne.lt_of_le hi (h_sup i) - · obtain ⟨j₀, -, h_j₀⟩ := Finset.exists_mem_eq_sup' Finset.univ_nonempty v - refine ⟨j₀, ?_⟩ - rw [← mul_le_mul_left (h_j₀ ▸ h_sup), Finset.mul_sum, mul_zero] - rw [show 0 = ∑ i, v i * A i j₀ from (congrFun h_vA j₀).symm] - refine Finset.sum_le_sum (fun i hi => ?_) - by_cases h : i = j₀ - · rw [h] - · exact (mul_le_mul_right_of_neg (h1 i j₀ h)).mpr (h_j₀ ▸ Finset.le_sup' v hi) - · have := (isEmpty_or_nonempty n).resolve_right hn - simp - -end Determinant - end Matrix diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 659f5508c5bda..cdad28a777cd2 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.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.LinearAlgebra.Matrix.Gershgorin import Mathlib.NumberTheory.NumberField.CanonicalEmbedding import Mathlib.NumberTheory.NumberField.Norm import Mathlib.RingTheory.Ideal.Norm @@ -420,8 +421,8 @@ theorem unit_lattice_span_eq_top : refine le_antisymm (le_top) ?_ -- The standard basis let B := Pi.basisFun ℝ {w : InfinitePlace K // w ≠ w₀} - -- The family of units constructed above - let v := fun w : { w : InfinitePlace K // w ≠ w₀ } => log_embedding K ((exists_unit K w).choose) + -- The image by log_embedding of the family of units constructed above + let v := fun w : { w : InfinitePlace K // w ≠ w₀ } => log_embedding K (exists_unit K w).choose -- To prove the result, it is enough to prove that the family `v` is linearly independent suffices B.det v ≠ 0 by rw [← isUnit_iff_ne_zero, ← is_basis_iff_det] at this @@ -430,16 +431,19 @@ theorem unit_lattice_span_eq_top : ⟨(exists_unit K w).choose, trivial, by rw [← hw]⟩) rw [Basis.det_apply] -- We use a specific lemma to prove that this determinant is nonzero - refine Matrix.det_ne_zero_of_neg (fun i j hij => ?_) (fun j => ?_) - · rw [Basis.coePiBasisFun.toMatrix_eq_transpose, Matrix.transpose_apply] - refine mul_neg_of_pos_of_neg ?_ ((exists_unit K j).choose_spec i ?_) - · rw [mult]; split_ifs <;> norm_num - · exact Subtype.ext_iff_val.not.mp hij - · simp_rw [Basis.coePiBasisFun.toMatrix_eq_transpose, Matrix.transpose_apply, - log_embedding_sum_component] - refine mul_pos_of_neg_of_neg ?_ ?_ + refine det_ne_zero_of_sum_col_lt_diag (fun w => ?_) + simp_rw [Real.norm_eq_abs, Basis.coePiBasisFun.toMatrix_eq_transpose, Matrix.transpose_apply] + rw [← sub_pos, Finset.sum_congr rfl (fun x hx => abs_of_neg ?_), Finset.sum_neg_distrib, + sub_neg_eq_add, Finset.sum_erase_eq_sub (Finset.mem_univ _), ← add_comm_sub] + refine add_pos_of_nonneg_of_pos ?_ ?_ + · rw [sub_nonneg] + exact le_abs_self _ + · rw [log_embedding_sum_component (exists_unit K w).choose] + refine mul_pos_of_neg_of_neg ?_ ((exists_unit K w).choose_spec _ w.prop.symm) + rw [mult]; split_ifs <;> norm_num + · refine mul_neg_of_pos_of_neg ?_ ((exists_unit K w).choose_spec x ?_) · rw [mult]; split_ifs <;> norm_num - · exact (exists_unit K j).choose_spec _ j.prop.symm + · exact Subtype.ext_iff_val.not.mp (Finset.ne_of_mem_erase hx) end span_top From bf2b10e67bec2dd5cc97149399ca40ca5270ab01 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sun, 13 Aug 2023 09:16:29 +0900 Subject: [PATCH 50/67] Docstrings --- Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 9b99ec576c36c..a7400434384e8 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -13,11 +13,15 @@ import Mathlib.RingTheory.Discriminant The canonical embedding of a number field `K` of degree `n` is the ring homomorphism `K →+* ℂ^n` that sends `x ∈ K` to `(φ_₁(x),...,φ_n(x))` where the `φ_i`'s are the complex -embeddings of `K`. +embeddings of `K`. Note that we do not choose an ordering of the embeddings, but instead map `K` +into the type `(K →+* ℂ) → ℂ` of `ℂ`-vectors indexed by the complex embeddings. ## Main definitions and results -* `NumberField.canonicalEmbedding.integerLattice.inter_ball_finite`: the intersection of the +* `canonicalEmbedding`: the ring homorphism `K →+* ((K →+* ℂ) → ℂ)` defined by sending `x : K` to +the vector `(φ x)` indexed by `φ : K →+* ℂ`. + +* `canonicalEmbedding.integerLattice.inter_ball_finite`: the intersection of the image of the ring of integers by the canonical embedding and any ball centered at `0` of finite radius is finite. From 30681d117077bea3263789390c143734494b8e19 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sun, 13 Aug 2023 09:27:39 +0900 Subject: [PATCH 51/67] Docstring --- Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 766acc6c0a32f..16f57e971e88c 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -28,10 +28,11 @@ the vector `(φ x)` indexed by `φ : K →+* ℂ`. image of the ring of integers by the canonical embedding and any ball centered at `0` of finite radius is finite. -* `NumberField.mixedEmbedding`: the ring homomorphism `K →+* ℝ^r₁ × ℂ^r₂` that sends `x ∈ K` to -`(φ_₁(x),...,φ_r₁(x)) × (ψ_₁(x),..., ψ_r₂(x)) ` where `(r₁, r₂)` is the signature of `K`, -`φ_₁,...,φ_r₁` are its real embeddings and `ψ_₁,..., ψ_r₂` are its complex embeddings (up to -complex conjugation). +* `mixedEmbedding`: the ring homomorphism from `K →+* ({ w // IsReal w } → ℝ) × +({ w // IsComplex w } → ℂ)` that sends `x ∈ K` to `(φ_w x)_w` where `φ_w` is the embedding +associated to the infinite place `w`. In particular, if `w` is real then `φ_w : K →+* ℝ` and, if +`w` is complex, `φ_w` is an arbitrary choice between the two complex emebeddings defining the place +`w`. * `exists_ne_zero_mem_ringOfIntegers_lt`: let `f : InfinitePlace K → ℝ≥0`, if the product `∏ w, f w` is large enough, proves that there exists a nonzero algebraic integer `a` such From 62c598558525a7f2bfa318d1694b3946a6865eb4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sun, 20 Aug 2023 15:44:45 +0200 Subject: [PATCH 52/67] fix names --- .../NumberField/CanonicalEmbedding.lean | 42 +++++++++---------- 1 file changed, 21 insertions(+), 21 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 7c1f1dc75f53c..19a7ae182428b 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -289,25 +289,25 @@ variable (f : InfinitePlace K → ℝ≥0) /-- The convex body defined by `f`: the set of points `x : E` such that `‖x w‖ < f w` for all infinite places `w`. -/ -def convex_body : Set (E K) := +abbrev convex_body_lt : Set (E K) := (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsReal w } => ball 0 (f w))) ×ˢ (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsComplex w } => ball 0 (f w))) -theorem convex_body_mem {x : K} : - mixedEmbedding K x ∈ (convex_body K f) ↔ ∀ w : InfinitePlace K, w x < f w := by - simp_rw [mixedEmbedding, RingHom.prod_apply, convex_body, Set.mem_prod, Set.mem_pi, - Set.mem_univ, forall_true_left, mem_ball_zero_iff, Pi.ringHom_apply, ← Complex.norm_real, +theorem convex_body_lt_mem {x : K} : + mixedEmbedding K x ∈ (convex_body_lt K f) ↔ ∀ w : InfinitePlace K, w x < f w := by + simp_rw [mixedEmbedding, RingHom.prod_apply, Set.mem_prod, Set.mem_pi, Set.mem_univ, + forall_true_left, mem_ball_zero_iff, Pi.ringHom_apply, ← Complex.norm_real, embedding_of_isReal_apply, Subtype.forall, ← ball_or_left, ← not_isReal_iff_isComplex, em, forall_true_left, norm_embedding_eq] -theorem convex_body_symmetric (x : E K) (hx : x ∈ (convex_body K f)) : - -x ∈ (convex_body K f) := by - simp only [convex_body, Set.mem_prod, Prod.fst_neg, Set.mem_pi, Set.mem_univ, Pi.neg_apply, +theorem convex_body_lt_symmetric (x : E K) (hx : x ∈ (convex_body_lt K f)) : + -x ∈ (convex_body_lt K f) := by + simp only [Set.mem_prod, Prod.fst_neg, Set.mem_pi, Set.mem_univ, Pi.neg_apply, mem_ball_zero_iff, norm_neg, Real.norm_eq_abs, forall_true_left, Subtype.forall, Prod.snd_neg, Complex.norm_eq_abs, hx] at hx ⊢ exact hx -theorem convex_body_convex : Convex ℝ (convex_body K f) := +theorem convex_body_lt_convex : Convex ℝ (convex_body_lt K f) := Convex.prod (convex_pi (fun _ _ => convex_ball _ _)) (convex_pi (fun _ _ => convex_ball _ _)) open Classical Fintype MeasureTheory MeasureTheory.Measure BigOperators @@ -317,7 +317,7 @@ local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) variable [NumberField K] -/-- The fudge factor that appears in the formula for the volume of `convex_body`. -/ +/-- The fudge factor that appears in the formula for the volume of `convex_body_lt`. -/ noncomputable def constant_factor : ℝ≥0∞ := (2 : ℝ≥0∞) ^ card {w : InfinitePlace K // IsReal w} * volume (ball (0 : ℂ) 1) ^ card {w : InfinitePlace K // IsComplex w} @@ -334,11 +334,11 @@ theorem constant_factor_lt_top : (constant_factor K) < ⊤ := by · exact ne_of_lt (pow_lt_top measure_ball_lt_top _) set_option maxHeartbeats 400000 in -/-- The volume of `(convex_body K f)` where `convex_body K f` is the set of points `x` such that -`‖x w‖ < f w` for all infinite places `w`. -/ +/-- The volume of `(convex_body_lt K f)` where `convex_body_lt K f` is the set of points `x` +such that `‖x w‖ < f w` for all infinite places `w`. -/ theorem convex_body_volume : - volume (convex_body K f) = (constant_factor K) * ∏ w, (f w) ^ (mult w) := by - rw [volume_eq_prod, convex_body, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi] + volume (convex_body_lt K f) = (constant_factor K) * ∏ w, (f w) ^ (mult w) := by + rw [volume_eq_prod, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi] conv_lhs => congr; congr; next => skip ext @@ -402,11 +402,11 @@ theorem minkowski_bound_lt_top : minkowski_bound K < ⊤ := by variable {f : InfinitePlace K → ℝ≥0} /-- Assume that `f : InfinitePlace K → ℝ≥0` is such that -`minkowski_bound K < volume (convex_body K f)` where `convex_body K f` is the set of points `x` -such that `‖x w‖ < f w` for all infinite places `w` (see `convex_body_volume` for the computation -of this volume), then there exists a nonzero algebraic integer `a` in `𝓞 K` such that -`w a < f w` for all infinite places `w`. -/ -theorem exists_ne_zero_mem_ringOfIntegers_lt (h : minkowski_bound K < volume (convex_body K f)) : +`minkowski_bound K < volume (convex_body_lt K f)` where `convex_body_lt K f` is the set of +points `x` such that `‖x w‖ < f w` for all infinite places `w` (see `convex_body_lt_volume` for +the computation of this volume), then there exists a nonzero algebraic integer `a` in `𝓞 K` such +that `w a < f w` for all infinite places `w`. -/ +theorem exists_ne_zero_mem_ringOfIntegers_lt (h : minkowski_bound K < volume (convex_body_lt K f)) : ∃ (a : 𝓞 K), a ≠ 0 ∧ ∀ w : InfinitePlace K, w a < f w := by have : @IsAddHaarMeasure (E K) _ _ _ volume := prod.instIsAddHaarMeasure volume volume have h_fund := Zspan.isAddFundamentalDomain (latticeBasis K) volume @@ -414,10 +414,10 @@ theorem exists_ne_zero_mem_ringOfIntegers_lt (h : minkowski_bound K < volume (co change Countable (Submodule.span ℤ (Set.range (latticeBasis K)): Set (E K)) infer_instance obtain ⟨⟨x, hx⟩, h_nzr, h_mem⟩ := exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure - h_fund h (convex_body_symmetric K f) (convex_body_convex K f) + h_fund h (convex_body_lt_symmetric K f) (convex_body_lt_convex K f) rw [Submodule.mem_toAddSubgroup, mem_span_latticeBasis] at hx obtain ⟨a, ha, rfl⟩ := hx - refine ⟨⟨a, ha⟩, ?_, (convex_body_mem K f).mp h_mem⟩ + refine ⟨⟨a, ha⟩, ?_, (convex_body_lt_mem K f).mp h_mem⟩ rw [ne_eq, AddSubgroup.mk_eq_zero_iff, map_eq_zero, ← ne_eq] at h_nzr exact Subtype.ne_of_val_ne h_nzr From f0076cfaa69676a37a3d15a151d0033c9038efb5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 25 Aug 2023 16:18:24 +0200 Subject: [PATCH 53/67] typo --- Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 19a7ae182428b..762f8f159f1c6 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -167,7 +167,7 @@ open NumberField NumberField.InfinitePlace NumberField.ComplexEmbedding local notation "E" K => ({ w : InfinitePlace K // IsReal w } → ℝ) × ({ w : InfinitePlace K // IsComplex w } → ℂ) -/-- The canonical embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ +/-- The mixed embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ noncomputable def _root_.NumberField.mixedEmbedding : K →+* (E K) := RingHom.prod (Pi.ringHom fun w => embedding_of_isReal w.prop) (Pi.ringHom fun w => w.val.embedding) From ed5dff296cd390d24bc1fea496f8aa221c52e761 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Mon, 28 Aug 2023 18:40:09 +0200 Subject: [PATCH 54/67] 1st commit --- Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean index 6fe6f3d6fe5cb..79d27ce3d07ba 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Complex.lean @@ -3,9 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.MeasureTheory.Constructions.BorelSpace.Complex -import Mathlib.MeasureTheory.Measure.Lebesgue.Basic -import Mathlib.MeasureTheory.Measure.Haar.OfBasis +import Mathlib.MeasureTheory.Measure.Haar.NormedSpace #align_import measure_theory.measure.lebesgue.complex from "leanprover-community/mathlib"@"fd5edc43dc4f10b85abfe544b88f82cf13c5f844" @@ -49,4 +47,7 @@ theorem volume_preserving_equiv_real_prod : MeasurePreserving measurableEquivRea (volume_preserving_finTwoArrow ℝ).comp volume_preserving_equiv_pi #align complex.volume_preserving_equiv_real_prod Complex.volume_preserving_equiv_real_prod +instance : Measure.IsAddHaarMeasure (@volume ℂ Complex.measureSpace) := + Measure.MapLinearEquiv.isAddHaarMeasure _ _ + end Complex From 5f48003721027d0c7300126a122cfb8bb007fb7e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Mon, 28 Aug 2023 18:52:15 +0200 Subject: [PATCH 55/67] Cleanup --- .../NumberField/CanonicalEmbedding.lean | 27 +++++++++++-------- 1 file changed, 16 insertions(+), 11 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 762f8f159f1c6..236602083a779 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -161,11 +161,11 @@ end NumberField.canonicalEmbedding namespace NumberField.mixedEmbedding -open NumberField NumberField.InfinitePlace NumberField.ComplexEmbedding +open NumberField NumberField.InfinitePlace NumberField.ComplexEmbedding FiniteDimensional /-- The ambient space `ℝ^r₁ × ℂ^r₂` with `(r₁, r₂)` the signature of `K`. -/ local notation "E" K => - ({ w : InfinitePlace K // IsReal w } → ℝ) × ({ w : InfinitePlace K // IsComplex w } → ℂ) + ({w : InfinitePlace K // IsReal w} → ℝ) × ({w : InfinitePlace K // IsComplex w} → ℂ) /-- The mixed embedding of a number field `K` of signature `(r₁, r₂)` into `ℝ^r₁ × ℂ^r₂`. -/ noncomputable def _root_.NumberField.mixedEmbedding : K →+* (E K) := @@ -180,6 +180,13 @@ instance [NumberField K] : Nontrivial (E K) := by · have : Nonempty {w : InfinitePlace K // IsComplex w} := ⟨⟨w, hw⟩⟩ exact nontrivial_prod_right +theorem rank_space [NumberField K] : finrank ℝ (E K) = finrank ℚ K := by + classical + rw [finrank_prod, finrank_pi, finrank_pi_fintype, Complex.finrank_real_complex, Finset.sum_const, + Finset.card_univ, ← card_real_embeddings, Algebra.id.smul_eq_mul, mul_comm, + ← card_complex_embeddings, ← NumberField.Embeddings.card K ℂ, Fintype.card_subtype_compl, + Nat.add_sub_of_le (Fintype.card_subtype_le _)] + theorem _root_.NumberField.mixedEmbedding_injective [NumberField K] : Function.Injective (NumberField.mixedEmbedding K) := by exact RingHom.injective _ @@ -281,7 +288,7 @@ theorem mem_span_latticeBasis [NumberField K] (x : (E K)) : end integerLattice -section convex_body +section convex_body_lt open Metric ENNReal NNReal @@ -290,8 +297,8 @@ variable (f : InfinitePlace K → ℝ≥0) /-- The convex body defined by `f`: the set of points `x : E` such that `‖x w‖ < f w` for all infinite places `w`. -/ abbrev convex_body_lt : Set (E K) := - (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsReal w } => ball 0 (f w))) ×ˢ - (Set.pi Set.univ (fun w : { w : InfinitePlace K // IsComplex w } => ball 0 (f w))) + (Set.univ.pi (fun w : { w : InfinitePlace K // IsReal w } => ball 0 (f w))) ×ˢ + (Set.univ.pi (fun w : { w : InfinitePlace K // IsComplex w } => ball 0 (f w))) theorem convex_body_lt_mem {x : K} : mixedEmbedding K x ∈ (convex_body_lt K f) ↔ ∀ w : InfinitePlace K, w x < f w := by @@ -322,8 +329,6 @@ noncomputable def constant_factor : ℝ≥0∞ := (2 : ℝ≥0∞) ^ card {w : InfinitePlace K // IsReal w} * volume (ball (0 : ℂ) 1) ^ card {w : InfinitePlace K // IsComplex w} -instance : IsAddHaarMeasure (@volume ℂ Complex.measureSpace) := MapLinearEquiv.isAddHaarMeasure _ _ - theorem constant_factor_pos : 0 < (constant_factor K) := by refine mul_pos (NeZero.ne _) ?_ exact ENNReal.pow_ne_zero (ne_of_gt (measure_ball_pos _ _ (by norm_num))) _ @@ -381,22 +386,22 @@ theorem adjust_f {w₁ : InfinitePlace K} (B : ℝ≥0) (hf : ∀ w, w ≠ w₁ exact fun w hw => pow_ne_zero _ (hf w (Finset.ne_of_mem_erase hw)) · rw [mult]; split_ifs <;> norm_num -end convex_body +end convex_body_lt section minkowski -open MeasureTheory MeasureTheory.Measure Classical NNReal ENNReal FiniteDimensional +open MeasureTheory MeasureTheory.Measure Classical NNReal ENNReal FiniteDimensional Zspan variable [NumberField K] /-- The bound that appears in Minkowski theorem, see `MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure`. -/ noncomputable def minkowski_bound : ℝ≥0∞ := - volume (Zspan.fundamentalDomain (latticeBasis K)) * 2 ^ (finrank ℝ (E K)) + volume (fundamentalDomain (latticeBasis K)) * 2 ^ (finrank ℝ (E K)) theorem minkowski_bound_lt_top : minkowski_bound K < ⊤ := by refine mul_lt_top ?_ ?_ - · exact ne_of_lt (Zspan.fundamentalDomain_bounded (latticeBasis K)).measure_lt_top + · exact ne_of_lt (fundamentalDomain_bounded (latticeBasis K)).measure_lt_top · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) variable {f : InfinitePlace K → ℝ≥0} From 809f8f174180b2fd054795fc5d6f6fbcda013f89 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Mon, 28 Aug 2023 22:25:23 +0200 Subject: [PATCH 56/67] Merge --- Mathlib/GroupTheory/QuotientGroup.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/GroupTheory/QuotientGroup.lean b/Mathlib/GroupTheory/QuotientGroup.lean index 9fe02bb339dea..73b3156bbe01b 100644 --- a/Mathlib/GroupTheory/QuotientGroup.lean +++ b/Mathlib/GroupTheory/QuotientGroup.lean @@ -191,6 +191,11 @@ theorem mk_zpow (a : G) (n : ℤ) : ((a ^ n : G) : Q ) = (a : Q) ^ n := #align quotient_group.coe_zpow QuotientGroup.mk_zpow #align quotient_add_group.coe_zsmul QuotientAddGroup.mk_zsmul + @[to_additive (attr := simp)] + theorem mk_prod {G ι : Type _} [CommGroup G] (N : Subgroup G) (s : Finset ι) {f : ι → G} : + ((Finset.prod s f : G) : G ⧸ N) = Finset.prod s (fun i => (f i : G ⧸ N)) := + map_prod (QuotientGroup.mk' N) _ _ + /-- A group homomorphism `φ : G →* M` with `N ⊆ ker(φ)` descends (i.e. `lift`s) to a group homomorphism `G/N →* M`. -/ @[to_additive "An `AddGroup` homomorphism `φ : G →+ M` with `N ⊆ ker(φ)` descends (i.e. `lift`s) From 3bd156cee40d068c1a0de601611c4762ce2864c3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 1 Sep 2023 17:30:15 +0200 Subject: [PATCH 57/67] Backport --- .../NumberField/CanonicalEmbedding.lean | 61 ++++++++++--------- 1 file changed, 31 insertions(+), 30 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index 236602083a779..c3c3ed677e78d 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -325,48 +325,49 @@ local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) variable [NumberField K] /-- The fudge factor that appears in the formula for the volume of `convex_body_lt`. -/ -noncomputable def constant_factor : ℝ≥0∞ := +noncomputable abbrev constant_factor_lt : ℝ≥0∞ := (2 : ℝ≥0∞) ^ card {w : InfinitePlace K // IsReal w} * volume (ball (0 : ℂ) 1) ^ card {w : InfinitePlace K // IsComplex w} -theorem constant_factor_pos : 0 < (constant_factor K) := by +theorem constant_factor_lt_pos : 0 < (constant_factor_lt K) := by refine mul_pos (NeZero.ne _) ?_ exact ENNReal.pow_ne_zero (ne_of_gt (measure_ball_pos _ _ (by norm_num))) _ -theorem constant_factor_lt_top : (constant_factor K) < ⊤ := by +theorem constant_factor_lt_lt_top : (constant_factor_lt K) < ⊤ := by refine mul_lt_top ?_ ?_ · exact ne_of_lt (pow_lt_top (lt_top_iff_ne_top.mpr two_ne_top) _) · exact ne_of_lt (pow_lt_top measure_ball_lt_top _) -set_option maxHeartbeats 400000 in /-- The volume of `(convex_body_lt K f)` where `convex_body_lt K f` is the set of points `x` such that `‖x w‖ < f w` for all infinite places `w`. -/ -theorem convex_body_volume : - volume (convex_body_lt K f) = (constant_factor K) * ∏ w, (f w) ^ (mult w) := by - rw [volume_eq_prod, prod_prod, volume_pi, volume_pi, pi_pi, pi_pi] - conv_lhs => - congr; congr; next => skip - ext - rw [Real.volume_ball, ofReal_mul (by norm_num), ofReal_coe_nnreal, mul_comm] - conv_lhs => - congr; next => skip - congr; next => skip - ext i - rw [addHaar_ball _ _ (by exact (f i).prop), Complex.finrank_real_complex, ← NNReal.coe_pow, - ofReal_coe_nnreal, mul_comm] - rw [Finset.prod_mul_distrib, Finset.prod_mul_distrib, Finset.prod_const, Finset.prod_const, - Finset.card_univ, Finset.card_univ, mul_assoc, mul_comm, ← mul_assoc, mul_assoc, ofReal_ofNat, - ← constant_factor] - simp_rw [mult, pow_ite, pow_one] - rw [Finset.prod_ite] - simp_rw [coe_mul, coe_finset_prod] - simp_rw [show (fun w : InfinitePlace K ↦ ¬IsReal w) = (fun w ↦ IsComplex w) - by funext; rw [not_isReal_iff_isComplex]] - congr 1; rw [mul_comm]; congr 1 - all_goals - · rw [← Finset.prod_subtype_eq_prod_filter] - congr; ext - exact ⟨fun _ => Finset.mem_subtype.mpr (Finset.mem_univ _), fun _ => Finset.mem_univ _⟩ +theorem convex_body_lt_volume : + volume (convex_body_lt K f) = (constant_factor_lt K) * ∏ w, (f w) ^ (mult w) := by + calc + _ = (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (2 * (f x.val))) * + ∏ x : {w // InfinitePlace.IsComplex w}, volume (ball (0 : ℂ) 1) * + ENNReal.ofReal (f x.val) ^ 2 := by + simp_rw [volume_eq_prod, prod_prod, volume_pi, pi_pi, Real.volume_ball] + conv_lhs => + congr; next => skip + congr; next => skip + ext i; rw [addHaar_ball _ _ (by exact (f i).prop), Complex.finrank_real_complex, mul_comm, + ENNReal.ofReal_pow (by exact (f i).prop)] + _ = (↑2 ^ card {w : InfinitePlace K // InfinitePlace.IsReal w} * + (∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (f x.val))) * + (volume (ball (0 : ℂ) 1) ^ card {w : InfinitePlace K // IsComplex w} * + (∏ x : {w // IsComplex w}, ENNReal.ofReal (f x.val) ^ 2)) := by + simp_rw [ofReal_mul (by norm_num : 0 ≤ (2 : ℝ)), Finset.prod_mul_distrib, Finset.prod_const, + Finset.card_univ, ofReal_ofNat] + _ = (constant_factor_lt K) * ((∏ x : {w // InfinitePlace.IsReal w}, ENNReal.ofReal (f x.val)) * + (∏ x : {w // IsComplex w}, ENNReal.ofReal (f x.val) ^ 2)) := by ring + _ = (constant_factor_lt K) * ∏ w, (f w) ^ (mult w) := by + simp_rw [mult, pow_ite, pow_one, Finset.prod_ite, ofReal_coe_nnreal, not_isReal_iff_isComplex, + coe_mul, coe_finset_prod, ENNReal.coe_pow] + congr 2 + · refine (Finset.prod_subtype (Finset.univ.filter _) ?_ (fun w => (f w : ℝ≥0∞))).symm + exact fun _ => by simp only [Finset.mem_univ, forall_true_left, Finset.mem_filter, true_and] + · refine (Finset.prod_subtype (Finset.univ.filter _) ?_ (fun w => (f w : ℝ≥0∞) ^ 2)).symm + exact fun _ => by simp only [Finset.mem_univ, forall_true_left, Finset.mem_filter, true_and] variable {f} From 433c63ea9345c66efc8ad219690f61873947e36e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 16 Sep 2023 15:59:04 +0200 Subject: [PATCH 58/67] Clean up after merge --- Mathlib/NumberTheory/NumberField/Units.lean | 51 ++++++++++----------- 1 file changed, 25 insertions(+), 26 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index ec5a3ae2c443e..29d71ff4e85d8 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -136,8 +136,8 @@ def torsionOrder [NumberField K] : ℕ+ := ⟨Fintype.card (torsion K), Fintype. /-- If `k` does not divide `torsionOrder` then there are no nontrivial roots of unity of order dividing `k`. -/ -theorem rootsOfUnity_eq_one [NumberField K] {k : ℕ+} (hc : Nat.coprime k (torsionOrder K)) : - ζ ∈ rootsOfUnity k (𝓞 K) ↔ ζ = 1 := by +theorem rootsOfUnity_eq_one [NumberField K] {k : ℕ+} (hc : Nat.coprime k (torsionOrder K)) + {ζ : (𝓞 K)ˣ} : ζ ∈ rootsOfUnity k (𝓞 K) ↔ ζ = 1 := by rw [mem_rootsOfUnity] refine ⟨fun h => ?_, fun h => by rw [h, one_pow]⟩ refine orderOf_eq_one_iff.mp (Nat.eq_one_of_dvd_coprimes hc ?_ ?_) @@ -310,7 +310,7 @@ open NumberField.mixedEmbedding NNReal -- See: https://github.com/leanprover/lean4/issues/2220 local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -variable (w₁ : InfinitePlace K) {B : ℕ} (hB : minkowski_bound K < (constant_factor K) * B) +variable (w₁ : InfinitePlace K) {B : ℕ} (hB : minkowski_bound K < (constant_factor_lt K) * B) /-- This result shows that there always exists a next term in the sequence. -/ theorem seq.next {x : 𝓞 K} (hx : x ≠ 0) : @@ -320,18 +320,18 @@ theorem seq.next {x : 𝓞 K} (hx : x ≠ 0) : suffices ∀ w, w ≠ w₁ → f w ≠ 0 by obtain ⟨g, h_geqf, h_gprod⟩ := adjust_f K B this obtain ⟨y, h_ynz, h_yle⟩ := exists_ne_zero_mem_ringOfIntegers_lt (f := g) - (by rw [convex_body_volume]; convert hB; exact congrArg ((↑): NNReal → ENNReal) h_gprod) + (by rw [convex_body_lt_volume]; convert hB; exact congrArg ((↑): NNReal → ENNReal) h_gprod) refine ⟨y, h_ynz, fun w hw => (h_geqf w hw ▸ h_yle w).trans ?_, ?_⟩ · rw [← Rat.cast_le (K := ℝ), Rat.cast_coe_nat] calc - _ = ∏ w : InfinitePlace K, w y ^ mult w := (prod_eq_abs_norm (y : K)).symm - _ ≤ ∏ w : InfinitePlace K, (g w : ℝ) ^ mult w := ?_ - _ ≤ (B : ℝ) := ?_ - · refine Finset.prod_le_prod ?_ ?_ - exact fun _ _ => pow_nonneg (by positivity) _ - exact fun w _ => pow_le_pow_of_le_left (by positivity) (le_of_lt (h_yle w)) (mult w) - · simp_rw [← NNReal.coe_pow, ← NNReal.coe_prod] - exact le_of_eq (congrArg toReal h_gprod) + _ = ∏ w : InfinitePlace K, w y ^ mult w := (prod_eq_abs_norm (y : K)).symm + _ ≤ ∏ w : InfinitePlace K, (g w : ℝ) ^ mult w := by + refine Finset.prod_le_prod ?_ ?_ + · exact fun _ _ => pow_nonneg (by positivity) _ + . exact fun w _ => pow_le_pow_of_le_left (by positivity) (le_of_lt (h_yle w)) (mult w) + _ ≤ (B : ℝ) := by + simp_rw [← NNReal.coe_pow, ← NNReal.coe_prod] + exact le_of_eq (congrArg toReal h_gprod) · refine div_lt_self ?_ (by norm_num) simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, hx] intro _ _ @@ -392,10 +392,10 @@ image by the `log_embedding` of these units is `ℝ`-linearly independent, see `unit_lattice_span_eq_top`. -/ theorem exists_unit (w₁ : InfinitePlace K ) : ∃ u : (𝓞 K)ˣ, (∀ w : InfinitePlace K, w ≠ w₁ → Real.log (w u) < 0) := by - obtain ⟨B, hB⟩ : ∃ B : ℕ, minkowski_bound K < (constant_factor K) * B := by + obtain ⟨B, hB⟩ : ∃ B : ℕ, minkowski_bound K < (constant_factor_lt K) * B := by simp_rw [mul_comm] refine ENNReal.exists_nat_mul_gt ?_ ?_ - exact ne_of_gt (constant_factor_pos K) + exact ne_of_gt (constant_factor_lt_pos K) exact ne_of_lt (minkowski_bound_lt_top 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) })) @@ -403,13 +403,13 @@ theorem exists_unit (w₁ : InfinitePlace K ) : refine ⟨hu.choose, fun w hw => Real.log_neg ?_ ?_⟩ · simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, ne_zero] · calc - _ = w ((seq K w₁ hB m : K) * (seq K w₁ hB n : K)⁻¹) := ?_ - _ = w (seq K w₁ hB m) * w (seq K w₁ hB n)⁻¹ := _root_.map_mul _ _ _ - _ < 1 := ?_ - · rw [← congrArg ((↑) : (𝓞 K) → K) hu.choose_spec, mul_comm, Submonoid.coe_mul, ← mul_assoc, - inv_mul_cancel (seq.ne_zero K w₁ hB n), one_mul] - · rw [map_inv₀, mul_inv_lt_iff (pos_iff.mpr (seq.ne_zero K w₁ hB n)), mul_one] - exact seq.antitone K w₁ hB hnm w hw + _ = w ((seq K w₁ hB m : K) * (seq K w₁ hB n : K)⁻¹) := by + rw [← congrArg ((↑) : (𝓞 K) → K) hu.choose_spec, mul_comm, Submonoid.coe_mul, ← mul_assoc, + inv_mul_cancel (seq.ne_zero K w₁ hB n), one_mul] + _ = w (seq K w₁ hB m) * w (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.antitone 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 => ?_) ?_ @@ -528,12 +528,11 @@ theorem fun_eq_repr {x ζ : (𝓞 K)ˣ} {f : Fin (rank K) → ℤ} (hζ : ζ ∈ rw [← (basis_mod_torsion K).repr_sum_self f, ← this] calc Additive.ofMul ↑x = ∑ i, (f i) • Additive.ofMul ↑(fund_system K i) := by - rw [h, QuotientGroup.mk_mul, (QuotientGroup.eq_one_iff _).mpr hζ, one_mul, - QuotientGroup.mk_prod, ofMul_prod]; rfl - _ = ∑ i, (f i) • (basis_mod_torsion K i) := by - simp_rw [fund_system, QuotientGroup.out_eq', ofMul_toMul] + rw [h, QuotientGroup.mk_mul, (QuotientGroup.eq_one_iff _).mpr hζ, one_mul, + QuotientGroup.mk_prod, ofMul_prod]; rfl + _ = ∑ i, (f i) • (basis_mod_torsion K i) := by + simp_rw [fund_system, QuotientGroup.out_eq', ofMul_toMul] -set_option maxHeartbeats 300000 in /-- Any unit `x` of `𝓞 K` can be written uniquely as a root of unity times the product of powers of the units of the fundamental system. -/ theorem exist_unique_eq_mul_prod (x : (𝓞 K)ˣ) : ∃! (ζ : torsion K) (e : Fin (rank K) → ℤ), From a7d7a8f480e31f74661e0a487581e96019ca3c23 Mon Sep 17 00:00:00 2001 From: Xavier Roblot <46200072+xroblot@users.noreply.github.com> Date: Sat, 16 Sep 2023 17:19:16 +0200 Subject: [PATCH 59/67] Update Mathlib/NumberTheory/NumberField/Units.lean Co-authored-by: github-actions[bot] <41898282+github-actions[bot]@users.noreply.github.com> --- Mathlib/NumberTheory/NumberField/Units.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 29d71ff4e85d8..62389201c7171 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -328,7 +328,7 @@ theorem seq.next {x : 𝓞 K} (hx : x ≠ 0) : _ ≤ ∏ w : InfinitePlace K, (g w : ℝ) ^ mult w := by refine Finset.prod_le_prod ?_ ?_ · exact fun _ _ => pow_nonneg (by positivity) _ - . exact fun w _ => pow_le_pow_of_le_left (by positivity) (le_of_lt (h_yle w)) (mult w) + · exact fun w _ => pow_le_pow_of_le_left (by positivity) (le_of_lt (h_yle w)) (mult w) _ ≤ (B : ℝ) := by simp_rw [← NNReal.coe_pow, ← NNReal.coe_prod] exact le_of_eq (congrArg toReal h_gprod) From b786b22276767640b9625dd4232708907cb211a8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Thu, 21 Sep 2023 17:38:02 +0200 Subject: [PATCH 60/67] Clean up after merge --- .../NumberField/CanonicalEmbedding.lean | 2 +- Mathlib/NumberTheory/NumberField/Units.lean | 190 +++++++++--------- 2 files changed, 96 insertions(+), 96 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean index d95a2afc3c268..569dfe9100a27 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding.lean @@ -329,7 +329,7 @@ noncomputable abbrev convexBodyLtFactor : ℝ≥0∞ := (2 : ℝ≥0∞) ^ card {w : InfinitePlace K // IsReal w} * volume (ball (0 : ℂ) 1) ^ card {w : InfinitePlace K // IsComplex w} -theorem convexBodyLtFactor_lt_pos : 0 < (convexBodyLtFactor K) := by +theorem convexBodyLtFactor_pos : 0 < (convexBodyLtFactor K) := by refine mul_pos (NeZero.ne _) ?_ exact ENNReal.pow_ne_zero (ne_of_gt (measure_ball_pos _ _ (by norm_num))) _ diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 62389201c7171..2255ba6bd7422 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -20,9 +20,9 @@ field `K`. * `Units.rank`: the unit rank of the number field `K`. -* `Units.fund_system`: a fundamental system of units of `K`. +* `Units.fundSystem`: a fundamental system of units of `K`. -* `Units.basis_mod_torsion`: a `ℤ`-basis of `(𝓞 K)ˣ ⧸ (torsion K)` (as an additive `ℤ`-module). +* `Units.basisModTorsion`: a `ℤ`-basis of `(𝓞 K)ˣ ⧸ (torsion K)` (as an additive `ℤ`-module). ## Main results @@ -136,7 +136,7 @@ def torsionOrder [NumberField K] : ℕ+ := ⟨Fintype.card (torsion K), Fintype. /-- If `k` does not divide `torsionOrder` then there are no nontrivial roots of unity of order dividing `k`. -/ -theorem rootsOfUnity_eq_one [NumberField K] {k : ℕ+} (hc : Nat.coprime k (torsionOrder K)) +theorem rootsOfUnity_eq_one [NumberField K] {k : ℕ+} (hc : Nat.Coprime k (torsionOrder K)) {ζ : (𝓞 K)ˣ} : ζ ∈ rootsOfUnity k (𝓞 K) ↔ ζ = 1 := by rw [mem_rootsOfUnity] refine ⟨fun h => ?_, fun h => by rw [h, one_pow]⟩ @@ -165,12 +165,12 @@ namespace dirichlet -- This section is devoted to the proof of Dirichlet's unit theorem -- We define a group morphism from `(𝓞 K)ˣ` to `{w : InfinitePlace K // w ≠ w₀} → ℝ` where `w₀` -- is a distinguished (arbitrary) infinite place, prove that its kernel is the torsion subgroup --- (see `log_embedding_eq_zero_iff`) and that its image, called `unit_lattice`, is a full --- `ℤ`-lattice. It follows that is a free `ℤ`-module (see `unit_lattice_moduleFree `) of --- rank `card (InfinitePlaces K) - 1` (see `unit_lattice_rank`). --- To prove that the `unit_lattice` is a full `ℤ`-lattice, we need to prove that it is discrete --- (see `unit_lattice_inter_ball_finite`) and that it spans the full space over `ℝ` --- (see ` unit_lattice_span_eq_top`); this is the main part of the proof, see the section +-- (see `logEmbedding_eq_zero_iff`) and that its image, called `unitLattice`, is a full +-- `ℤ`-lattice. It follows that is a free `ℤ`-module (see `unitLattice_moduleFree `) of +-- rank `card (InfinitePlaces K) - 1` (see `unitLattice_rank`). +-- To prove that the `unitLattice` is a full `ℤ`-lattice, we need to prove that it is discrete +-- (see `unitLattice_inter_ball_finite`) and that it spans the full space over `ℝ` +-- (see ` unitLattice_span_eq_top`); this is the main part of the proof, see the section -- `span_top` below for more details. open scoped Classical @@ -185,7 +185,7 @@ def w₀ : InfinitePlace K := (inferInstance : Nonempty (InfinitePlace K)).some variable (K) /-- The logarithmic embedding of the units (seen as an `Additive` group). -/ -def log_embedding : Additive ((𝓞 K)ˣ) →+ ({w : InfinitePlace K // w ≠ w₀} → ℝ) := +def logEmbedding : Additive ((𝓞 K)ˣ) →+ ({w : InfinitePlace K // w ≠ w₀} → ℝ) := { toFun := fun x w => mult w.val * Real.log (w.val (Additive.toMul x)) map_zero' := by simp; rfl map_add' := by @@ -196,12 +196,12 @@ def log_embedding : Additive ((𝓞 K)ˣ) →+ ({w : InfinitePlace K // w ≠ w variable {K} @[simp] -theorem log_embedding_component (x : (𝓞 K)ˣ) (w : {w : InfinitePlace K // w ≠ w₀}) : - (log_embedding K x) w = mult w.val * Real.log (w.val x) := rfl +theorem logEmbedding_component (x : (𝓞 K)ˣ) (w : {w : InfinitePlace K // w ≠ w₀}) : + (logEmbedding K x) w = mult w.val * Real.log (w.val x) := rfl -theorem log_embedding_sum_component (x : (𝓞 K)ˣ) : - ∑ w, log_embedding K x w = - mult (w₀ : InfinitePlace K) * Real.log (w₀ (x : K)) := by - have h := congrArg Real.log (prod_eq_abs_norm (x : K)) +theorem logEmbedding_sum_component (x : (𝓞 K)ˣ) : + ∑ w, logEmbedding K x w = - mult (w₀ : InfinitePlace K) * Real.log (w₀ (x : K)) := by + have h := congr_arg Real.log (prod_eq_abs_norm (x : K)) rw [show |(Algebra.norm ℚ) (x : K)| = 1 from isUnit_iff_norm.mp x.isUnit, Rat.cast_one, Real.log_one, Real.log_prod] at h · simp_rw [Real.log_pow] at h @@ -222,46 +222,46 @@ theorem mult_log_place_eq_zero {x : (𝓞 K)ˣ} {w : InfinitePlace K} : · refine (ne_of_gt ?_) rw [mult]; split_ifs <;> norm_num -theorem log_embedding_eq_zero_iff (x : (𝓞 K)ˣ) : - log_embedding K x = 0 ↔ x ∈ torsion K := by +theorem logEmbedding_eq_zero_iff (x : (𝓞 K)ˣ) : + logEmbedding K x = 0 ↔ x ∈ torsion K := by rw [mem_torsion] 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 [← log_embedding_sum_component, Finset.sum_eq_zero] + rw [← logEmbedding_sum_component, Finset.sum_eq_zero] exact fun w _ => congrFun h w · exact mult_log_place_eq_zero.mp (congrFun h ⟨w, hw⟩) · ext w - rw [log_embedding_component, h w.val, Real.log_one, mul_zero, Pi.zero_apply] + rw [logEmbedding_component, h w.val, Real.log_one, mul_zero, Pi.zero_apply] -theorem log_embedding_component_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖log_embedding K x‖ ≤ r) - (w : {w : InfinitePlace K // w ≠ w₀}) : |log_embedding K x w| ≤ r := by +theorem logEmbedding_component_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖logEmbedding K x‖ ≤ r) + (w : {w : InfinitePlace K // w ≠ w₀}) : |logEmbedding K x w| ≤ r := by lift r to NNReal using hr simp_rw [Pi.norm_def, NNReal.coe_le_coe, Finset.sup_le_iff, ← NNReal.coe_le_coe] at h exact h w (Finset.mem_univ _) -theorem log_le_of_log_embedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖log_embedding K x‖ ≤ r) +theorem log_le_of_logEmbedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖logEmbedding K x‖ ≤ r) (w : InfinitePlace K) : |Real.log (w x)| ≤ (Fintype.card (InfinitePlace K)) * r := by have tool : ∀ x : ℝ, 0 ≤ x → x ≤ mult w * x := fun x hx => by nth_rw 1 [← one_mul x] refine mul_le_mul ?_ le_rfl hx ?_ all_goals { rw [mult]; split_ifs <;> norm_num } by_cases hw : w = w₀ - · have hyp := congrArg (‖·‖) (log_embedding_sum_component x).symm + · have hyp := congr_arg (‖·‖) (logEmbedding_sum_component x).symm replace hyp := (le_of_eq hyp).trans (norm_sum_le _ _) simp_rw [norm_mul, norm_neg, Real.norm_eq_abs, Nat.abs_cast] at hyp refine (le_trans ?_ hyp).trans ?_ · rw [← hw] exact tool _ (abs_nonneg _) · refine (Finset.sum_le_card_nsmul Finset.univ _ _ - (fun w _ => log_embedding_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 [Finset.card_univ] - · have hyp := log_embedding_component_le hr h ⟨w, hw⟩ - rw [log_embedding_component, abs_mul, Nat.abs_cast] at hyp + · have hyp := logEmbedding_component_le hr h ⟨w, hw⟩ + rw [logEmbedding_component, abs_mul, Nat.abs_cast] at hyp refine (le_trans ?_ hyp).trans ?_ · exact tool _ (abs_nonneg _) · nth_rw 1 [← one_mul r] @@ -270,11 +270,11 @@ theorem log_le_of_log_embedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h variable (K) /-- The lattice formed by the image of the logarithmic embedding. -/ -noncomputable def unit_lattice : AddSubgroup ({w : InfinitePlace K // w ≠ w₀} → ℝ) := - AddSubgroup.map (log_embedding K) ⊤ +noncomputable def unitLattice : AddSubgroup ({w : InfinitePlace K // w ≠ w₀} → ℝ) := + AddSubgroup.map (logEmbedding K) ⊤ -theorem unit_lattice_inter_ball_finite (r : ℝ) : - ((unit_lattice K : Set ({ w : InfinitePlace K // w ≠ w₀} → ℝ)) ∩ +theorem unitLattice_inter_ball_finite (r : ℝ) : + ((unitLattice K : Set ({ w : InfinitePlace K // w ≠ w₀} → ℝ)) ∩ Metric.closedBall 0 r).Finite := by obtain hr | hr := lt_or_le r 0 · convert Set.finite_empty @@ -282,12 +282,12 @@ theorem unit_lattice_inter_ball_finite (r : ℝ) : exact Set.inter_empty _ · suffices Set.Finite {x : (𝓞 K)ˣ | IsIntegral ℤ (x : K) ∧ ∀ (φ : K →+* ℂ), ‖φ x‖ ≤ Real.exp ((Fintype.card (InfinitePlace K)) * r)} by - refine (Set.Finite.image (log_embedding K) this).subset ?_ + refine (Set.Finite.image (logEmbedding K) this).subset ?_ rintro _ ⟨⟨x, ⟨_, rfl⟩⟩, hx⟩ refine ⟨x, ⟨x.val.prop, (le_iff_le _ _).mp (fun w => (Real.log_le_iff_le_exp ?_).mp ?_)⟩, rfl⟩ · exact pos_iff.mpr (coe_ne_zero x) · rw [mem_closedBall_zero_iff] at hx - exact (le_abs_self _).trans (log_le_of_log_embedding_le hr hx w) + exact (le_abs_self _).trans (log_le_of_logEmbedding_le hr hx w) refine Set.Finite.of_finite_image ?_ ((coe_injective K).injOn _) refine (Embeddings.finite_of_norm_le K ℂ (Real.exp ((Fintype.card (InfinitePlace K)) * r))).subset ?_ @@ -299,7 +299,7 @@ section span_top -- for each infinite place `w₁ ≠ w₀` an unit `u_w₁` of `K` such that, for all infinite places -- `w` such that `w ≠ w₁`, we have `Real.log w (u_w₁) < 0` (and thus `Real.log w₁ (u_w₁) > 0`). -- It follows then from a determinant computation (using `Matrix.det_ne_zero_of_neg`) that the --- image by `log_embedding` of these units is a `ℝ`-linearly independent family. +-- image by `logEmbedding` of these units is a `ℝ`-linearly independent family. -- The unit `u_w₁` is obtained by construction a sequence `seq n` of nonzero algebraic integers -- that is strictly decreasing at infinite places distinct from `w₁` and of norm `≤ B`. Since -- there are finitely many ideals of norm `≤ B`, there exists two terms in the sequence defining @@ -310,17 +310,17 @@ open NumberField.mixedEmbedding NNReal -- See: https://github.com/leanprover/lean4/issues/2220 local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -variable (w₁ : InfinitePlace K) {B : ℕ} (hB : minkowski_bound K < (constant_factor_lt K) * B) +variable (w₁ : InfinitePlace K) {B : ℕ} (hB : minkowskiBound K < (convexBodyLtFactor K) * B) /-- This result shows that there always exists a next term in the sequence. -/ -theorem seq.next {x : 𝓞 K} (hx : x ≠ 0) : +theorem seq_next {x : 𝓞 K} (hx : x ≠ 0) : ∃ y : 𝓞 K, y ≠ 0 ∧ (∀ w, w ≠ w₁ → w y < w x) ∧ |Algebra.norm ℚ (y : K)| ≤ B := by let f : InfinitePlace K → ℝ≥0 := fun w => ⟨(w x) / 2, div_nonneg (AbsoluteValue.nonneg _ _) (by norm_num)⟩ suffices ∀ w, w ≠ w₁ → f w ≠ 0 by obtain ⟨g, h_geqf, h_gprod⟩ := adjust_f K B this obtain ⟨y, h_ynz, h_yle⟩ := exists_ne_zero_mem_ringOfIntegers_lt (f := g) - (by rw [convex_body_lt_volume]; convert hB; exact congrArg ((↑): NNReal → ENNReal) h_gprod) + (by rw [convexBodyLt_volume]; convert hB; exact congr_arg ((↑): NNReal → ENNReal) h_gprod) refine ⟨y, h_ynz, fun w hw => (h_geqf w hw ▸ h_yle w).trans ?_, ?_⟩ · rw [← Rat.cast_le (K := ℝ), Rat.cast_coe_nat] calc @@ -331,7 +331,7 @@ theorem seq.next {x : 𝓞 K} (hx : x ≠ 0) : · exact fun w _ => pow_le_pow_of_le_left (by positivity) (le_of_lt (h_yle w)) (mult w) _ ≤ (B : ℝ) := by simp_rw [← NNReal.coe_pow, ← NNReal.coe_prod] - exact le_of_eq (congrArg toReal h_gprod) + exact le_of_eq (congr_arg toReal h_gprod) · refine div_lt_self ?_ (by norm_num) simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, hx] intro _ _ @@ -345,15 +345,15 @@ theorem seq.next {x : 𝓞 K} (hx : x ≠ 0) : def seq : ℕ → { x : 𝓞 K // x ≠ 0 } | 0 => ⟨1, by norm_num⟩ | n + 1 => - ⟨(seq.next K w₁ hB (seq n).prop).choose, (seq.next K w₁ hB (seq n).prop).choose_spec.1⟩ + ⟨(seq_next K w₁ hB (seq n).prop).choose, (seq_next K w₁ hB (seq n).prop).choose_spec.1⟩ /-- The terms of the sequence are nonzero. -/ -theorem seq.ne_zero (n : ℕ) : (seq K w₁ hB n : K) ≠ 0 := by +theorem seq_ne_zero (n : ℕ) : (seq K w₁ hB n : K) ≠ 0 := by refine (map_ne_zero_iff (algebraMap (𝓞 K) K) ?_).mpr (seq K w₁ hB n).prop exact IsFractionRing.injective { x // x ∈ 𝓞 K } K -/-- The sequence is strictly decreasing at infinite places different from `w₁`. -/ -theorem seq.antitone {n m : ℕ} (h : n < m) : +/-- The sequence is strictly decreasing at infinite places distinct from `w₁`. -/ +theorem seq_antitone {n m : ℕ} (h : n < m) : ∀ w : InfinitePlace K, w ≠ w₁ → w (seq K w₁ hB m) < w (seq K w₁ hB n) := by induction m with | zero => @@ -364,13 +364,13 @@ theorem seq.antitone {n m : ℕ} (h : n < m) : cases eq_or_lt_of_le (Nat.le_of_lt_succ h) with | inl hr => rw [hr] - exact (seq.next K w₁ hB (seq K w₁ hB m).prop).choose_spec.2.1 w hw + exact (seq_next K w₁ hB (seq K w₁ hB m).prop).choose_spec.2.1 w hw | inr hr => refine lt_trans ?_ (m_ih hr w hw) - exact (seq.next K w₁ hB (seq K w₁ hB m).prop).choose_spec.2.1 w hw + exact (seq_next K w₁ hB (seq K w₁ hB m).prop).choose_spec.2.1 w hw /-- The terms of the sequence have norm bounded by `B`. -/ -theorem seq.norm_bdd (n : ℕ) : +theorem seq_norm_bdd (n : ℕ) : 1 ≤ Int.natAbs (Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K)) ∧ Int.natAbs (Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K)) ≤ B := by cases n with @@ -385,18 +385,18 @@ theorem seq.norm_bdd (n : ℕ) : · rw [← Nat.cast_le (α := ℚ), Int.cast_natAbs, Int.cast_abs] change |algebraMap ℤ ℚ _| ≤ _ rw [← Algebra.norm_localization ℤ (Sₘ := K) (nonZeroDivisors ℤ)] - exact (seq.next K w₁ hB (seq K w₁ hB n).prop).choose_spec.2.2 + exact (seq_next K w₁ hB (seq K w₁ hB n).prop).choose_spec.2.2 /-- Construct an unit associated to the place `w₁`. The family, for `w₁ ≠ w₀`, formed by the -image by the `log_embedding` of these units is `ℝ`-linearly independent, see +image by the `logEmbedding` of these units is `ℝ`-linearly independent, see `unit_lattice_span_eq_top`. -/ theorem exists_unit (w₁ : InfinitePlace K ) : ∃ u : (𝓞 K)ˣ, (∀ w : InfinitePlace K, w ≠ w₁ → Real.log (w u) < 0) := by - obtain ⟨B, hB⟩ : ∃ B : ℕ, minkowski_bound K < (constant_factor_lt K) * B := by + obtain ⟨B, hB⟩ : ∃ B : ℕ, minkowskiBound K < (convexBodyLtFactor K) * B := by simp_rw [mul_comm] refine ENNReal.exists_nat_mul_gt ?_ ?_ - exact ne_of_gt (constant_factor_lt_pos K) - exact ne_of_lt (minkowski_bound_lt_top K) + exact ne_of_gt (convexBodyLtFactor_pos K) + exact ne_of_lt (minkowskiBound_lt_top 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 @@ -404,28 +404,28 @@ theorem exists_unit (w₁ : InfinitePlace K ) : · simp only [pos_iff, ne_eq, ZeroMemClass.coe_eq_zero, ne_zero] · calc _ = w ((seq K w₁ hB m : K) * (seq K w₁ hB n : K)⁻¹) := by - rw [← congrArg ((↑) : (𝓞 K) → K) hu.choose_spec, mul_comm, Submonoid.coe_mul, ← mul_assoc, - inv_mul_cancel (seq.ne_zero K w₁ hB n), one_mul] + rw [← congr_arg ((↑) : (𝓞 K) → K) hu.choose_spec, mul_comm, Submonoid.coe_mul, + ← mul_assoc, inv_mul_cancel (seq_ne_zero K w₁ hB n), one_mul] _ = w (seq K w₁ hB m) * w (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.antitone K w₁ hB hnm w hw + rw [map_inv₀, mul_inv_lt_iff (pos_iff.mpr (seq_ne_zero K w₁ hB n)), mul_one] + exact seq_antitone 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] - exact seq.norm_bdd K w₁ hB n + exact seq_norm_bdd 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) -theorem unit_lattice_span_eq_top : - Submodule.span ℝ (unit_lattice K : Set ({w : InfinitePlace K // w ≠ w₀} → ℝ)) = ⊤ := by +theorem unitLattice_span_eq_top : + Submodule.span ℝ (unitLattice K : Set ({w : InfinitePlace K // w ≠ w₀} → ℝ)) = ⊤ := by refine le_antisymm (le_top) ?_ -- The standard basis let B := Pi.basisFun ℝ {w : InfinitePlace K // w ≠ w₀} -- The image by log_embedding of the family of units constructed above - let v := fun w : { w : InfinitePlace K // w ≠ w₀ } => log_embedding K (exists_unit K w).choose + let v := fun w : { w : InfinitePlace K // w ≠ w₀ } => logEmbedding K (exists_unit K w).choose -- To prove the result, it is enough to prove that the family `v` is linearly independent suffices B.det v ≠ 0 by rw [← isUnit_iff_ne_zero, ← is_basis_iff_det] at this @@ -441,7 +441,7 @@ theorem unit_lattice_span_eq_top : refine add_pos_of_nonneg_of_pos ?_ ?_ · rw [sub_nonneg] exact le_abs_self _ - · rw [log_embedding_sum_component (exists_unit K w).choose] + · rw [logEmbedding_sum_component (exists_unit K w).choose] refine mul_pos_of_neg_of_neg ?_ ((exists_unit K w).choose_spec _ w.prop.symm) rw [mult]; split_ifs <;> norm_num · refine mul_neg_of_pos_of_neg ?_ ((exists_unit K w).choose_spec x ?_) @@ -455,50 +455,50 @@ def _root_.NumberField.Units.rank : ℕ := Fintype.card (InfinitePlace K) - 1 open FiniteDimensional -theorem unit_lattice_discrete : DiscreteTopology (unit_lattice K) := by +theorem unitLattice_discrete : DiscreteTopology (unitLattice K) := by refine discreteTopology_of_open_singleton_zero ?_ refine isOpen_singleton_of_finite_mem_nhds 0 (s := Metric.closedBall 0 1) ?_ ?_ · exact Metric.closedBall_mem_nhds _ (by norm_num) · refine Set.Finite.of_finite_image ?_ (Set.injOn_of_injective Subtype.val_injective _) - convert unit_lattice_inter_ball_finite K 1 + convert unitLattice_inter_ball_finite K 1 ext x refine ⟨?_, fun ⟨hx1, hx2⟩ => ⟨⟨x, hx1⟩, hx2, rfl⟩⟩ rintro ⟨x, hx, rfl⟩ exact ⟨Subtype.mem x, hx⟩ -theorem rank_space : +protected theorem finrank : finrank ℝ ({w : InfinitePlace K // w ≠ w₀} → ℝ) = Units.rank K := by simp only [finrank_fintype_fun_eq_card, Fintype.card_subtype_compl, Fintype.card_ofSubsingleton, rank] -theorem unit_lattice_moduleFree : Module.Free ℤ (unit_lattice K) := by - have := unit_lattice_discrete K - exact Zlattice.module_free ℝ (unit_lattice_span_eq_top K) +theorem unitLattice_moduleFree : Module.Free ℤ (unitLattice K) := by + have := unitLattice_discrete K + exact Zlattice.module_free ℝ (unitLattice_span_eq_top K) -theorem unit_lattice_moduleFinite : Module.Finite ℤ (unit_lattice K) := by - have := unit_lattice_discrete K - exact Zlattice.module_finite ℝ (unit_lattice_span_eq_top K) +theorem unitLattice_moduleFinite : Module.Finite ℤ (unitLattice K) := by + have := unitLattice_discrete K + exact Zlattice.module_finite ℝ (unitLattice_span_eq_top K) -theorem unit_lattice_rank : finrank ℤ (unit_lattice K) = Units.rank K := by - have := unit_lattice_discrete K - rw [← rank_space] - exact Zlattice.rank ℝ (unit_lattice_span_eq_top K) +theorem unitLattice_rank : finrank ℤ (unitLattice K) = Units.rank K := by + have := unitLattice_discrete K + rw [← dirichlet.finrank] + exact Zlattice.rank ℝ (unitLattice_span_eq_top K) end dirichlet variable [NumberField K] /-- A basis of the quotient `(𝓞 K)ˣ ⧸ (torsion K)` seen as an additive ℤ-module. -/ -def basis_mod_torsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by - let f : (dirichlet.unit_lattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (torsion K)) := by +def basisModTorsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by + let f : (dirichlet.unitLattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (torsion K)) := by refine AddEquiv.toIntLinearEquiv ?_ - rw [dirichlet.unit_lattice, ← AddMonoidHom.range_eq_map (dirichlet.log_embedding K)] - refine (QuotientAddGroup.quotientKerEquivRange (dirichlet.log_embedding K)).symm.trans ?_ + rw [dirichlet.unitLattice, ← AddMonoidHom.range_eq_map (dirichlet.logEmbedding K)] + refine (QuotientAddGroup.quotientKerEquivRange (dirichlet.logEmbedding K)).symm.trans ?_ refine (QuotientAddGroup.quotientAddEquivOfEq ?_).trans (QuotientAddGroup.quotientKerEquivOfSurjective (MonoidHom.toAdditive (QuotientGroup.mk' (torsion K))) (fun x => ?_)) · ext - rw [AddMonoidHom.mem_ker, AddMonoidHom.mem_ker, dirichlet.log_embedding_eq_zero_iff, + rw [AddMonoidHom.mem_ker, AddMonoidHom.mem_ker, dirichlet.logEmbedding_eq_zero_iff, MonoidHom.toAdditive_apply_apply, ofMul_eq_zero, QuotientGroup.mk'_apply, QuotientGroup.eq_one_iff] rfl @@ -507,45 +507,45 @@ def basis_mod_torsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (tors QuotientGroup.out_eq'] rfl have : Module.Free ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := - (dirichlet.unit_lattice_moduleFree K).of_equiv' f + (dirichlet.unitLattice_moduleFree K).of_equiv' f have : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by - have := dirichlet.unit_lattice_moduleFinite K + have := dirichlet.unitLattice_moduleFinite K exact Module.Finite.equiv f have : FiniteDimensional.finrank ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) = rank K := by - rw [← LinearEquiv.finrank_eq f, dirichlet.unit_lattice_rank] + rw [← LinearEquiv.finrank_eq f, dirichlet.unitLattice_rank] refine Basis.reindex (Module.Free.chooseBasis ℤ _) (Fintype.equivOfCardEq ?_) rw [← FiniteDimensional.finrank_eq_card_chooseBasisIndex, this, Fintype.card_fin] -/-- A fundamental system of units of `K`. The units of `fund_system` are arbitrary lifts of the -units in `basis_mod_torsion`. -/ -def fund_system : Fin (rank K) → (𝓞 K)ˣ := - fun i => Quotient.out' (Additive.toMul (basis_mod_torsion K i)) +/-- A fundamental system of units of `K`. The units of `fundSystem` are arbitrary lifts of the +units in `basisModTorsion`. -/ +def fundSystem : Fin (rank K) → (𝓞 K)ˣ := + fun i => Quotient.out' (Additive.toMul (basisModTorsion K i)) theorem fun_eq_repr {x ζ : (𝓞 K)ˣ} {f : Fin (rank K) → ℤ} (hζ : ζ ∈ torsion K) - (h : x = ζ * ∏ i, (fund_system K i) ^ (f i)) : - f = (basis_mod_torsion K).repr (Additive.ofMul ↑x) := by - suffices Additive.ofMul ↑x = ∑ i, (f i) • (basis_mod_torsion K i) by - rw [← (basis_mod_torsion K).repr_sum_self f, ← this] + (h : x = ζ * ∏ i, (fundSystem K i) ^ (f i)) : + f = (basisModTorsion K).repr (Additive.ofMul ↑x) := by + suffices Additive.ofMul ↑x = ∑ i, (f i) • (basisModTorsion K i) by + rw [← (basisModTorsion K).repr_sum_self f, ← this] calc - Additive.ofMul ↑x = ∑ i, (f i) • Additive.ofMul ↑(fund_system K i) := by + Additive.ofMul ↑x = ∑ i, (f i) • Additive.ofMul ↑(fundSystem K i) := by rw [h, QuotientGroup.mk_mul, (QuotientGroup.eq_one_iff _).mpr hζ, one_mul, QuotientGroup.mk_prod, ofMul_prod]; rfl - _ = ∑ i, (f i) • (basis_mod_torsion K i) := by - simp_rw [fund_system, QuotientGroup.out_eq', ofMul_toMul] + _ = ∑ i, (f i) • (basisModTorsion K i) := by + simp_rw [fundSystem, QuotientGroup.out_eq', ofMul_toMul] /-- Any unit `x` of `𝓞 K` can be written uniquely as a root of unity times the product of powers of the units of the fundamental system. -/ theorem exist_unique_eq_mul_prod (x : (𝓞 K)ˣ) : ∃! (ζ : torsion K) (e : Fin (rank K) → ℤ), - x = ζ * ∏ i, (fund_system K i) ^ (e i) := by - let ζ := x * (∏ i, (fund_system K i) ^ ((basis_mod_torsion K).repr (Additive.ofMul ↑x) i))⁻¹ + x = ζ * ∏ i, (fundSystem K i) ^ (e i) := by + let ζ := x * (∏ i, (fundSystem K i) ^ ((basisModTorsion K).repr (Additive.ofMul ↑x) i))⁻¹ have h_tors : ζ ∈ torsion K := by rw [← QuotientGroup.eq_one_iff, QuotientGroup.mk_mul, QuotientGroup.mk_inv, ← ofMul_eq_zero, ofMul_mul, ofMul_inv, QuotientGroup.mk_prod, ofMul_prod] - simp_rw [QuotientGroup.mk_zpow, ofMul_zpow, fund_system, QuotientGroup.out_eq', ofMul_toMul] + simp_rw [QuotientGroup.mk_zpow, ofMul_zpow, fundSystem, QuotientGroup.out_eq', ofMul_toMul] rw [add_eq_zero_iff_eq_neg, neg_neg] - exact ((basis_mod_torsion K).sum_repr (Additive.ofMul ↑x)).symm + exact ((basisModTorsion K).sum_repr (Additive.ofMul ↑x)).symm refine ⟨⟨ζ, h_tors⟩, ?_, ?_⟩ - · refine ⟨((basis_mod_torsion K).repr (Additive.ofMul ↑x) : Fin (rank K) → ℤ), ?_, ?_⟩ + · refine ⟨((basisModTorsion K).repr (Additive.ofMul ↑x) : Fin (rank K) → ℤ), ?_, ?_⟩ · simp only [_root_.inv_mul_cancel_right] · exact fun _ hf => fun_eq_repr K h_tors hf · rintro η ⟨_, hf, _⟩ From b61a7b7b7343f9bc40cc78686a368e100b97d6f8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Fri, 22 Sep 2023 17:00:57 +0200 Subject: [PATCH 61/67] Review --- Mathlib/NumberTheory/NumberField/Units.lean | 64 +++++++++++---------- 1 file changed, 33 insertions(+), 31 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 2255ba6bd7422..2d009430f24d6 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -28,9 +28,9 @@ field `K`. * `isUnit_iff_norm`: an algebraic integer `x : 𝓞 K` is a unit if and only if `|norm ℚ x| = 1`. -* `Units.mem_torsion`: an unit `x : (𝓞 K)ˣ` is torsion iff `w x = 1` for all infinite places of `K`. +* `Units.mem_torsion`: a unit `x : (𝓞 K)ˣ` is torsion iff `w x = 1` for all infinite places of `K`. -* `Units.exist_unique_eq_mul_prod`: an unit `x : (𝓞 K)ˣ` can be written in a unique way as the +* `Units.exist_unique_eq_mul_prod`: a unit `x : (𝓞 K)ˣ` can be written in a unique way as the product of a root of unity times the product of powers of units of the fundamental system. ## Tags @@ -161,7 +161,8 @@ theorem rootsOfUnity_eq_torsion [NumberField K] : end torsion -namespace dirichlet +namespace dirichletUnitTheorem + -- This section is devoted to the proof of Dirichlet's unit theorem -- We define a group morphism from `(𝓞 K)ˣ` to `{w : InfinitePlace K // w ≠ w₀} → ℝ` where `w₀` -- is a distinguished (arbitrary) infinite place, prove that its kernel is the torsion subgroup @@ -199,7 +200,7 @@ variable {K} theorem logEmbedding_component (x : (𝓞 K)ˣ) (w : {w : InfinitePlace K // w ≠ w₀}) : (logEmbedding K x) w = mult w.val * Real.log (w.val x) := rfl -theorem logEmbedding_sum_component (x : (𝓞 K)ˣ) : +theorem sum_logEmbedding_component (x : (𝓞 K)ˣ) : ∑ w, logEmbedding K x w = - mult (w₀ : InfinitePlace K) * Real.log (w₀ (x : K)) := by have h := congr_arg Real.log (prod_eq_abs_norm (x : K)) rw [show |(Algebra.norm ℚ) (x : K)| = 1 from isUnit_iff_norm.mp x.isUnit, Rat.cast_one, @@ -230,7 +231,7 @@ theorem logEmbedding_eq_zero_iff (x : (𝓞 K)ˣ) : · 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 [← logEmbedding_sum_component, Finset.sum_eq_zero] + rw [← sum_logEmbedding_component, Finset.sum_eq_zero] exact fun w _ => congrFun h w · exact mult_log_place_eq_zero.mp (congrFun h ⟨w, hw⟩) · ext w @@ -249,7 +250,7 @@ theorem log_le_of_logEmbedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : refine mul_le_mul ?_ le_rfl hx ?_ all_goals { rw [mult]; split_ifs <;> norm_num } by_cases hw : w = w₀ - · have hyp := congr_arg (‖·‖) (logEmbedding_sum_component x).symm + · have hyp := congr_arg (‖·‖) (sum_logEmbedding_component x).symm replace hyp := (le_of_eq hyp).trans (norm_sum_le _ _) simp_rw [norm_mul, norm_neg, Real.norm_eq_abs, Nat.abs_cast] at hyp refine (le_trans ?_ hyp).trans ?_ @@ -270,7 +271,8 @@ theorem log_le_of_logEmbedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : variable (K) /-- The lattice formed by the image of the logarithmic embedding. -/ -noncomputable def unitLattice : AddSubgroup ({w : InfinitePlace K // w ≠ w₀} → ℝ) := +noncomputable def _root_.NumberField.Units.unitLattice : + AddSubgroup ({w : InfinitePlace K // w ≠ w₀} → ℝ) := AddSubgroup.map (logEmbedding K) ⊤ theorem unitLattice_inter_ball_finite (r : ℝ) : @@ -295,8 +297,9 @@ theorem unitLattice_inter_ball_finite (r : ℝ) : exact ⟨h_int, h_le⟩ section span_top --- To prove that the span over `ℝ` of the `unit_lattice` is equal to the full space, we construct --- for each infinite place `w₁ ≠ w₀` an unit `u_w₁` of `K` such that, for all infinite places + +-- To prove that the span over `ℝ` of the `unitLattice` is equal to the full space, we construct +-- for each infinite place `w₁ ≠ w₀` a unit `u_w₁` of `K` such that, for all infinite places -- `w` such that `w ≠ w₁`, we have `Real.log w (u_w₁) < 0` (and thus `Real.log w₁ (u_w₁) > 0`). -- It follows then from a determinant computation (using `Matrix.det_ne_zero_of_neg`) that the -- image by `logEmbedding` of these units is a `ℝ`-linearly independent family. @@ -387,7 +390,7 @@ theorem seq_norm_bdd (n : ℕ) : rw [← Algebra.norm_localization ℤ (Sₘ := K) (nonZeroDivisors ℤ)] exact (seq_next K w₁ hB (seq K w₁ hB n).prop).choose_spec.2.2 -/-- Construct an unit associated to the place `w₁`. The family, for `w₁ ≠ w₀`, formed by the +/-- Construct a unit associated to the place `w₁`. The family, for `w₁ ≠ w₀`, formed by the image by the `logEmbedding` of these units is `ℝ`-linearly independent, see `unit_lattice_span_eq_top`. -/ theorem exists_unit (w₁ : InfinitePlace K ) : @@ -441,7 +444,7 @@ theorem unitLattice_span_eq_top : refine add_pos_of_nonneg_of_pos ?_ ?_ · rw [sub_nonneg] exact le_abs_self _ - · rw [logEmbedding_sum_component (exists_unit K w).choose] + · rw [sum_logEmbedding_component (exists_unit K w).choose] refine mul_pos_of_neg_of_neg ?_ ((exists_unit K w).choose_spec _ w.prop.symm) rw [mult]; split_ifs <;> norm_num · refine mul_neg_of_pos_of_neg ?_ ((exists_unit K w).choose_spec x ?_) @@ -455,7 +458,7 @@ def _root_.NumberField.Units.rank : ℕ := Fintype.card (InfinitePlace K) - 1 open FiniteDimensional -theorem unitLattice_discrete : DiscreteTopology (unitLattice K) := by +instance _root_.NumberField.Units.instDiscrete_unitLattice : DiscreteTopology (unitLattice K) := by refine discreteTopology_of_open_singleton_zero ?_ refine isOpen_singleton_of_finite_mem_nhds 0 (s := Metric.closedBall 0 1) ?_ ?_ · exact Metric.closedBall_mem_nhds _ (by norm_num) @@ -471,34 +474,35 @@ protected theorem finrank : simp only [finrank_fintype_fun_eq_card, Fintype.card_subtype_compl, Fintype.card_ofSubsingleton, rank] -theorem unitLattice_moduleFree : Module.Free ℤ (unitLattice K) := by - have := unitLattice_discrete K - exact Zlattice.module_free ℝ (unitLattice_span_eq_top K) +instance _root_.NumberField.Units.instModuleFree_unitLattice : Module.Free ℤ (unitLattice K) := + Zlattice.module_free ℝ (unitLattice_span_eq_top K) -theorem unitLattice_moduleFinite : Module.Finite ℤ (unitLattice K) := by - have := unitLattice_discrete K - exact Zlattice.module_finite ℝ (unitLattice_span_eq_top K) +instance _root_.NumberField.Units.instModuleFinite_unitLattice : Module.Finite ℤ (unitLattice K) := + Zlattice.module_finite ℝ (unitLattice_span_eq_top K) -theorem unitLattice_rank : finrank ℤ (unitLattice K) = Units.rank K := by - have := unitLattice_discrete K - rw [← dirichlet.finrank] +@[simp] +theorem _root_.NumberField.Units.unitLattice_rank : + finrank ℤ (unitLattice K) = Units.rank K := by + rw [← dirichletUnitTheorem.finrank] exact Zlattice.rank ℝ (unitLattice_span_eq_top K) -end dirichlet +end dirichletUnitTheorem variable [NumberField K] +open dirichletUnitTheorem + /-- A basis of the quotient `(𝓞 K)ˣ ⧸ (torsion K)` seen as an additive ℤ-module. -/ def basisModTorsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by - let f : (dirichlet.unitLattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (torsion K)) := by + let f : (unitLattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (torsion K)) := by refine AddEquiv.toIntLinearEquiv ?_ - rw [dirichlet.unitLattice, ← AddMonoidHom.range_eq_map (dirichlet.logEmbedding K)] - refine (QuotientAddGroup.quotientKerEquivRange (dirichlet.logEmbedding K)).symm.trans ?_ + rw [unitLattice, ← AddMonoidHom.range_eq_map (logEmbedding K)] + refine (QuotientAddGroup.quotientKerEquivRange (logEmbedding K)).symm.trans ?_ refine (QuotientAddGroup.quotientAddEquivOfEq ?_).trans (QuotientAddGroup.quotientKerEquivOfSurjective (MonoidHom.toAdditive (QuotientGroup.mk' (torsion K))) (fun x => ?_)) · ext - rw [AddMonoidHom.mem_ker, AddMonoidHom.mem_ker, dirichlet.logEmbedding_eq_zero_iff, + rw [AddMonoidHom.mem_ker, AddMonoidHom.mem_ker, logEmbedding_eq_zero_iff, MonoidHom.toAdditive_apply_apply, ofMul_eq_zero, QuotientGroup.mk'_apply, QuotientGroup.eq_one_iff] rfl @@ -507,12 +511,10 @@ def basisModTorsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsio QuotientGroup.out_eq'] rfl have : Module.Free ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := - (dirichlet.unitLattice_moduleFree K).of_equiv' f - have : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by - have := dirichlet.unitLattice_moduleFinite K - exact Module.Finite.equiv f + (instModuleFree_unitLattice K).of_equiv' f + have : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := Module.Finite.equiv f have : FiniteDimensional.finrank ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) = rank K := by - rw [← LinearEquiv.finrank_eq f, dirichlet.unitLattice_rank] + rw [← LinearEquiv.finrank_eq f, unitLattice_rank] refine Basis.reindex (Module.Free.chooseBasis ℤ _) (Fintype.equivOfCardEq ?_) rw [← FiniteDimensional.finrank_eq_card_chooseBasisIndex, this, Fintype.card_fin] From 872d1e0beb8f187a71fb6fbddd016c9cee7ba6dd Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 23 Sep 2023 14:04:44 +0200 Subject: [PATCH 62/67] Fix docstrings --- Mathlib/NumberTheory/NumberField/Units.lean | 25 ++++++++++++--------- 1 file changed, 14 insertions(+), 11 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 2d009430f24d6..92e67990dc7bb 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -18,20 +18,23 @@ field `K`. ## Main definitions -* `Units.rank`: the unit rank of the number field `K`. +* `NumberField.Units.rank`: the unit rank of the number field `K`. -* `Units.fundSystem`: a fundamental system of units of `K`. +* `NumberField.Units.fundSystem`: a fundamental system of units of `K`. -* `Units.basisModTorsion`: a `ℤ`-basis of `(𝓞 K)ˣ ⧸ (torsion K)` (as an additive `ℤ`-module). +* `NumberField.Units.basisModTorsion`: a `ℤ`-basis of `(𝓞 K)ˣ ⧸ (torsion K)` +as an additive `ℤ`-module. ## Main results -* `isUnit_iff_norm`: an algebraic integer `x : 𝓞 K` is a unit if and only if `|norm ℚ x| = 1`. +* `NumberField.isUnit_iff_norm`: an algebraic integer `x : 𝓞 K` is a unit if and only if `|norm ℚ x| = 1`. -* `Units.mem_torsion`: a unit `x : (𝓞 K)ˣ` is torsion iff `w x = 1` for all infinite places of `K`. +* `NumberField.Units.mem_torsion`: a unit `x : (𝓞 K)ˣ` is torsion iff `w x = 1` for all infinite +places of `K`. -* `Units.exist_unique_eq_mul_prod`: a unit `x : (𝓞 K)ˣ` can be written in a unique way as the -product of a root of unity times the product of powers of units of the fundamental system. +* `NumberField.Units.exist_unique_eq_mul_prod`: **Dirichlet Unit Theorem**. Any unit `x` of `𝓞 K` +can be written uniquely as the product of a root of unity and powers of the units of a +fundamental system. ## Tags number field, units @@ -60,11 +63,11 @@ section IsUnit variable {K} -theorem isUnit_iff_norm [NumberField K] {x : 𝓞 K} : +theorem NumberField.isUnit_iff_norm [NumberField K] {x : 𝓞 K} : IsUnit x ↔ |(RingOfIntegers.norm ℚ x : ℚ)| = 1 := by convert (RingOfIntegers.isUnit_norm ℚ (F := K)).symm rw [← abs_one, abs_eq_abs, ← Rat.RingOfIntegers.isUnit_iff] -#align is_unit_iff_norm isUnit_iff_norm +#align is_unit_iff_norm NumberField.isUnit_iff_norm end IsUnit @@ -535,8 +538,8 @@ theorem fun_eq_repr {x ζ : (𝓞 K)ˣ} {f : Fin (rank K) → ℤ} (hζ : ζ ∈ _ = ∑ i, (f i) • (basisModTorsion K i) := by simp_rw [fundSystem, QuotientGroup.out_eq', ofMul_toMul] -/-- Any unit `x` of `𝓞 K` can be written uniquely as a root of unity times the product of powers -of the units of the fundamental system. -/ +/-- **Dirichlet Unit Theorem**. Any unit `x` of `𝓞 K` can be written uniquely as the product of +a root of unity and powers of the units of a fundamental system. -/ theorem exist_unique_eq_mul_prod (x : (𝓞 K)ˣ) : ∃! (ζ : torsion K) (e : Fin (rank K) → ℤ), x = ζ * ∏ i, (fundSystem K i) ^ (e i) := by let ζ := x * (∏ i, (fundSystem K i) ^ ((basisModTorsion K).repr (Additive.ofMul ↑x) i))⁻¹ From ae68a0a7a946cd0a4356707e6972a07493bb535c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 23 Sep 2023 14:15:29 +0200 Subject: [PATCH 63/67] new section --- Mathlib/NumberTheory/NumberField/Units.lean | 41 ++++++++++++--------- 1 file changed, 23 insertions(+), 18 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 92e67990dc7bb..ee9c274dd82de 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -33,8 +33,8 @@ as an additive `ℤ`-module. places of `K`. * `NumberField.Units.exist_unique_eq_mul_prod`: **Dirichlet Unit Theorem**. Any unit `x` of `𝓞 K` -can be written uniquely as the product of a root of unity and powers of the units of a -fundamental system. +can be written uniquely as the product of a root of unity and powers of the units of of the +fundamental system `fundSystem`. ## Tags number field, units @@ -456,12 +456,18 @@ theorem unitLattice_span_eq_top : end span_top -/-- The unit rank of the number field `K`, it is equal to `card (InfinitePlace K) - 1`. -/ -def _root_.NumberField.Units.rank : ℕ := Fintype.card (InfinitePlace K) - 1 +end dirichletUnitTheorem -open FiniteDimensional +section statements + +variable [NumberField K] + +open dirichletUnitTheorem FiniteDimensional Classical + +/-- The unit rank of the number field `K`, it is equal to `card (InfinitePlace K) - 1`. -/ +def rank : ℕ := Fintype.card (InfinitePlace K) - 1 -instance _root_.NumberField.Units.instDiscrete_unitLattice : DiscreteTopology (unitLattice K) := by +instance instDiscrete_unitLattice : DiscreteTopology (unitLattice K) := by refine discreteTopology_of_open_singleton_zero ?_ refine isOpen_singleton_of_finite_mem_nhds 0 (s := Metric.closedBall 0 1) ?_ ?_ · exact Metric.closedBall_mem_nhds _ (by norm_num) @@ -472,29 +478,23 @@ instance _root_.NumberField.Units.instDiscrete_unitLattice : DiscreteTopology (u rintro ⟨x, hx, rfl⟩ exact ⟨Subtype.mem x, hx⟩ -protected theorem finrank : +protected theorem finrank_eq_rank : finrank ℝ ({w : InfinitePlace K // w ≠ w₀} → ℝ) = Units.rank K := by simp only [finrank_fintype_fun_eq_card, Fintype.card_subtype_compl, Fintype.card_ofSubsingleton, rank] -instance _root_.NumberField.Units.instModuleFree_unitLattice : Module.Free ℤ (unitLattice K) := +instance instModuleFree_unitLattice : Module.Free ℤ (unitLattice K) := Zlattice.module_free ℝ (unitLattice_span_eq_top K) -instance _root_.NumberField.Units.instModuleFinite_unitLattice : Module.Finite ℤ (unitLattice K) := +instance instModuleFinite_unitLattice : Module.Finite ℤ (unitLattice K) := Zlattice.module_finite ℝ (unitLattice_span_eq_top K) @[simp] -theorem _root_.NumberField.Units.unitLattice_rank : +theorem unitLattice_rank : finrank ℤ (unitLattice K) = Units.rank K := by - rw [← dirichletUnitTheorem.finrank] + rw [← Units.finrank_eq_rank] exact Zlattice.rank ℝ (unitLattice_span_eq_top K) -end dirichletUnitTheorem - -variable [NumberField K] - -open dirichletUnitTheorem - /-- A basis of the quotient `(𝓞 K)ˣ ⧸ (torsion K)` seen as an additive ℤ-module. -/ def basisModTorsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by let f : (unitLattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (torsion K)) := by @@ -526,6 +526,9 @@ units in `basisModTorsion`. -/ def fundSystem : Fin (rank K) → (𝓞 K)ˣ := fun i => Quotient.out' (Additive.toMul (basisModTorsion K i)) +/-- The exponents that appear in the unique decomposition of a unit as the product of +a root of unity and powers of the units of the fundamental system `fundSystem` (see +`exist_unique_eq_mul_prod`) are given by the representation of the unit of `basisModTorsion`. -/ theorem fun_eq_repr {x ζ : (𝓞 K)ˣ} {f : Fin (rank K) → ℤ} (hζ : ζ ∈ torsion K) (h : x = ζ * ∏ i, (fundSystem K i) ^ (f i)) : f = (basisModTorsion K).repr (Additive.ofMul ↑x) := by @@ -539,7 +542,7 @@ theorem fun_eq_repr {x ζ : (𝓞 K)ˣ} {f : Fin (rank K) → ℤ} (hζ : ζ ∈ simp_rw [fundSystem, QuotientGroup.out_eq', ofMul_toMul] /-- **Dirichlet Unit Theorem**. Any unit `x` of `𝓞 K` can be written uniquely as the product of -a root of unity and powers of the units of a fundamental system. -/ +a root of unity and powers of the units of the fundamental system `fundSystem`. -/ theorem exist_unique_eq_mul_prod (x : (𝓞 K)ˣ) : ∃! (ζ : torsion K) (e : Fin (rank K) → ℤ), x = ζ * ∏ i, (fundSystem K i) ^ (e i) := by let ζ := x * (∏ i, (fundSystem K i) ^ ((basisModTorsion K).repr (Additive.ofMul ↑x) i))⁻¹ @@ -560,4 +563,6 @@ theorem exist_unique_eq_mul_prod (x : (𝓞 K)ˣ) : ∃! (ζ : torsion K) (e : F nth_rewrite 1 [hf] rw [_root_.mul_inv_cancel_right] +end statements + end NumberField.Units From 1020c107958544d89dd01d7a9d62eeac6dbc039f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Sat, 23 Sep 2023 14:23:07 +0200 Subject: [PATCH 64/67] Long line --- Mathlib/NumberTheory/NumberField/Units.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index ee9c274dd82de..47a8844ec6871 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -27,7 +27,8 @@ as an additive `ℤ`-module. ## Main results -* `NumberField.isUnit_iff_norm`: an algebraic integer `x : 𝓞 K` is a unit if and only if `|norm ℚ x| = 1`. +* `NumberField.isUnit_iff_norm`: an algebraic integer `x : 𝓞 K` is a unit if and only if +`|norm ℚ x| = 1`. * `NumberField.Units.mem_torsion`: a unit `x : (𝓞 K)ˣ` is torsion iff `w x = 1` for all infinite places of `K`. From 78ce8ffde4828979b55b22949d07319b52b1d9a9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 27 Sep 2023 10:24:25 +0200 Subject: [PATCH 65/67] Review --- Mathlib/NumberTheory/NumberField/Units.lean | 181 +++++++++++--------- 1 file changed, 96 insertions(+), 85 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 47a8844ec6871..4956969b05e0c 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -31,7 +31,7 @@ as an additive `ℤ`-module. `|norm ℚ x| = 1`. * `NumberField.Units.mem_torsion`: a unit `x : (𝓞 K)ˣ` is torsion iff `w x = 1` for all infinite -places of `K`. +places `w` of `K`. * `NumberField.Units.exist_unique_eq_mul_prod`: **Dirichlet Unit Theorem**. Any unit `x` of `𝓞 K` can be written uniquely as the product of a root of unity and powers of the units of of the @@ -140,7 +140,7 @@ def torsionOrder [NumberField K] : ℕ+ := ⟨Fintype.card (torsion K), Fintype. /-- If `k` does not divide `torsionOrder` then there are no nontrivial roots of unity of order dividing `k`. -/ -theorem rootsOfUnity_eq_one [NumberField K] {k : ℕ+} (hc : Nat.Coprime k (torsionOrder K)) +theorem rootsOfUnity_eq_one [NumberField K] {k : ℕ+} (hc : Nat.Coprime k (torsionOrder K)) {ζ : (𝓞 K)ˣ} : ζ ∈ rootsOfUnity k (𝓞 K) ↔ ζ = 1 := by rw [mem_rootsOfUnity] refine ⟨fun h => ?_, fun h => by rw [h, one_pow]⟩ @@ -167,18 +167,21 @@ end torsion namespace dirichletUnitTheorem --- This section is devoted to the proof of Dirichlet's unit theorem --- We define a group morphism from `(𝓞 K)ˣ` to `{w : InfinitePlace K // w ≠ w₀} → ℝ` where `w₀` --- is a distinguished (arbitrary) infinite place, prove that its kernel is the torsion subgroup --- (see `logEmbedding_eq_zero_iff`) and that its image, called `unitLattice`, is a full --- `ℤ`-lattice. It follows that is a free `ℤ`-module (see `unitLattice_moduleFree `) of --- rank `card (InfinitePlaces K) - 1` (see `unitLattice_rank`). --- To prove that the `unitLattice` is a full `ℤ`-lattice, we need to prove that it is discrete --- (see `unitLattice_inter_ball_finite`) and that it spans the full space over `ℝ` --- (see ` unitLattice_span_eq_top`); this is the main part of the proof, see the section --- `span_top` below for more details. +/-! +### Dirichlet Unit Theorem +This section is devoted to the proof of Dirichlet's unit theorem. + +We define a group morphism from `(𝓞 K)ˣ` to `{w : InfinitePlace K // w ≠ w₀} → ℝ` where `w₀` is a +distinguished (arbitrary) infinite place, prove that its kernel is the torsion subgroup (see +`logEmbedding_eq_zero_iff`) and that its image, called `unitLattice`, is a full `ℤ`-lattice. It +follows that `unitLattice` is a free `ℤ`-module (see `unitLattice_moduleFree `) of rank +`card (InfinitePlaces K) - 1` (see `unitLattice_rank`). To prove that the `unitLattice` is a full +`ℤ`-lattice, we need to prove that it is discrete (see `unitLattice_inter_ball_finite`) and that it +spans the full space over `ℝ` (see ` unitLattice_span_eq_top`); this is the main part of the proof, +see the section `span_top` below for more details. +-/ -open scoped Classical +open Classical Finset variable [NumberField K] @@ -193,10 +196,7 @@ variable (K) def logEmbedding : Additive ((𝓞 K)ˣ) →+ ({w : InfinitePlace K // w ≠ w₀} → ℝ) := { toFun := fun x w => mult w.val * Real.log (w.val (Additive.toMul x)) map_zero' := by simp; rfl - map_add' := by - intro _ _ - simp [ne_eq, toMul_add, map_mul, map_eq_zero, coe_ne_zero, Real.log_mul, mul_add] - rfl } + map_add' := fun _ _ => by simp [Real.log_mul, mul_add]; rfl } variable {K} @@ -210,32 +210,31 @@ theorem sum_logEmbedding_component (x : (𝓞 K)ˣ) : rw [show |(Algebra.norm ℚ) (x : K)| = 1 from isUnit_iff_norm.mp x.isUnit, Rat.cast_one, Real.log_one, Real.log_prod] at h · simp_rw [Real.log_pow] at h - rw [← Finset.insert_erase (Finset.mem_univ w₀), Finset.sum_insert (Finset.not_mem_erase w₀ - Finset.univ), add_comm, add_eq_zero_iff_eq_neg] at h + 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 (Finset.sum_subtype _ (fun w => ?_) (fun w => (mult w) * (Real.log (w (x : K))))).symm - exact ⟨Finset.ne_of_mem_erase, fun h => Finset.mem_erase_of_ne_of_mem h (Finset.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)) 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] - · have : 0 ≤ w x := map_nonneg _ _ - linarith + · linarith [(map_nonneg _ _ : 0 ≤ w x)] · simp only [ne_eq, map_eq_zero, coe_ne_zero x] · refine (ne_of_gt ?_) rw [mult]; split_ifs <;> norm_num -theorem logEmbedding_eq_zero_iff (x : (𝓞 K)ˣ) : +theorem logEmbedding_eq_zero_iff {x : (𝓞 K)ˣ} : logEmbedding K x = 0 ↔ x ∈ torsion K := by rw [mem_torsion] refine ⟨fun h w => ?_, fun h => ?_⟩ · by_cases hw : w = w₀ - · suffices - mult w₀ * Real.log (w₀ (x : K)) = 0 by + · 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, Finset.sum_eq_zero] + rw [← sum_logEmbedding_component, sum_eq_zero] exact fun w _ => congrFun h w · exact mult_log_place_eq_zero.mp (congrFun h ⟨w, hw⟩) · ext w @@ -245,7 +244,7 @@ theorem logEmbedding_component_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : (w : {w : InfinitePlace K // w ≠ w₀}) : |logEmbedding K x w| ≤ r := by lift r to NNReal using hr simp_rw [Pi.norm_def, NNReal.coe_le_coe, Finset.sup_le_iff, ← NNReal.coe_le_coe] at h - exact h w (Finset.mem_univ _) + exact h w (mem_univ _) theorem log_le_of_logEmbedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : ‖logEmbedding K x‖ ≤ r) (w : InfinitePlace K) : |Real.log (w x)| ≤ (Fintype.card (InfinitePlace K)) * r := by @@ -260,11 +259,11 @@ theorem log_le_of_logEmbedding_le {r : ℝ} {x : (𝓞 K)ˣ} (hr : 0 ≤ r) (h : refine (le_trans ?_ hyp).trans ?_ · rw [← hw] exact tool _ (abs_nonneg _) - · refine (Finset.sum_le_card_nsmul Finset.univ _ _ + · refine (sum_le_card_nsmul univ _ _ (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 [Finset.card_univ] + simp [card_univ] · have hyp := logEmbedding_component_le hr h ⟨w, hw⟩ rw [logEmbedding_component, abs_mul, Nat.abs_cast] at hyp refine (le_trans ?_ hyp).trans ?_ @@ -286,8 +285,8 @@ theorem unitLattice_inter_ball_finite (r : ℝ) : · convert Set.finite_empty rw [Metric.closedBall_eq_empty.mpr hr] exact Set.inter_empty _ - · suffices Set.Finite {x : (𝓞 K)ˣ | IsIntegral ℤ (x : K) ∧ - ∀ (φ : K →+* ℂ), ‖φ x‖ ≤ Real.exp ((Fintype.card (InfinitePlace K)) * r)} by + · suffices {x : (𝓞 K)ˣ | IsIntegral ℤ (x : K) ∧ + ∀ (φ : K →+* ℂ), ‖φ x‖ ≤ Real.exp ((Fintype.card (InfinitePlace K)) * r)}.Finite by refine (Set.Finite.image (logEmbedding K) this).subset ?_ rintro _ ⟨⟨x, ⟨_, rfl⟩⟩, hx⟩ refine ⟨x, ⟨x.val.prop, (le_iff_le _ _).mp (fun w => (Real.log_le_iff_le_exp ?_).mp ?_)⟩, rfl⟩ @@ -302,15 +301,19 @@ theorem unitLattice_inter_ball_finite (r : ℝ) : section span_top --- To prove that the span over `ℝ` of the `unitLattice` is equal to the full space, we construct --- for each infinite place `w₁ ≠ w₀` a unit `u_w₁` of `K` such that, for all infinite places --- `w` such that `w ≠ w₁`, we have `Real.log w (u_w₁) < 0` (and thus `Real.log w₁ (u_w₁) > 0`). --- It follows then from a determinant computation (using `Matrix.det_ne_zero_of_neg`) that the --- image by `logEmbedding` of these units is a `ℝ`-linearly independent family. --- The unit `u_w₁` is obtained by construction a sequence `seq n` of nonzero algebraic integers --- that is strictly decreasing at infinite places distinct from `w₁` and of norm `≤ B`. Since --- there are finitely many ideals of norm `≤ B`, there exists two terms in the sequence defining --- the same ideal and their quotient is the desired unit `u_w₁` (see `exists_unit`). +/-! +## Section `span_top` + +In this section, we prove that the span over `ℝ` of the `unitLattice` is equal to the full space. +For this, we construct for each infinite place `w₁ ≠ w₀` a unit `u_w₁` of `K` such that, for all +infinite places `w` such that `w ≠ w₁`, we have `Real.log w (u_w₁) < 0` +(and thus `Real.log w₁ (u_w₁) > 0`). It follows then from a determinant computation +(using `Matrix.det_ne_zero_of_sum_col_lt_diag`) that the image by `logEmbedding` of these units is +a `ℝ`-linearly independent family. The unit `u_w₁` is obtained by constructing a sequence `seq n` +of nonzero algebraic integers that is strictly decreasing at infinite places distinct from `w₁` and +of norm `≤ B`. Since there are finitely many ideals of norm `≤ B`, there exists two term in the +sequence defining the same ideal and their quotient is the desired unit `u_w₁` (see `exists_unit`). +-/ open NumberField.mixedEmbedding NNReal @@ -333,7 +336,7 @@ theorem seq_next {x : 𝓞 K} (hx : x ≠ 0) : calc _ = ∏ w : InfinitePlace K, w y ^ mult w := (prod_eq_abs_norm (y : K)).symm _ ≤ ∏ w : InfinitePlace K, (g w : ℝ) ^ mult w := by - refine Finset.prod_le_prod ?_ ?_ + refine prod_le_prod ?_ ?_ · exact fun _ _ => pow_nonneg (by positivity) _ · exact fun w _ => pow_le_pow_of_le_left (by positivity) (le_of_lt (h_yle w)) (mult w) _ ≤ (B : ℝ) := by @@ -359,46 +362,46 @@ theorem seq_ne_zero (n : ℕ) : (seq K w₁ hB n : K) ≠ 0 := by refine (map_ne_zero_iff (algebraMap (𝓞 K) K) ?_).mpr (seq K w₁ hB n).prop exact IsFractionRing.injective { x // x ∈ 𝓞 K } K +/-- 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.ne_of_val_ne (seq_ne_zero K w₁ hB n)) + /-- The sequence is strictly decreasing at infinite places distinct from `w₁`. -/ -theorem seq_antitone {n m : ℕ} (h : n < m) : - ∀ w : InfinitePlace K, w ≠ w₁ → w (seq K w₁ hB m) < w (seq K w₁ hB n) := by +theorem seq_decreasing {n m : ℕ} (h : n < m) (w : InfinitePlace K) (hw : w ≠ w₁) : + w (seq K w₁ hB m) < w (seq K w₁ hB n) := by induction m with | zero => exfalso exact Nat.not_succ_le_zero n h | succ m m_ih => - intro w hw cases eq_or_lt_of_le (Nat.le_of_lt_succ h) with | inl hr => rw [hr] exact (seq_next K w₁ hB (seq K w₁ hB m).prop).choose_spec.2.1 w hw | inr hr => - refine lt_trans ?_ (m_ih hr w hw) + refine lt_trans ?_ (m_ih hr) exact (seq_next K w₁ hB (seq K w₁ hB m).prop).choose_spec.2.1 w hw /-- The terms of the sequence have norm bounded by `B`. -/ -theorem seq_norm_bdd (n : ℕ) : - 1 ≤ Int.natAbs (Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K)) ∧ - Int.natAbs (Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K)) ≤ B := by +theorem seq_norm_le (n : ℕ) : + Int.natAbs (Algebra.norm ℤ (seq K w₁ hB n : 𝓞 K)) ≤ B := by cases n with | zero => have : 1 ≤ B := by contrapose! hB simp only [Nat.lt_one_iff.mp hB, CharP.cast_eq_zero, mul_zero, zero_le] - simp only [ne_eq, seq, _root_.map_one, Int.natAbs_one, le_refl, this, and_self] + simp only [ne_eq, seq, map_one, Int.natAbs_one, this] | succ n => - refine ⟨Nat.succ_le_iff.mpr (Int.natAbs_pos.mpr ?_), ?_⟩ - · exact Algebra.norm_ne_zero_iff.mpr (seq K w₁ hB n.succ).prop - · rw [← Nat.cast_le (α := ℚ), Int.cast_natAbs, Int.cast_abs] - change |algebraMap ℤ ℚ _| ≤ _ - rw [← Algebra.norm_localization ℤ (Sₘ := K) (nonZeroDivisors ℤ)] - exact (seq_next K w₁ hB (seq K w₁ hB n).prop).choose_spec.2.2 + rw [← Nat.cast_le (α := ℚ), Int.cast_natAbs, Int.cast_abs] + change |algebraMap ℤ ℚ _| ≤ _ + rw [← Algebra.norm_localization ℤ (Sₘ := K) (nonZeroDivisors ℤ)] + exact (seq_next K w₁ hB (seq K w₁ hB n).prop).choose_spec.2.2 /-- Construct a unit associated to the place `w₁`. The family, for `w₁ ≠ w₀`, formed by the image by the `logEmbedding` of these units is `ℝ`-linearly independent, see `unit_lattice_span_eq_top`. -/ theorem exists_unit (w₁ : InfinitePlace K ) : - ∃ u : (𝓞 K)ˣ, (∀ w : InfinitePlace K, w ≠ w₁ → Real.log (w u) < 0) := by + ∃ u : (𝓞 K)ˣ, ∀ w : InfinitePlace K, w ≠ w₁ → Real.log (w u) < 0 := by obtain ⟨B, hB⟩ : ∃ B : ℕ, minkowskiBound K < (convexBodyLtFactor K) * B := by simp_rw [mul_comm] refine ENNReal.exists_nat_mul_gt ?_ ?_ @@ -416,12 +419,13 @@ theorem exists_unit (w₁ : InfinitePlace K ) : _ = w (seq K w₁ hB m) * w (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_antitone K w₁ hB hnm w hw + 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] - exact seq_norm_bdd K w₁ hB n + 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) @@ -443,8 +447,8 @@ theorem unitLattice_span_eq_top : -- We use a specific lemma to prove that this determinant is nonzero refine det_ne_zero_of_sum_col_lt_diag (fun w => ?_) simp_rw [Real.norm_eq_abs, Basis.coePiBasisFun.toMatrix_eq_transpose, Matrix.transpose_apply] - rw [← sub_pos, Finset.sum_congr rfl (fun x hx => abs_of_neg ?_), Finset.sum_neg_distrib, - sub_neg_eq_add, Finset.sum_erase_eq_sub (Finset.mem_univ _), ← add_comm_sub] + rw [← sub_pos, sum_congr rfl (fun x hx => abs_of_neg ?_), sum_neg_distrib, sub_neg_eq_add, + sum_erase_eq_sub (mem_univ _), ← add_comm_sub] refine add_pos_of_nonneg_of_pos ?_ ?_ · rw [sub_nonneg] exact le_abs_self _ @@ -453,7 +457,7 @@ theorem unitLattice_span_eq_top : rw [mult]; split_ifs <;> norm_num · refine mul_neg_of_pos_of_neg ?_ ((exists_unit K w).choose_spec x ?_) · rw [mult]; split_ifs <;> norm_num - · exact Subtype.ext_iff_val.not.mp (Finset.ne_of_mem_erase hx) + · exact Subtype.ext_iff_val.not.mp (ne_of_mem_erase hx) end span_top @@ -496,31 +500,39 @@ theorem unitLattice_rank : rw [← Units.finrank_eq_rank] exact Zlattice.rank ℝ (unitLattice_span_eq_top K) +/-- The linear equivalence between `unitLattice` and `(𝓞 K)ˣ ⧸ (torsion K)` as an additive +`ℤ`-module. -/ +def unitLatticeEquiv : (unitLattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (torsion K)) := by + refine AddEquiv.toIntLinearEquiv ?_ + rw [unitLattice, ← AddMonoidHom.range_eq_map (logEmbedding K)] + refine (QuotientAddGroup.quotientKerEquivRange (logEmbedding K)).symm.trans ?_ + refine (QuotientAddGroup.quotientAddEquivOfEq ?_).trans + (QuotientAddGroup.quotientKerEquivOfSurjective + (MonoidHom.toAdditive (QuotientGroup.mk' (torsion K))) (fun x => ?_)) + · ext + rw [AddMonoidHom.mem_ker, AddMonoidHom.mem_ker, logEmbedding_eq_zero_iff, + MonoidHom.toAdditive_apply_apply, ofMul_eq_zero, QuotientGroup.mk'_apply, + QuotientGroup.eq_one_iff] + rfl + · refine ⟨Additive.ofMul x.out', ?_⟩ + simp only [MonoidHom.toAdditive_apply_apply, toMul_ofMul, QuotientGroup.mk'_apply, + QuotientGroup.out_eq'] + rfl + +instance : Module.Free ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := + (instModuleFree_unitLattice K).of_equiv' (unitLatticeEquiv K) + +instance : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := + Module.Finite.equiv (unitLatticeEquiv K) + +theorem rank_modTorsion : + FiniteDimensional.finrank ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) = rank K := by + rw [← LinearEquiv.finrank_eq (unitLatticeEquiv K), unitLattice_rank] + /-- A basis of the quotient `(𝓞 K)ˣ ⧸ (torsion K)` seen as an additive ℤ-module. -/ def basisModTorsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by - let f : (unitLattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (torsion K)) := by - refine AddEquiv.toIntLinearEquiv ?_ - rw [unitLattice, ← AddMonoidHom.range_eq_map (logEmbedding K)] - refine (QuotientAddGroup.quotientKerEquivRange (logEmbedding K)).symm.trans ?_ - refine (QuotientAddGroup.quotientAddEquivOfEq ?_).trans - (QuotientAddGroup.quotientKerEquivOfSurjective - (MonoidHom.toAdditive (QuotientGroup.mk' (torsion K))) (fun x => ?_)) - · ext - rw [AddMonoidHom.mem_ker, AddMonoidHom.mem_ker, logEmbedding_eq_zero_iff, - MonoidHom.toAdditive_apply_apply, ofMul_eq_zero, QuotientGroup.mk'_apply, - QuotientGroup.eq_one_iff] - rfl - · refine ⟨Additive.ofMul x.out', ?_⟩ - simp only [MonoidHom.toAdditive_apply_apply, toMul_ofMul, QuotientGroup.mk'_apply, - QuotientGroup.out_eq'] - rfl - have : Module.Free ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := - (instModuleFree_unitLattice K).of_equiv' f - have : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := Module.Finite.equiv f - have : FiniteDimensional.finrank ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) = rank K := by - rw [← LinearEquiv.finrank_eq f, unitLattice_rank] refine Basis.reindex (Module.Free.chooseBasis ℤ _) (Fintype.equivOfCardEq ?_) - rw [← FiniteDimensional.finrank_eq_card_chooseBasisIndex, this, Fintype.card_fin] + rw [← FiniteDimensional.finrank_eq_card_chooseBasisIndex, rank_modTorsion, Fintype.card_fin] /-- A fundamental system of units of `K`. The units of `fundSystem` are arbitrary lifts of the units in `basisModTorsion`. -/ @@ -558,8 +570,7 @@ theorem exist_unique_eq_mul_prod (x : (𝓞 K)ˣ) : ∃! (ζ : torsion K) (e : F · simp only [_root_.inv_mul_cancel_right] · exact fun _ hf => fun_eq_repr K h_tors hf · rintro η ⟨_, hf, _⟩ - have f_eq := fun_eq_repr K η.prop hf - simp_rw [f_eq] at hf + simp_rw [fun_eq_repr K η.prop hf] at hf ext1; dsimp only nth_rewrite 1 [hf] rw [_root_.mul_inv_cancel_right] From 9b6a91499649a404f2bd152fae8844053fc41470 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 27 Sep 2023 10:31:26 +0200 Subject: [PATCH 66/67] whitespace --- Mathlib/NumberTheory/NumberField/Units.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index 4956969b05e0c..f36311fa4cfee 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -522,7 +522,7 @@ def unitLatticeEquiv : (unitLattice K) ≃ₗ[ℤ] Additive ((𝓞 K)ˣ ⧸ (tor instance : Module.Free ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := (instModuleFree_unitLattice K).of_equiv' (unitLatticeEquiv K) -instance : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := +instance : Module.Finite ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := Module.Finite.equiv (unitLatticeEquiv K) theorem rank_modTorsion : From 606d4113af853e66f409fbc1a97b2b93a54a30d7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Xavier-Fran=C3=A7ois=20Roblot?= <46200072+xroblot@users.noreply.github.com> Date: Wed, 27 Sep 2023 15:44:56 +0200 Subject: [PATCH 67/67] Review --- Mathlib/NumberTheory/NumberField/Units.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/NumberTheory/NumberField/Units.lean b/Mathlib/NumberTheory/NumberField/Units.lean index f36311fa4cfee..76a6fa59d3b8c 100644 --- a/Mathlib/NumberTheory/NumberField/Units.lean +++ b/Mathlib/NumberTheory/NumberField/Units.lean @@ -302,7 +302,7 @@ theorem unitLattice_inter_ball_finite (r : ℝ) : section span_top /-! -## Section `span_top` +#### Section `span_top` In this section, we prove that the span over `ℝ` of the `unitLattice` is equal to the full space. For this, we construct for each infinite place `w₁ ≠ w₀` a unit `u_w₁` of `K` such that, for all @@ -530,9 +530,9 @@ theorem rank_modTorsion : rw [← LinearEquiv.finrank_eq (unitLatticeEquiv K), unitLattice_rank] /-- A basis of the quotient `(𝓞 K)ˣ ⧸ (torsion K)` seen as an additive ℤ-module. -/ -def basisModTorsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := by - refine Basis.reindex (Module.Free.chooseBasis ℤ _) (Fintype.equivOfCardEq ?_) - rw [← FiniteDimensional.finrank_eq_card_chooseBasisIndex, rank_modTorsion, Fintype.card_fin] +def basisModTorsion : Basis (Fin (rank K)) ℤ (Additive ((𝓞 K)ˣ ⧸ (torsion K))) := + Basis.reindex (Module.Free.chooseBasis ℤ _) (Fintype.equivOfCardEq <| by + rw [← FiniteDimensional.finrank_eq_card_chooseBasisIndex, rank_modTorsion, Fintype.card_fin]) /-- A fundamental system of units of `K`. The units of `fundSystem` are arbitrary lifts of the units in `basisModTorsion`. -/