Presentation of the module of differentials #
Given a presentation of a R
-algebra S
, we obtain a presentation
of the S
-module Ω[S⁄R]
.
Assume pres : Algebra.Presentation R S
is a presentation of S
as an R
-algebra.
We then have a type pres.vars
for the generators, a type pres.rels
for the relations,
and for each r : pres.rels
we have the relation pres.relation r
in pres.Ring
(which is
a ring of polynomials over R
with variables indexed by pres.vars
).
Then, Ω[pres.Ring⁄R]
identifies to the free module on the type pres.vars
, and
for each r : pres.rels
, we may consider the image of the differential of pres.relation r
in this free module, and by using the (surjective) map pres.Ring → S
, we obtain
an element pres.differentialsRelations.relation r
in the free S
-module on pres.vars
.
The main result in this file is that Ω[S⁄R]
identifies to the quotient of the
free S
-module on pres.vars
by the submodule generated by these
elements pres.differentialsRelations.relation r
. We show this as a consequence
of Algebra.Extension.exact_cotangentComplex_toKaehler
from the file Mathlib.RingTheory.Kaehler.CotangentComplex
.
The shape of the presentation by generators and relations of the S
-module Ω[S⁄R]
that is obtained from a presentation of S
as an R
-algebra.
Equations
- One or more equations did not get rendered due to their size.
Instances For
We obtain here a few compatibilities which allow to compare the two sequences
(pres.rels →₀ S) → (pres.vars →₀ S) → Ω[S⁄R]
and
pres.toExtension.Cotangent →ₗ[S] pres.toExtension.CotangentSpace → Ω[S⁄R]
.
Indeed, there is commutative diagram with a surjective map
hom₁ : (pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent
on the left and
bijections on the middle and on the right. Then, the exactness of the first
sequence shall follow from the exactness of the second which is
Algebra.Extension.exact_cotangentComplex_toKaehler
.
Same as comm₂₃
below, but here we have not yet constructed differentialsSolution
.
The canonical map (pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent
.
Equations
- Algebra.Presentation.differentials.hom₁ pres = Finsupp.linearCombination S fun (r : pres.rels) => Algebra.Extension.Cotangent.mk ⟨pres.relation r, ⋯⟩
Instances For
The S
-module Ω[S⁄R]
contains an obvious solution to the system of linear
equations pres.differentialsRelations.Solution
when pres
is a presentation
of S
as an R
-algebra.
Equations
- pres.differentialsSolution = { var := fun (g : pres.differentialsRelations.G) => (KaehlerDifferential.D R S) (pres.val g), linearCombination_var_relation := ⋯ }
Instances For
The presentation of the S
-module Ω[S⁄R]
deduced from a presentation
of S
as a R
-algebra.
Equations
- One or more equations did not get rendered due to their size.