The module I ⧸ I ^ 2
#
In this file, we provide special API support for the module I ⧸ I ^ 2
. The official
definition is a quotient module of I
, but the alternative definition as an ideal of R ⧸ I ^ 2
is
also given, and the two are R
-equivalent as in Ideal.cotangentEquivIdeal
.
Additional support is also given to the cotangent space m ⧸ m ^ 2
of a local ring.
instance
Ideal.instAddCommGroupCotangent
{R : Type u}
[CommRing R]
(I : Ideal R)
:
AddCommGroup I.Cotangent
instance
Ideal.Cotangent.moduleOfTower
{R : Type u}
{S : Type v}
[CommRing R]
[CommSemiring S]
[Algebra S R]
(I : Ideal R)
:
Module S I.Cotangent
Equations
instance
Ideal.Cotangent.isScalarTower
{R : Type u}
{S : Type v}
{S' : Type w}
[CommRing R]
[CommSemiring S]
[Algebra S R]
[CommSemiring S']
[Algebra S' R]
[Algebra S S']
[IsScalarTower S S' R]
(I : Ideal R)
:
IsScalarTower S S' I.Cotangent
Equations
- ⋯ = ⋯
instance
Ideal.instIsNoetherianCotangentOfSubtypeMem
{R : Type u}
[CommRing R]
(I : Ideal R)
[IsNoetherian R ↥I]
:
IsNoetherian R I.Cotangent
Equations
- ⋯ = ⋯
theorem
Ideal.toCotangent_apply
{R : Type u}
[CommRing R]
(I : Ideal R)
(a✝ : ↥I)
:
I.toCotangent a✝ = Submodule.Quotient.mk a✝
theorem
Ideal.map_toCotangent_ker
{R : Type u}
[CommRing R]
(I : Ideal R)
:
Submodule.map (Submodule.subtype I) (LinearMap.ker I.toCotangent) = I ^ 2
theorem
Ideal.mem_toCotangent_ker
{R : Type u}
[CommRing R]
(I : Ideal R)
{x : ↥I}
:
x ∈ LinearMap.ker I.toCotangent ↔ ↑x ∈ I ^ 2
theorem
Ideal.toCotangent_surjective
{R : Type u}
[CommRing R]
(I : Ideal R)
:
Function.Surjective ⇑I.toCotangent
theorem
Ideal.toCotangent_range
{R : Type u}
[CommRing R]
(I : Ideal R)
:
LinearMap.range I.toCotangent = ⊤
theorem
Ideal.cotangent_subsingleton_iff
{R : Type u}
[CommRing R]
(I : Ideal R)
:
Subsingleton I.Cotangent ↔ IsIdempotentElem I
theorem
Ideal.to_quotient_square_comp_toCotangent
{R : Type u}
[CommRing R]
(I : Ideal R)
:
I.cotangentToQuotientSquare ∘ₗ I.toCotangent = Submodule.mkQ (I ^ 2) ∘ₗ Submodule.subtype I
@[simp]
theorem
Ideal.toCotangent_to_quotient_square
{R : Type u}
[CommRing R]
(I : Ideal R)
(x : ↥I)
:
I.cotangentToQuotientSquare (I.toCotangent x) = (Submodule.mkQ (I ^ 2)) ↑x
theorem
Ideal.isTorsionBySet_cotangent
{R : Type u}
[CommRing R]
(I : Ideal R)
:
Module.IsTorsionBySet R I.Cotangent ↑I
I ⧸ I ^ 2
as an ideal of R ⧸ I ^ 2
.
Equations
- I.cotangentIdeal = Submodule.map (Ideal.Quotient.mk (I ^ 2)).toSemilinearMap I
Instances For
theorem
Ideal.to_quotient_square_range
{R : Type u}
[CommRing R]
(I : Ideal R)
:
LinearMap.range I.cotangentToQuotientSquare = Submodule.restrictScalars R I.cotangentIdeal
The equivalence of the two definitions of I / I ^ 2
, either as the quotient of I
or the
ideal of R / I ^ 2
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
theorem
Ideal.cotangentEquivIdeal_symm_apply
{R : Type u}
[CommRing R]
(I : Ideal R)
(x : R)
(hx : x ∈ I)
:
I.cotangentEquivIdeal.symm ⟨(Submodule.mkQ (I ^ 2)) x, ⋯⟩ = I.toCotangent ⟨x, hx⟩
def
AlgHom.kerSquareLift
{R : Type u}
[CommRing R]
{A : Type u_1}
{B : Type u_2}
[CommRing A]
[CommRing B]
[Algebra R A]
[Algebra R B]
(f : A →ₐ[R] B)
:
A ⧸ RingHom.ker f.toRingHom ^ 2 →ₐ[R] B
The lift of f : A →ₐ[R] B
to A ⧸ J ^ 2 →ₐ[R] B
with J
being the kernel of f
.
Equations
- f.kerSquareLift = { toRingHom := Ideal.Quotient.lift (RingHom.ker f.toRingHom ^ 2) f.toRingHom ⋯, commutes' := ⋯ }
Instances For
theorem
AlgHom.ker_kerSquareLift
{R : Type u}
[CommRing R]
{A : Type u_1}
{B : Type u_2}
[CommRing A]
[CommRing B]
[Algebra R A]
[Algebra R B]
(f : A →ₐ[R] B)
:
RingHom.ker f.kerSquareLift.toRingHom = (RingHom.ker f.toRingHom).cotangentIdeal
instance
Ideal.Algebra.kerSquareLift
{R : Type u}
[CommRing R]
{A : Type u_1}
[CommRing A]
[Algebra R A]
:
Algebra (R ⧸ RingHom.ker (algebraMap R A) ^ 2) A
Equations
- Ideal.Algebra.kerSquareLift = (Algebra.ofId R A).kerSquareLift.toAlgebra
instance
Ideal.instIsScalarTowerQuotientHPowKerRingHomAlgebraMapOfNat
{R : Type u}
[CommRing R]
{A : Type u_1}
{B : Type u_2}
[CommRing A]
[CommRing B]
[Algebra R A]
[Algebra R B]
[Algebra A B]
[IsScalarTower R A B]
:
IsScalarTower R (A ⧸ RingHom.ker (algebraMap A B) ^ 2) B
Equations
- ⋯ = ⋯
The quotient ring of I ⧸ I ^ 2
is R ⧸ I
.
Equations
- I.quotCotangent = (Ideal.quotEquivOfEq ⋯).trans ((DoubleQuot.quotQuotEquivQuotSup (I ^ 2) I).trans (Ideal.quotEquivOfEq ⋯))
Instances For
def
Ideal.mapCotangent
{R : Type u}
[CommRing R]
{A : Type u_1}
{B : Type u_2}
[CommRing A]
[CommRing B]
[Algebra R A]
[Algebra R B]
(I₁ : Ideal A)
(I₂ : Ideal B)
(f : A →ₐ[R] B)
(h : I₁ ≤ Ideal.comap f I₂)
:
The map I/I² → J/J²
if I ≤ f⁻¹(J)
.
Equations
- I₁.mapCotangent I₂ f h = (Submodule.restrictScalars R (I₁ • ⊤)).mapQ (Submodule.restrictScalars R (I₂ • ⊤)) (f.toLinearMap.restrict h) ⋯
Instances For
@[simp]
theorem
Ideal.mapCotangent_toCotangent
{R : Type u}
[CommRing R]
{A : Type u_1}
{B : Type u_2}
[CommRing A]
[CommRing B]
[Algebra R A]
[Algebra R B]
(I₁ : Ideal A)
(I₂ : Ideal B)
(f : A →ₐ[R] B)
(h : I₁ ≤ Ideal.comap f I₂)
(x : ↥I₁)
:
(I₁.mapCotangent I₂ f h) (I₁.toCotangent x) = I₂.toCotangent ⟨f ↑x, ⋯⟩
@[reducible, inline]
The A ⧸ I
-vector space I ⧸ I ^ 2
.
Equations
- IsLocalRing.CotangentSpace R = (IsLocalRing.maximalIdeal R).Cotangent
Instances For
instance
IsLocalRing.instModuleResidueFieldCotangentSpace
(R : Type u_1)
[CommRing R]
[IsLocalRing R]
:
Equations
- IsLocalRing.instModuleResidueFieldCotangentSpace R = (IsLocalRing.maximalIdeal R).cotangentModule
instance
IsLocalRing.instIsScalarTowerResidueFieldCotangentSpace
(R : Type u_1)
[CommRing R]
[IsLocalRing R]
:
Equations
- ⋯ = ⋯
instance
IsLocalRing.instFiniteDimensionalResidueFieldCotangentSpaceOfIsNoetherianRing
(R : Type u_1)
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
:
Equations
- ⋯ = ⋯
theorem
IsLocalRing.subsingleton_cotangentSpace_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
:
theorem
IsLocalRing.CotangentSpace.map_eq_top_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
{M : Submodule R ↥(IsLocalRing.maximalIdeal R)}
:
Submodule.map (IsLocalRing.maximalIdeal R).toCotangent M = ⊤ ↔ M = ⊤
theorem
IsLocalRing.CotangentSpace.span_image_eq_top_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
{s : Set ↥(IsLocalRing.maximalIdeal R)}
:
Submodule.span (IsLocalRing.ResidueField R) (⇑(IsLocalRing.maximalIdeal R).toCotangent '' s) = ⊤ ↔ Submodule.span R s = ⊤
theorem
IsLocalRing.finrank_cotangentSpace_eq_zero_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
:
theorem
IsLocalRing.finrank_cotangentSpace_le_one_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
:
@[deprecated IsLocalRing.CotangentSpace]
Alias of IsLocalRing.CotangentSpace
.
The A ⧸ I
-vector space I ⧸ I ^ 2
.
Instances For
@[deprecated IsLocalRing.subsingleton_cotangentSpace_iff]
theorem
LocalRing.subsingleton_cotangentSpace_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
:
@[deprecated IsLocalRing.CotangentSpace.map_eq_top_iff]
theorem
LocalRing.map_eq_top_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
{M : Submodule R ↥(IsLocalRing.maximalIdeal R)}
:
Submodule.map (IsLocalRing.maximalIdeal R).toCotangent M = ⊤ ↔ M = ⊤
@[deprecated IsLocalRing.CotangentSpace.span_image_eq_top_iff]
theorem
LocalRing.span_image_eq_top_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
{s : Set ↥(IsLocalRing.maximalIdeal R)}
:
Submodule.span (IsLocalRing.ResidueField R) (⇑(IsLocalRing.maximalIdeal R).toCotangent '' s) = ⊤ ↔ Submodule.span R s = ⊤
@[deprecated IsLocalRing.finrank_cotangentSpace_eq_zero_iff]
theorem
LocalRing.finrank_cotangentSpace_eq_zero_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
:
@[deprecated IsLocalRing.finrank_cotangentSpace_eq_zero]
@[deprecated IsLocalRing.finrank_cotangentSpace_le_one_iff]
theorem
LocalRing.finrank_cotangentSpace_le_one_iff
{R : Type u_1}
[CommRing R]
[IsLocalRing R]
[IsNoetherianRing R]
: