Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

feat: generalize Ideal.spanNorm to allow non free extensions #19244

Open
wants to merge 12 commits into
base: master
Choose a base branch
from
5 changes: 5 additions & 0 deletions Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean
Original file line number Diff line number Diff line change
Expand Up @@ -253,4 +253,9 @@ instance integralClosure.isDedekindDomain_fractionRing [IsDedekindDomain A] :
IsDedekindDomain (integralClosure A L) :=
integralClosure.isDedekindDomain A (FractionRing A) L

attribute [local instance] FractionRing.liftAlgebra in
instance [NoZeroSMulDivisors A C] [Module.Finite A C] [IsIntegrallyClosed C] :
IsLocalization (Algebra.algebraMapSubmonoid C A⁰) (FractionRing C) :=
IsIntegralClosure.isLocalization _ (FractionRing A) _ _

end IsIntegralClosure
178 changes: 123 additions & 55 deletions Mathlib/RingTheory/Ideal/Norm/RelNorm.lean
Original file line number Diff line number Diff line change
Expand Up @@ -6,6 +6,7 @@
import Mathlib.LinearAlgebra.FreeModule.PID
import Mathlib.RingTheory.DedekindDomain.PID
import Mathlib.RingTheory.Localization.NormTrace
import Mathlib.RingTheory.IntegralClosure.IntegralRestrict

/-!

Expand Down Expand Up @@ -34,44 +35,53 @@

open Submodule

variable (R : Type*) [CommRing R] {S : Type*} [CommRing S] [Algebra R S]
attribute [local instance] FractionRing.liftAlgebra

/-- `Ideal.spanNorm R (I : Ideal S)` is the ideal generated by mapping `Algebra.norm R` over `I`.
variable (R S : Type*) [CommRing R] [IsDomain R] {S : Type*} [CommRing S] [IsDomain S]
variable [IsIntegrallyClosed R] [IsIntegrallyClosed S] [Algebra R S] [Module.Finite R S]
variable [NoZeroSMulDivisors R S] [Algebra.IsSeparable (FractionRing R) (FractionRing S)]

/-- `Ideal.spanNorm R (I : Ideal S)` is the ideal generated by mapping `Algebra.intNorm R S`
over `I`.

See also `Ideal.relNorm`.
-/
def spanNorm (I : Ideal S) : Ideal R :=
Ideal.span (Algebra.norm R '' (I : Set S))
Ideal.span (Algebra.intNorm R S '' (I : Set S))

@[simp]
theorem spanNorm_bot [Nontrivial S] [Module.Free R S] [Module.Finite R S] :
theorem spanNorm_bot :
spanNorm R (⊥ : Ideal S) = ⊥ := span_eq_bot.mpr fun x hx => by simpa using hx

variable {R}

@[simp]
theorem spanNorm_eq_bot_iff [IsDomain R] [IsDomain S] [Module.Free R S] [Module.Finite R S]
{I : Ideal S} : spanNorm R I = ⊥ ↔ I = ⊥ := by
simp only [spanNorm, Ideal.span_eq_bot, Set.mem_image, SetLike.mem_coe, forall_exists_index,
and_imp, forall_apply_eq_imp_iff₂,
Algebra.norm_eq_zero_iff_of_basis (Module.Free.chooseBasis R S), @eq_bot_iff _ _ _ I,
SetLike.le_def]
theorem spanNorm_eq_bot_iff {I : Ideal S} : spanNorm R I = ⊥ ↔ I = ⊥ := by
simp only [spanNorm, span_eq_bot, Set.mem_image, SetLike.mem_coe, forall_exists_index, and_imp,
forall_apply_eq_imp_iff₂, Algebra.intNorm_eq_zero, @eq_bot_iff _ _ _ I, SetLike.le_def]
rfl

variable (R)

theorem norm_mem_spanNorm {I : Ideal S} (x : S) (hx : x ∈ I) : Algebra.norm R x ∈ I.spanNorm R :=
theorem intNorm_mem_spanNorm {I : Ideal S} (x : S) (hx : x ∈ I) :
Algebra.intNorm R S x ∈ I.spanNorm R :=
subset_span (Set.mem_image_of_mem _ hx)

theorem norm_mem_spanNorm [Module.Free R S] {I : Ideal S} (x : S) (hx : x ∈ I) :
Algebra.norm R x ∈ I.spanNorm R := by
refine subset_span ⟨x, hx, ?_⟩
rw [Algebra.intNorm_eq_norm]

