Skip to content

Commit 63ba377

Browse files
bump
1 parent 33953c1 commit 63ba377

File tree

7 files changed

+15
-229
lines changed

7 files changed

+15
-229
lines changed

FltRegular.lean

-2
Original file line numberDiff line numberDiff line change
@@ -6,7 +6,6 @@ import FltRegular.CaseII.Statement
66
import FltRegular.FLT5
77
import FltRegular.FltRegular
88
import FltRegular.MayAssume.Lemmas
9-
import FltRegular.NumberTheory.AuxLemmas
109
import FltRegular.NumberTheory.Cyclotomic.CaseI
1110
import FltRegular.NumberTheory.Cyclotomic.CyclRat
1211
import FltRegular.NumberTheory.Cyclotomic.CyclotomicUnits
@@ -15,7 +14,6 @@ import FltRegular.NumberTheory.Cyclotomic.GaloisActionOnCyclo
1514
import FltRegular.NumberTheory.Cyclotomic.MoreLemmas
1615
import FltRegular.NumberTheory.Cyclotomic.UnitLemmas
1716
import FltRegular.NumberTheory.CyclotomicRing
18-
import FltRegular.NumberTheory.Different
1917
import FltRegular.NumberTheory.Finrank
2018
import FltRegular.NumberTheory.GaloisPrime
2119
import FltRegular.NumberTheory.Hilbert90

FltRegular/NumberTheory/AuxLemmas.lean

-110
This file was deleted.

FltRegular/NumberTheory/Different.lean

-104
This file was deleted.

FltRegular/NumberTheory/GaloisPrime.lean

+5-5
Original file line numberDiff line numberDiff line change
@@ -1,11 +1,12 @@
1-
import FltRegular.NumberTheory.AuxLemmas
21
import Mathlib.RingTheory.IntegralClosure.IntegralRestrict
32
import Mathlib.Data.Set.Card
43
import Mathlib.FieldTheory.PurelyInseparable
54
import Mathlib.Order.CompletePartialOrder
65
import Mathlib.RingTheory.DedekindDomain.Dvr
76
import Mathlib.Algebra.Order.Star.Basic
87
import Mathlib.RingTheory.SimpleRing.Basic
8+
import Mathlib.NumberTheory.RamificationInertia
9+
import Mathlib.Algebra.Lie.OfAssociative
910

1011
/-!
1112
# Galois action on primes
@@ -237,15 +238,14 @@ lemma Ideal.ramificationIdxIn_bot : (⊥ : Ideal R).ramificationIdxIn S = 0 := b
237238
lemma Ideal.inertiaDegIn_bot [Nontrivial R] [IsDomain S] [NoZeroSMulDivisors R S] [IsNoetherian R S]
238239
[Algebra.IsIntegral R S] [H : (⊥ : Ideal R).IsMaximal] :
239240
(⊥ : Ideal R).inertiaDegIn S = Module.finrank R S := by
240-
delta inertiaDegIn
241+
dsimp [inertiaDegIn]
241242
rw [primesOver_bot]
242243
have : ({⊥} : Set (Ideal S)).Nonempty := by simp
243244
rw [dif_pos this, this.choose_spec]
244245
have hR := not_imp_not.mp (Ring.ne_bot_of_isMaximal_of_not_isField H) rfl
245246
have hS := isField_of_isIntegral_of_isField' (S := S) hR
246247
letI : Field R := hR.toField
247248
letI : Field S := hS.toField
248-
have : IsIntegralClosure S R S := isIntegralClosure_self
249249
rw [← Ideal.map_bot (f := algebraMap R S), ← finrank_quotient_map (R := R) (S := S) ⊥ R S]
250250
exact inertiaDeg_algebraMap _ _
251251

