1
- import FltRegular.NumberTheory.AuxLemmas
2
1
import Mathlib.RingTheory.IntegralClosure.IntegralRestrict
3
2
import Mathlib.Data.Set.Card
4
3
import Mathlib.FieldTheory.PurelyInseparable
5
4
import Mathlib.Order.CompletePartialOrder
6
5
import Mathlib.RingTheory.DedekindDomain.Dvr
7
6
import Mathlib.Algebra.Order.Star.Basic
8
7
import Mathlib.RingTheory.SimpleRing.Basic
8
+ import Mathlib.NumberTheory.RamificationInertia
9
+ import Mathlib.Algebra.Lie.OfAssociative
9
10
10
11
/-!
11
12
# Galois action on primes
@@ -237,15 +238,14 @@ lemma Ideal.ramificationIdxIn_bot : (⊥ : Ideal R).ramificationIdxIn S = 0 := b
237
238
lemma Ideal.inertiaDegIn_bot [Nontrivial R] [IsDomain S] [NoZeroSMulDivisors R S] [IsNoetherian R S]
238
239
[Algebra.IsIntegral R S] [H : (⊥ : Ideal R).IsMaximal] :
239
240
(⊥ : Ideal R).inertiaDegIn S = Module.finrank R S := by
240
- delta inertiaDegIn
241
+ dsimp [ inertiaDegIn]
241
242
rw [primesOver_bot]
242
243
have : ({⊥} : Set (Ideal S)).Nonempty := by simp
243
244
rw [dif_pos this, this.choose_spec]
244
245
have hR := not_imp_not.mp (Ring.ne_bot_of_isMaximal_of_not_isField H) rfl
245
246
have hS := isField_of_isIntegral_of_isField' (S := S) hR
246
247
letI : Field R := hR.toField
247
248
letI : Field S := hS.toField
248
- have : IsIntegralClosure S R S := isIntegralClosure_self
249
249
rw [← Ideal.map_bot (f := algebraMap R S), ← finrank_quotient_map (R := R) (S := S) ⊥ R S]
250
250
exact inertiaDeg_algebraMap _ _
251
251
@@ -261,7 +261,7 @@ lemma Ideal.ramificationIdxIn_eq_ramificationIdx [IsGalois K L] (p : Ideal R) (P
261
261
rw [dif_pos this]
262
262
have ⟨σ, hσ⟩ := exists_comap_galRestrict_eq R K L S hP this.choose_spec
263
263
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
265
265
266
266
lemma Ideal.inertiaDegIn_eq_inertiaDeg [IsGalois K L] (p : Ideal R) (P : Ideal S)
267
267
(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
271
271
rw [dif_pos this]
272
272
have ⟨σ, hσ⟩ := exists_comap_galRestrict_eq R K L S hP this.choose_spec
273
273
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
275
275
276
276
open Module
277
277
0 commit comments