@[simp]
theorem spanNorm_singleton {r : S} : spanNorm R (span ({r} : Set S)) = span {Algebra.norm R r} :=
theorem spanNorm_singleton {r : S} :
spanNorm R (span ({r} : Set S)) = span {Algebra.intNorm R S r} :=
le_antisymm
(span_le.mpr fun x hx =>
mem_span_singleton.mpr
(by
obtain ⟨x, hx', rfl⟩ := (Set.mem_image _ _ _).mp hx
exact map_dvd _ (mem_span_singleton.mp hx')))
((span_singleton_le_iff_mem _).mpr (norm_mem_spanNorm _ _ (mem_span_singleton_self _)))
((span_singleton_le_iff_mem _).mpr (intNorm_mem_spanNorm _ _ (mem_span_singleton_self _)))

@[simp]
theorem spanNorm_top : spanNorm R (⊤ : Ideal S) = ⊤ := by
Expand All @@ -80,8 +90,8 @@
rw [← Ideal.span_singleton_one, spanNorm_singleton]
simp

theorem map_spanNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) :
map f (spanNorm R I) = span (f ∘ Algebra.norm R '' (I : Set S)) := by
theorem map_spanIntNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) :
map f (spanNorm R I) = span (f ∘ Algebra.intNorm R S '' (I : Set S)) := by
rw [spanNorm, map_span, Set.image_image]
-- Porting note: `Function.comp` reducibility
rfl
Expand All @@ -90,31 +100,66 @@
theorem spanNorm_mono {I J : Ideal S} (h : I ≤ J) : spanNorm R I ≤ spanNorm R J :=
Ideal.span_mono (Set.monotone_image h)

theorem spanNorm_localization (I : Ideal S) [Module.Finite R S] [Module.Free R S] (M : Submonoid R)
theorem spanIntNorm_localization (I : Ideal S) (M : Submonoid R) (hM : M ≤ R⁰)
{Rₘ : Type*} (Sₘ : Type*) [CommRing Rₘ] [Algebra R Rₘ] [CommRing Sₘ] [Algebra S Sₘ]
[Algebra Rₘ Sₘ] [Algebra R Sₘ] [IsScalarTower R Rₘ Sₘ] [IsScalarTower R S Sₘ]
[IsLocalization M Rₘ] [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ] :
[IsLocalization M Rₘ] [IsLocalization (Algebra.algebraMapSubmonoid S M) Sₘ]
[IsIntegrallyClosed Rₘ] [IsDomain Rₘ] [IsDomain Sₘ] [NoZeroSMulDivisors Rₘ Sₘ]
[Module.Finite Rₘ Sₘ] [IsIntegrallyClosed Sₘ]
[Algebra.IsSeparable (FractionRing Rₘ) (FractionRing Sₘ)] :
spanNorm Rₘ (I.map (algebraMap S Sₘ)) = (spanNorm R I).map (algebraMap R Rₘ) := by
cases subsingleton_or_nontrivial R
· haveI := IsLocalization.unique R Rₘ M
simp [eq_iff_true_of_subsingleton]
let b := Module.Free.chooseBasis R S
rw [map_spanNorm]
let K := FractionRing R
let f : Rₘ →+* K := IsLocalization.map _ (T := R⁰) (RingHom.id R) hM
let _ := f.toAlgebra
have : IsScalarTower R Rₘ K := IsScalarTower.of_algebraMap_eq'
(by rw [RingHom.algebraMap_toAlgebra, IsLocalization.map_comp, RingHomCompTriple.comp_eq])
let _ := IsFractionRing.isFractionRing_of_isDomain_of_isLocalization M Rₘ K
let L := FractionRing S
let g : Sₘ →+* L := IsLocalization.map _ (M := Algebra.algebraMapSubmonoid S M) (T := S⁰)
(RingHom.id S) (Submonoid.map_le_of_le_comap _ <| hM.trans
(nonZeroDivisors_le_comap_nonZeroDivisors_of_injective _
(NoZeroSMulDivisors.algebraMap_injective _ _)))
letI := g.toAlgebra
have : IsScalarTower S Sₘ L := IsScalarTower.of_algebraMap_eq'
(by rw [RingHom.algebraMap_toAlgebra, IsLocalization.map_comp, RingHomCompTriple.comp_eq])
letI := ((algebraMap K L).comp f).toAlgebra
haveI : IsScalarTower Rₘ K L := IsScalarTower.of_algebraMap_eq' rfl
haveI : IsScalarTower Rₘ Sₘ L := by
apply IsScalarTower.of_algebraMap_eq'
apply IsLocalization.ringHom_ext M
rw [RingHom.algebraMap_toAlgebra, RingHom.algebraMap_toAlgebra (R := Sₘ), RingHom.comp_assoc,
RingHom.comp_assoc, ← IsScalarTower.algebraMap_eq, IsScalarTower.algebraMap_eq R S Sₘ,
IsLocalization.map_comp, RingHom.comp_id, ← RingHom.comp_assoc, IsLocalization.map_comp,
RingHom.comp_id, ← IsScalarTower.algebraMap_eq, ← IsScalarTower.algebraMap_eq]
letI := IsFractionRing.isFractionRing_of_isDomain_of_isLocalization
(Algebra.algebraMapSubmonoid S M) Sₘ L
haveI : IsIntegralClosure Sₘ Rₘ L :=
IsIntegralClosure.of_isIntegrallyClosed _ _ _
rw [map_spanIntNorm]
refine span_eq_span (Set.image_subset_iff.mpr ?_) (Set.image_subset_iff.mpr ?_)
· rintro a' ha'
simp only [Set.mem_preimage, submodule_span_eq, ← map_spanNorm, SetLike.mem_coe,
simp only [Set.mem_preimage, submodule_span_eq, ← map_spanIntNorm, SetLike.mem_coe,
IsLocalization.mem_map_algebraMap_iff (Algebra.algebraMapSubmonoid S M) Sₘ,
IsLocalization.mem_map_algebraMap_iff M Rₘ, Prod.exists] at ha' ⊢
obtain ⟨⟨a, ha⟩, ⟨_, ⟨s, hs, rfl⟩⟩, has⟩ := ha'
refine ⟨⟨Algebra.norm R a, norm_mem_spanNorm _ _ ha⟩,
⟨s ^ Fintype.card (Module.Free.ChooseBasisIndex R S), pow_mem hs _⟩, ?_⟩
refine ⟨⟨Algebra.intNorm R S a, intNorm_mem_spanNorm _ _ ha⟩,
⟨s ^ Module.finrank K L, pow_mem hs _⟩, ?_⟩
simp only [Submodule.coe_mk, Subtype.coe_mk, map_pow] at has ⊢
apply_fun Algebra.norm Rₘ at has
rwa [_root_.map_mul, ← IsScalarTower.algebraMap_apply, IsScalarTower.algebraMap_apply R Rₘ,
Algebra.norm_algebraMap_of_basis (b.localizationLocalization Rₘ M Sₘ),
Algebra.norm_localization R M a] at has
apply_fun algebraMap _ L at has
apply_fun Algebra.norm K at has
simp only [_root_.map_mul, IsScalarTower.algebraMap_apply R Rₘ Sₘ] at has
rw [← IsScalarTower.algebraMap_apply, ← IsScalarTower.algebraMap_apply,
← IsScalarTower.algebraMap_apply,
IsScalarTower.algebraMap_apply R K L,
Algebra.norm_algebraMap] at has
apply IsFractionRing.injective Rₘ K
simp only [_root_.map_mul, map_pow]
have : FiniteDimensional K L := Module.Finite_of_isLocalization R S _ _ R⁰
rwa [Algebra.algebraMap_intNorm (L := L), ← IsScalarTower.algebraMap_apply,
← IsScalarTower.algebraMap_apply, Algebra.algebraMap_intNorm (L := L)]
· intro a ha
rw [Set.mem_preimage, Function.comp_apply, ← Algebra.norm_localization (Sₘ := Sₘ) R M a]
rw [Set.mem_preimage, Function.comp_apply, Algebra.intNorm_eq_of_isLocalization
(A := R) (B := S) M (Aₘ := Rₘ) (Bₘ := Sₘ)]
exact subset_span (Set.mem_image_of_mem _ (mem_map_of_mem _ ha))

theorem spanNorm_mul_spanNorm_le (I J : Ideal S) :
Expand All @@ -124,10 +169,10 @@
rintro _ ⟨x, hxI, y, hyJ, rfl⟩
exact Ideal.mul_mem_mul hxI hyJ

/-- This condition `eq_bot_or_top` is equivalent to being a field.

Check failure on line 172 in Mathlib/RingTheory/Ideal/Norm/RelNorm.lean

View workflow job for this annotation

GitHub Actions / Build

Ideal.spanNorm_mul_of_bot_or_top argument 13 inst✝ : Module.Free
However, `Ideal.spanNorm_mul_of_field` is harder to apply since we'd need to upgrade a `CommRing R`
instance to a `Field R` instance. -/
theorem spanNorm_mul_of_bot_or_top [IsDomain R] [IsDomain S] [Module.Free R S] [Module.Finite R S]
theorem spanNorm_mul_of_bot_or_top [Module.Free R S]
(eq_bot_or_top : ∀ I : Ideal R, I = ⊥ ∨ I = ⊤) (I J : Ideal S) :
spanNorm R (I * J) = spanNorm R I * spanNorm R J := by
refine le_antisymm ?_ (spanNorm_mul_spanNorm_le R _ _)
Expand All @@ -140,47 +185,69 @@
rw [hJ]
exact le_top

@[simp]
theorem spanNorm_mul_of_field {K : Type*} [Field K] [Algebra K S] [IsDomain S] [Module.Finite K S]
(I J : Ideal S) : spanNorm K (I * J) = spanNorm K I * spanNorm K J :=
spanNorm_mul_of_bot_or_top K eq_bot_or_top I J

variable [IsDomain R] [IsDomain S] [IsDedekindDomain R] [IsDedekindDomain S]
variable [Module.Finite R S] [Module.Free R S]
variable [IsDedekindDomain R] [IsDedekindDomain S] [Module.Free R S]

/-- Multiplicativity of `Ideal.spanNorm`. simp-normal form is `map_mul (Ideal.relNorm R)`. -/
theorem spanNorm_mul (I J : Ideal S) : spanNorm R (I * J) = spanNorm R I * spanNorm R J := by
nontriviality R
cases subsingleton_or_nontrivial S
· have : ∀ I : Ideal S, I = ⊤ := fun I => Subsingleton.elim I ⊤
· have : ∀ I : Ideal S, I = ⊤ := fun I Subsingleton.elim I ⊤
simp [this I, this J, this (I * J)]
refine eq_of_localization_maximal ?_
intro P hP
by_cases hP0 : P = ⊥
· subst hP0
rw [spanNorm_mul_of_bot_or_top]
intro I
refine or_iff_not_imp_right.mpr fun hI => ?_
refine or_iff_not_imp_right.mpr fun hI ?_
exact (hP.eq_of_le hI bot_le).symm
let P' := Algebra.algebraMapSubmonoid S P.primeCompl
letI : Algebra (Localization.AtPrime P) (Localization P') := localizationAlgebra P.primeCompl S
haveI : IsScalarTower R (Localization.AtPrime P) (Localization P') :=
let Rₚ := Localization.AtPrime P
let Sₚ := Localization P'
let _ : Algebra Rₚ Sₚ := localizationAlgebra P.primeCompl S
have : IsScalarTower R Rₚ Sₚ :=
IsScalarTower.of_algebraMap_eq (fun x =>
(IsLocalization.map_eq (T := P') (Q := Localization P') P.primeCompl.le_comap_map x).symm)
have h : P' ≤ S⁰ :=
map_le_nonZeroDivisors_of_injective _ (NoZeroSMulDivisors.algebraMap_injective _ _)
P.primeCompl_le_nonZeroDivisors
haveI : IsDomain (Localization P') := IsLocalization.isDomain_localization h
haveI : IsDedekindDomain (Localization P') := IsLocalization.isDedekindDomain S h _
letI := Classical.decEq (Ideal (Localization P'))
haveI : IsPrincipalIdealRing (Localization P') :=
have : IsDomain Sₚ := IsLocalization.isDomain_localization h
have : IsDedekindDomain Sₚ := IsLocalization.isDedekindDomain S h _
have : IsPrincipalIdealRing Sₚ :=
Copy link
Member Author

@riccardobrasca riccardobrasca Nov 19, 2024

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

This sequence of have is not very nice. I am open to any suggestion.

IsDedekindDomain.isPrincipalIdealRing_localization_over_prime S P hP0
rw [Ideal.map_mul, ← spanNorm_localization R I P.primeCompl (Localization P'),
← spanNorm_localization R J P.primeCompl (Localization P'),
← spanNorm_localization R (I * J) P.primeCompl (Localization P'), Ideal.map_mul,
← (I.map _).span_singleton_generator, ← (J.map _).span_singleton_generator,
have := isIntegrallyClosed_of_isLocalization Rₚ P.primeCompl P.primeCompl_le_nonZeroDivisors
have := NoZeroSMulDivisors_of_isLocalization R S Rₚ Sₚ _ P.primeCompl_le_nonZeroDivisors
have := Module.Finite_of_isLocalization R S Rₚ Sₚ P.primeCompl
let L := FractionRing S
let g : Sₚ →+* L := IsLocalization.map _ (M := P') (T := S⁰) (RingHom.id S) h
letI := g.toAlgebra
haveI : IsScalarTower S Sₚ (FractionRing S) := IsScalarTower.of_algebraMap_eq'
(by rw [RingHom.algebraMap_toAlgebra, IsLocalization.map_comp, RingHom.comp_id])
letI := IsFractionRing.isFractionRing_of_isDomain_of_isLocalization
P' Sₚ (FractionRing S)
haveI : Algebra.IsSeparable (FractionRing Rₚ) (FractionRing Sₚ) := by
haveI : IsScalarTower R Rₚ (FractionRing R) := IsScalarTower.of_algebraMap_eq'
(by rw [RingHom.algebraMap_toAlgebra, IsLocalization.lift_comp])
letI := IsFractionRing.isFractionRing_of_isDomain_of_isLocalization
P.primeCompl Rₚ (FractionRing R)
apply Algebra.IsSeparable.of_equiv_equiv
(FractionRing.algEquiv Rₚ (FractionRing R)).symm.toRingEquiv
(FractionRing.algEquiv Sₚ (FractionRing S)).symm.toRingEquiv
apply IsLocalization.ringHom_ext R⁰
ext
simp only [AlgEquiv.toRingEquiv_eq_coe, RingHom.coe_comp,
RingHom.coe_coe, Function.comp_apply, ← IsScalarTower.algebraMap_apply]
simp only [IsScalarTower.algebraMap_apply R Rₚ (FractionRing R), AlgEquiv.coe_ringEquiv,
AlgEquiv.commutes, IsScalarTower.algebraMap_apply R S L,
IsScalarTower.algebraMap_apply S Sₚ L]
simp only [← IsScalarTower.algebraMap_apply]
rw [IsScalarTower.algebraMap_apply R Rₚ (FractionRing Rₚ),
← IsScalarTower.algebraMap_apply Rₚ, ← IsScalarTower.algebraMap_apply]
simp only [Ideal.map_mul, ← spanIntNorm_localization (R := R) (S := S)
(Rₘ := Localization.AtPrime P) (Sₘ := Localization P') _ _ P.primeCompl_le_nonZeroDivisors]
rw [← (I.map _).span_singleton_generator, ← (J.map _).span_singleton_generator,
span_singleton_mul_span_singleton, spanNorm_singleton, spanNorm_singleton,
spanNorm_singleton, span_singleton_mul_span_singleton, _root_.map_mul]
spanNorm_singleton, span_singleton_mul_span_singleton, _root_.map_mul]

/-- The relative norm `Ideal.relNorm R (I : Ideal S)`, where `R` and `S` are Dedekind domains,
and `S` is an extension of `R` that is finite and free as a module. -/
Expand All @@ -190,7 +257,8 @@
map_one' := by dsimp only; rw [one_eq_top, spanNorm_top R, one_eq_top]
map_mul' := spanNorm_mul R

theorem relNorm_apply (I : Ideal S) : relNorm R I = span (Algebra.norm R '' (I : Set S) : Set R) :=
theorem relNorm_apply (I : Ideal S) :
relNorm R I = span (Algebra.intNorm R S '' (I : Set S) : Set R) :=
rfl

@[simp]
Expand All @@ -216,12 +284,12 @@
norm_mem_spanNorm R x hx

@[simp]
theorem relNorm_singleton (r : S) : relNorm R (span ({r} : Set S)) = span {Algebra.norm R r} :=
theorem relNorm_singleton (r : S) : relNorm R (span ({r} : Set S)) = span {Algebra.intNorm R S r} :=
spanNorm_singleton R

theorem map_relNorm (I : Ideal S) {T : Type*} [CommRing T] (f : R →+* T) :
map f (relNorm R I) = span (f ∘ Algebra.norm R '' (I : Set S)) :=
map_spanNorm R I f
map f (relNorm R I) = span (f ∘ Algebra.intNorm R S '' (I : Set S)) :=
map_spanIntNorm R I f

@[mono]
theorem relNorm_mono {I J : Ideal S} (h : I ≤ J) : relNorm R I ≤ relNorm R J :=
Expand Down
26 changes: 26 additions & 0 deletions Mathlib/RingTheory/Localization/Ideal.lean
Original file line number Diff line number Diff line change
Expand Up @@ -231,6 +231,32 @@ theorem ideal_eq_iInf_comap_map_away {S : Finset R} (hS : Ideal.span (α := R) S
rw [pow_add, mul_assoc, ← mul_comm x, e]
exact I.mul_mem_left _ y.2

variable (R) in
lemma _root_.NoZeroSMulDivisors_of_isLocalization (Rₚ Sₚ : Type*) [CommRing Rₚ] [CommRing Sₚ]
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I might misunderstand the situation, but I would say that this lemma holds without assuming that R_p and S_p are rings, you just need to know that they are localizations of R and S at M and at the image of M (and that M is made of non-zero divisors, of course). No?

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

Do you mean I don't need addition? IsLocalization is defined only for semirings.

Copy link
Collaborator

@faenuccio faenuccio Dec 2, 2024

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

No, what I meant is that I do not think that you need S (and so S_p) to be R-algebras, they only need to be R-modules. I would then expect to use IsLocalizedModule instead of IsLocalization for S_p; but you would need to manually provide the morphism from R to S, and perhaps to ask that the image of M is a monoid and that it satisfies NoZeroSmulDivisors.

I agree this can look a bit strange, but since the conclusion only pertains to the R_p-module structure of S_p, it seems to me that you are only using the algebra structures involved to have nice morphisms (which can be achieved by asking that such morphisms exist), and not because you really need S or S_p to be themselves rings.

Copy link
Member Author

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I get the idea, but in practice I don't see a precise statement. If S is not a ring but only an R-module, what is the morphism from R to S? And If S is a ring then any localization as R-module is automatically a ring.

[Algebra R Rₚ] [Algebra R Sₚ] [Algebra S Sₚ] [Algebra Rₚ Sₚ] [IsScalarTower R S Sₚ]
[IsScalarTower R Rₚ Sₚ] (M : Submonoid R) (hM : M ≤ R⁰) [IsLocalization M Rₚ]
[IsLocalization (Algebra.algebraMapSubmonoid S M) Sₚ] [NoZeroSMulDivisors R S] [IsDomain S] :
NoZeroSMulDivisors Rₚ Sₚ := by
have e : Algebra.algebraMapSubmonoid S M ≤ S⁰ :=
Submonoid.map_le_of_le_comap _ <| hM.trans
(nonZeroDivisors_le_comap_nonZeroDivisors_of_injective _
(NoZeroSMulDivisors.algebraMap_injective _ _))
haveI : IsDomain Sₚ := IsLocalization.isDomain_of_le_nonZeroDivisors S e
have : algebraMap Rₚ Sₚ = IsLocalization.map (T := Algebra.algebraMapSubmonoid S M) Sₚ
riccardobrasca marked this conversation as resolved.
Show resolved Hide resolved
(algebraMap R S) (Submonoid.le_comap_map M) := by
apply IsLocalization.ringHom_ext M
simp only [IsLocalization.map_comp, ← IsScalarTower.algebraMap_eq]
rw [NoZeroSMulDivisors.iff_algebraMap_injective, RingHom.injective_iff_ker_eq_bot,
RingHom.ker_eq_bot_iff_eq_zero]
intro x hx
obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective M x
simp only [RingHom.algebraMap_toAlgebra, IsLocalization.map_mk', IsLocalization.mk'_eq_zero_iff,
Subtype.exists, exists_prop, this] at hx ⊢
obtain ⟨_, ⟨a, ha, rfl⟩, H⟩ := hx
simp only [← _root_.map_mul,
(injective_iff_map_eq_zero' _).mp (NoZeroSMulDivisors.algebraMap_injective R S)] at H
exact ⟨a, ha, H⟩

end CommRing

end IsLocalization
Loading