@@ -261,7 +261,7 @@ lemma Ideal.ramificationIdxIn_eq_ramificationIdx [IsGalois K L] (p : Ideal R) (P
261261
rw [dif_pos this]
262262
have ⟨σ, hσ⟩ := exists_comap_galRestrict_eq R K L S hP this.choose_spec
263263
rw [← hσ]
264-
exact Ideal.ramificationIdx_comap_eq (galRestrict R K L S σ) p P
264+
exact Ideal.ramificationIdx_comap_eq (galRestrict R K L S σ) P
265265

266266
lemma Ideal.inertiaDegIn_eq_inertiaDeg [IsGalois K L] (p : Ideal R) (P : Ideal S)
267267
(hP : P ∈ primesOver S p) [p.IsMaximal] :
@@ -271,7 +271,7 @@ lemma Ideal.inertiaDegIn_eq_inertiaDeg [IsGalois K L] (p : Ideal R) (P : Ideal S
271271
rw [dif_pos this]
272272
have ⟨σ, hσ⟩ := exists_comap_galRestrict_eq R K L S hP this.choose_spec
273273
rw [← hσ]
274-
exact Ideal.inertiaDeg_comap_eq (galRestrict R K L S σ) p P
274+
exact Ideal.inertiaDeg_comap_eq p (galRestrict R K L S σ) P
275275

276276
open Module
277277

FltRegular/NumberTheory/KummersLemma/Field.lean

+4-3
Original file line numberDiff line numberDiff line change
@@ -101,9 +101,10 @@ lemma map_poly : (poly hp hζ u hcong).map (algebraMap (𝓞 K) K) =
101101
include hu in
102102
lemma irreducible_map_poly :
103103
Irreducible ((poly hp hζ u hcong).map (algebraMap (𝓞 K) K)) := by
104-
rw [map_poly, ← irreducible_taylor_iff (r := 1 / (ζ - 1))]
105-
simp only [taylor, one_div, map_add, LinearMap.coe_mk, AddHom.coe_mk, pow_comp, sub_comp,
106-
X_comp, C_comp, add_sub_cancel_right]
104+
rw [map_poly]
105+
refine Irreducible.of_map (f := algEquivAevalXAddC (1 / (ζ - 1))) ?_
106+
simp only [one_div, map_add, algEquivAevalXAddC_apply, map_pow, map_sub, aeval_X, aeval_C,
107+
algebraMap_eq, add_sub_cancel_right]
107108
rw [← sub_neg_eq_add, ← (C : K →+* _).map_neg]
108109
apply X_pow_sub_C_irreducible_of_prime hpri.out
109110
intro b hb

FltRegular/NumberTheory/Unramified.lean

+4-3
Original file line numberDiff line numberDiff line change
@@ -1,6 +1,7 @@
11
import FltRegular.NumberTheory.GaloisPrime
22
import Mathlib.NumberTheory.KummerDedekind
3-
import FltRegular.NumberTheory.Different
3+
import Mathlib.RingTheory.DedekindDomain.Different
4+
45
/-!
56
# Unramified extensions
67
@@ -166,8 +167,8 @@ lemma isUnramifiedAt_of_Separable_minpoly' [Algebra.IsSeparable K L]
166167
have hxP : aeval x (derivative (minpoly R x)) ∈ P := by
167168
have : differentIdeal R S ≤ P := by
168169
rw [← Ideal.dvd_iff_le]
169-
exact (dvd_pow_self _ H).trans (pow_sub_one_dvd_differentIdeal R S P _ hpbot
170-
(Ideal.dvd_iff_le.mpr <| Ideal.le_pow_ramificationIdx (p := p) (f := algebraMap R S)))
170+
exact (dvd_pow_self _ H).trans (pow_sub_one_dvd_differentIdeal R P _ hpbot <|
171+
Ideal.dvd_iff_le.mpr <| Ideal.le_pow_ramificationIdx)
171172
exact this (aeval_derivative_mem_differentIdeal R K L _ hx')
172173
rw [← Ideal.Quotient.eq_zero_iff_mem, ← Ideal.Quotient.algebraMap_eq] at hxP
173174

lake-manifest.json

+2-2
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,7 @@
55
"type": "git",
66
"subDir": null,
77
"scope": "leanprover-community",
8-
"rev": "dc72dcdb8e97b3c56bd70f06f043ed2dee3258e6",
8+
"rev": "eb6c831c05bc6ca6d37454ca97531a9b780dfcb5",
99
"name": "batteries",
1010
"manifestFile": "lake-manifest.json",
1111
"inputRev": "main",
@@ -75,7 +75,7 @@
7575
"type": "git",
7676
"subDir": null,
7777
"scope": "",
78-
"rev": "9d502188e7df20d93aded2b312b28259ff1da308",
78+
"rev": "d27f22898a74975360dd1387155c6031b501a90b",
7979
"name": "mathlib",
8080
"manifestFile": "lake-manifest.json",
8181
"inputRev": null,

0 commit comments

Comments
 (0)