Unramified algebras over local rings #
Main results #
Algebra.FormallyUnramified.iff_map_maximalIdeal_eq
: LetR
be a local ring,A
be a localR
-algebra essentially of finite type. ThenA/R
is unramified if and only ifκA/κR
is separable, andm_R S = m_S
.Algebra.isUnramifiedAt_iff_map_eq
: LetA
be an essentially of finite typeR
-algebra,q
be a prime overp
. ThenA
is unramified atp
if and only ifκ(q)/κ(p)
is separable, andpS_q = qS_q
.
instance
Algebra.instFormallyUnramifiedResidueField_1
{R : Type u_1}
{S : Type u_2}
[CommRing R]
[CommRing S]
[Algebra R S]
[IsLocalRing R]
[IsLocalRing S]
[IsLocalHom (algebraMap R S)]
[FormallyUnramified R S]
:
instance
Algebra.instFiniteResidueFieldOfFormallyUnramified
{R : Type u_1}
{S : Type u_2}
[CommRing R]
[CommRing S]
[Algebra R S]
[IsLocalRing R]
[IsLocalRing S]
[IsLocalHom (algebraMap R S)]
[EssFiniteType R S]
[FormallyUnramified R S]
:
Stacks Tag 00UW ((2))
instance
Algebra.instIsSeparableResidueFieldOfFormallyUnramified
{R : Type u_1}
{S : Type u_2}
[CommRing R]
[CommRing S]
[Algebra R S]
[IsLocalRing R]
[IsLocalRing S]
[IsLocalHom (algebraMap R S)]
[EssFiniteType R S]
[FormallyUnramified R S]
:
Stacks Tag 00UW ((2))
theorem
Algebra.FormallyUnramified.isField_quotient_map_maximalIdeal
{R : Type u_1}
{S : Type u_2}
[CommRing R]
[CommRing S]
[Algebra R S]
[IsLocalRing R]
[IsLocalRing S]
[IsLocalHom (algebraMap R S)]
[EssFiniteType R S]
[FormallyUnramified R S]
:
IsField (S ⧸ Ideal.map (algebraMap R S) (IsLocalRing.maximalIdeal R))
theorem
Algebra.FormallyUnramified.map_maximalIdeal
{R : Type u_1}
{S : Type u_2}
[CommRing R]
[CommRing S]
[Algebra R S]
[IsLocalRing R]
[IsLocalRing S]
[IsLocalHom (algebraMap R S)]
[EssFiniteType R S]
[FormallyUnramified R S]
:
Stacks Tag 00UW ((1))
theorem
Algebra.FormallyUnramified.of_map_maximalIdeal
{R : Type u_1}
{S : Type u_2}
[CommRing R]
[CommRing S]
[Algebra R S]
[IsLocalRing R]
[IsLocalRing S]
[IsLocalHom (algebraMap R S)]
[EssFiniteType R S]
[Algebra.IsSeparable (IsLocalRing.ResidueField R) (IsLocalRing.ResidueField S)]
(H : Ideal.map (algebraMap R S) (IsLocalRing.maximalIdeal R) = IsLocalRing.maximalIdeal S)
:
theorem
Algebra.FormallyUnramified.iff_map_maximalIdeal_eq
{R : Type u_1}
{S : Type u_2}
[CommRing R]
[CommRing S]
[Algebra R S]
[IsLocalRing R]
[IsLocalRing S]
[IsLocalHom (algebraMap R S)]
[EssFiniteType R S]
:
noncomputable instance
Algebra.instResidueFieldOfLiesOver
(R : Type u)
{S : Type u}
[CommRing R]
[CommRing S]
[Algebra R S]
(p : Ideal R)
[p.IsPrime]
(q : Ideal S)
[q.IsPrime]
[q.LiesOver p]
:
Equations
- Algebra.instResidueFieldOfLiesOver R p q = (Ideal.ResidueField.mapₐ p q ⋯).toAlgebra
theorem
Algebra.isUnramifiedAt_iff_map_eq
(R : Type u)
{S : Type u}
[CommRing R]
[CommRing S]
[Algebra R S]
[EssFiniteType R S]
(p : Ideal R)
[p.IsPrime]
(q : Ideal S)
[q.IsPrime]
[q.LiesOver p]
:
Let A
be an essentially of finite type R
-algebra, q
be a prime over p
.
Then A
is unramified at p
if and only if κ(q)/κ(p)
is separable, and pS_q = qS_q
.