Naive cotangent complex associated to a presentation. #
Given a presentation 0 → I → R[x₁,...,xₙ] → S → 0 (or equivalently a closed embedding S ↪ Aⁿ
defined by I), we may define the (naive) cotangent complex I/I² → ⨁ᵢ S dxᵢ → Ω[S/R] → 0.
Main results #
Algebra.Extension.Cotangent: The conormal spaceI/I². (Defined inGenerators/Basic)Algebra.Extension.CotangentSpace: The cotangent space⨁ᵢ S dxᵢ.Algebra.Generators.cotangentSpaceBasis: The canonical basis on⨁ᵢ S dxᵢ.Algebra.Extension.CotangentComplex: The mapI/I² → ⨁ᵢ S dxᵢ.Algebra.Extension.toKaehler: The projection⨁ᵢ S dxᵢ → Ω[S/R].Algebra.Extension.toKaehler_surjective: The map⨁ᵢ S dxᵢ → Ω[S/R]is surjective.Algebra.Extension.exact_cotangentComplex_toKaehler:I/I² → ⨁ᵢ S dxᵢ → Ω[S/R]is exact.Algebra.Extension.Hom.Sub: Iffandgare two maps between presentations,f - ginduces a map⨁ᵢ S dxᵢ → I/I²that makesfandghomotopic.Algebra.Extension.H1Cotangent: The first homology of the (naive) cotangent complex ofSoverR, induced by a given presentation.Algebra.H1Cotangent:H¹(L_{S/R}), the first homology of the (naive) cotangent complex ofSoverR.
Implementation detail #
We actually develop these material for general extensions (i.e. surjection P → S) so that we can
apply them to infinitesimal smooth (or versal) extensions later.
The cotangent space on P = R[X].
This is isomorphic to Sⁿ with n being the number of variables of P.
Equations
- P.CotangentSpace = TensorProduct P.Ring S Ω[P.Ring⁄R]
Instances For
The cotangent complex given by a presentation R[X] → S (i.e. a closed embedding S ↪ Aⁿ).
Equations
- One or more equations did not get rendered due to their size.
Instances For
This is (isomorphic to) the base change of the contangent complex to A, but
the domain and codomains of this are more manageable.
Equations
Instances For
This is the map on the cotangent space associated to a map of presentation.
The matrix associated to this map is the Jacobian matrix. See CotangentSpace.repr_map.
Equations
- Algebra.Extension.CotangentSpace.map f = LinearMap.liftBaseChange S (↑P.Ring ((TensorProduct.mk P'.Ring S' Ω[P'.Ring⁄R']) 1) ∘ₗ KaehlerDifferential.map R R' P.Ring P'.Ring)
Instances For
If f and g are two maps P → P' between presentations,
then the image of f - g is in the kernel of P' → S.
Equations
- f.subToKer g = LinearMap.codRestrict (Submodule.restrictScalars R P'.ker) (f.toAlgHom.toLinearMap - g.toAlgHom.toLinearMap) ⋯
Instances For
If f and g are two maps P → P' between presentations,
their difference induces a map P.CotangentSpace →ₗ[S] P'.Cotangent that makes two maps
between the cotangent complexes homotopic.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The first homology of the (naive) cotangent complex of S over R,
induced by a given presentation 0 → I → P → R → 0,
defined as the kernel of I/I² → S ⊗[P] Ω[P⁄R].
Equations
Instances For
The induced map on the first homology of the (naive) cotangent complex.
Equations
Instances For
Maps P₁ → P₂ and P₂ → P₁ between extensions
induce an isomorphism between H¹(L_P₁) and H¹(L_P₂).
Equations
- Algebra.Extension.H1Cotangent.equiv f₁ f₂ = { toLinearMap := Algebra.Extension.H1Cotangent.map f₁, invFun := ⇑(Algebra.Extension.H1Cotangent.map f₂), left_inv := ⋯, right_inv := ⋯ }
Instances For
The canonical basis on the CotangentSpace.
Equations
Instances For
Given generators R[xᵢ] → S and an injective map σ → ι, this is the
composition I/I² → ⊕ S dxᵢ → ⊕ S dxᵢ where the second i only runs over σ.
Equations
- P.cotangentRestrict hu = Finsupp.lcomapDomain u hu ∘ₗ ↑P.cotangentSpaceBasis.repr ∘ₗ P.toExtension.cotangentComplex
Instances For
H¹(L_{S/R}) is independent of the presentation chosen.
Equations
Instances For
The induced map on the first homology of the (naive) cotangent complex of S over R.
Equations
Instances For
Isomorphic algebras induce isomorphic H¹(L_{S/R}).
Equations
- One or more equations did not get rendered due to their size.
Instances For
H¹(L_{S/R}) is independent of the presentation chosen.