Cotangent and localization away #
Let R → S → T be algebras such that T is the localization of S away from one
element, where S is generated over R by P : R[X] → S with kernel I and
Q : S[Y] → T is the canonical S-presentation of T with kernel K.
Denote by J the kernel of the composition R[X,Y] → T.
This file proves J/J² ≃ₗ[T] T ⊗[S] (I/I²) × K/K². For this we establish the exact sequence:
0 → T ⊗[S] (I/I²) → J/J² → K/K² → 0
and use that K/K² is free, so the sequence splits. The first part of the file
shows the exactness on the left and the rest of the file deduces the exact sequence
and the splitting from the Jacobi Zariski sequence.
Main results #
Algebra.Generators.liftBaseChange_injective:T ⊗[S] (I/I²) → J/J²is injective ifTis the localization ofSaway from an element.Algebra.Generators.cotangentCompLocalizationAwayEquiv:J/J² ≃ₗ[T] T ⊗[S] (I/I²) × K/K².
If R[X] → S generates S, T is the localization of S away from g and
f is a pre-image of g in R[X], this is the R-algebra map R[X,Y] →ₐ[R] (R[X]/I²)[1/f]
defined via mapping Y to 1/f.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Let R → S → T be algebras such that T is the localization of S away from one
element, where S is generated over R by P with kernel I and Q is the
canonical S-presentation of T. Denote by J the kernel of the composition
R[X,Y] → T. Then T ⊗[S] (I/I²) → J/J² is injective.
In the notation of the module docstring: Since T is standard smooth
of relative dimension one over S, K/K² is free of rank one generated
by the image of g * X - 1.
This is the section K/K² → J/J² defined by sending the image of g * X - 1 to x : J/J².
Equations
- Algebra.Generators.cotangentCompAwaySec g P x = ((Algebra.Generators.basisCotangentAway T g).constr T) fun (x_1 : Unit) => x
Instances For
By construction, the section cotangentCompAwaySec sends g * X - 1 to x.
The section cotangentCompAwaySec is indeed a section of the canonical map J/J² → K/K².
Let S be generated over R by P : R[X] → S with kernel I and let T
be the localization of S away from g generated over S by S[Y] → T with
kernel K.
Denote by J the kernel of the induced R[X, Y] → T. Then
J/J² ≃ₗ[T] T ⊗[S] (I/I²) × (K/K²).
This is the splitting characterised by x ↦ (0, g * X - 1).