Documentation

Mathlib.RingTheory.Extension.Cotangent.Basic

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 #

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.

@[reducible, inline]
abbrev Algebra.Extension.CotangentSpace {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Extension R S) :
Type (max w v)

The cotangent space on P = R[X]. This is isomorphic to Sⁿ with n being the number of variables of P.

Equations
Instances For
    noncomputable def Algebra.Extension.cotangentComplex {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Extension R S) :

    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
      @[simp]
      noncomputable def KaehlerDifferential.cotangentComplexBaseChange (R : Type u) (S : Type v) [CommRing R] [CommRing S] (P : Type u_2) (A : Type u_3) [CommRing P] [CommRing A] [Algebra P S] [Algebra P A] [Algebra R P] [Algebra S A] [IsScalarTower P S A] :

      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
        theorem KaehlerDifferential.cotangentComplexBaseChange_tmul {R : Type u} {S : Type v} [CommRing R] [CommRing S] {P : Type u_2} {A : Type u_3} [CommRing P] [CommRing A] [Algebra P S] [Algebra P A] [Algebra R P] [Algebra S A] [IsScalarTower P S A] (a : A) (b : (RingHom.ker (algebraMap P S))) :
        (cotangentComplexBaseChange R S P A) (a ⊗ₜ[P] b) = a (kerToTensor R P A) b,
        noncomputable def Algebra.Extension.CotangentSpace.map {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] (f : P.Hom P') :

        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
        Instances For
          @[simp]
          theorem Algebra.Extension.CotangentSpace.map_tmul {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] (f : P.Hom P') (x : S) (y : P.Ring) :
          theorem Algebra.Extension.CotangentSpace.map_comp {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] {R'' : Type u''} {S'' : Type v''} [CommRing R''] [CommRing S''] [Algebra R'' S''] {P'' : Extension R'' S''} [Algebra R R''] [Algebra S S''] [Algebra R S''] [IsScalarTower R R'' S''] [Algebra R' R''] [Algebra S' S''] [Algebra R' S''] [IsScalarTower R' R'' S''] [IsScalarTower R R' R''] [IsScalarTower S S' S''] (f : P.Hom P') (g : P'.Hom P'') :
          theorem Algebra.Extension.CotangentSpace.map_comp_apply {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] {R'' : Type u''} {S'' : Type v''} [CommRing R''] [CommRing S''] [Algebra R'' S''] {P'' : Extension R'' S''} [Algebra R R''] [Algebra S S''] [Algebra R S''] [IsScalarTower R R'' S''] [Algebra R' R''] [Algebra S' S''] [Algebra R' S''] [IsScalarTower R' R'' S''] [IsScalarTower R R' R''] [IsScalarTower S S' S''] (f : P.Hom P') (g : P'.Hom P'') (x : P.CotangentSpace) :
          theorem Algebra.Extension.CotangentSpace.map_cotangentComplex {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] (f : P.Hom P') (x : P.Cotangent) :
          theorem Algebra.Extension.CotangentSpace.map_comp_cotangentComplex {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] (f : P.Hom P') :
          theorem Algebra.Extension.Hom.sub_aux {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] (f g : P.Hom P') (x y : P.Ring) :
          f.toAlgHom (x * y) - g.toAlgHom (x * y) - (P'.σ ((algebraMap P.Ring S') x) * (f.toAlgHom y - g.toAlgHom y) + P'.σ ((algebraMap P.Ring S') y) * (f.toAlgHom x - g.toAlgHom x)) P'.ker ^ 2
          noncomputable def Algebra.Extension.Hom.subToKer {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] (f g : P.Hom P') :
          P.Ring →ₗ[R] P'.ker

          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
          Instances For
            @[simp]
            theorem Algebra.Extension.Hom.subToKer_apply_coe {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] (f g : P.Hom P') (c : P.Ring) :
            ((f.subToKer g) c) = f.toRingHom c - g.toRingHom c
            noncomputable def Algebra.Extension.Hom.sub {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] (f g : P.Hom P') :

            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
              theorem Algebra.Extension.Hom.sub_one_tmul {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] (f g : P.Hom P') (x : P.Ring) :
              @[simp]
              theorem Algebra.Extension.Hom.sub_tmul {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] (f g : P.Hom P') (r : S) (x : P.Ring) :
              theorem Algebra.Extension.CotangentSpace.map_sub_map {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] (f g : P.Hom P') :
              theorem Algebra.Extension.Cotangent.map_sub_map {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] (f g : P.Hom P') :
              @[reducible, inline]
              noncomputable abbrev Algebra.Extension.toKaehler {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Extension R S) :

              The projection map from the relative cotangent space to the module of differentials.

              Equations
              Instances For
                noncomputable def Algebra.Extension.H1Cotangent {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Extension R S) :

                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
                  @[simp]
                  theorem Algebra.Extension.H1Cotangent.val_add {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} (x y : P.H1Cotangent) :
                  ↑(x + y) = x + y
                  @[simp]
                  theorem Algebra.Extension.H1Cotangent.val_zero {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} :
                  0 = 0
                  @[simp]
                  theorem Algebra.Extension.H1Cotangent.val_smul {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R₀ : Type u_1} [CommRing R₀] [Algebra R₀ S] [Module R₀ P.Cotangent] [IsScalarTower R₀ S P.Cotangent] (r : R₀) (x : P.H1Cotangent) :
                  ↑(r x) = r x
                  instance Algebra.Extension.instIsScalarTowerH1CotangentOfCotangent {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {R₁ : Type u_1} {R₂ : Type u_2} [CommRing R₁] [CommRing R₂] [Algebra R₁ R₂] [Algebra R₁ S] [Algebra R₂ S] [Module R₁ P.Cotangent] [IsScalarTower R₁ S P.Cotangent] [Module R₂ P.Cotangent] [IsScalarTower R₂ S P.Cotangent] [IsScalarTower R₁ R₂ P.Cotangent] :
                  noncomputable def Algebra.Extension.h1Cotangentι {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} :

                  The inclusion of H¹(L_{S/R}) into the conormal space of a presentation.

                  Equations
                  Instances For
                    @[simp]
                    theorem Algebra.Extension.h1Cotangentι_apply {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} (self : (LinearMap.ker P.cotangentComplex)) :
                    h1Cotangentι self = self
                    theorem Algebra.Extension.h1Cotangentι_ext {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} (x y : P.H1Cotangent) (e : x = y) :
                    x = y
                    theorem Algebra.Extension.h1Cotangentι_ext_iff {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P : Extension R S} {x y : P.H1Cotangent} :
                    x = y x = y
                    noncomputable def Algebra.Extension.H1Cotangent.map {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] {P : Extension R S} (f : P.Hom P') :

                    The induced map on the first homology of the (naive) cotangent complex.

                    Equations
                    Instances For
                      @[simp]
                      theorem Algebra.Extension.H1Cotangent.map_apply_coe {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] {P : Extension R S} (f : P.Hom P') (c : (LinearMap.ker P.cotangentComplex)) :
                      ((map f) c) = (Cotangent.map f) c
                      theorem Algebra.Extension.H1Cotangent.map_eq {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] {P : Extension R S} (f g : P.Hom P') :
                      map f = map g
                      @[simp]
                      theorem Algebra.Extension.H1Cotangent.map_comp {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {R' : Type u'} {S' : Type v'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Extension R' S'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] {R'' : Type u''} {S'' : Type v''} [CommRing R''] [CommRing S''] [Algebra R'' S''] {P'' : Extension R'' S''} [Algebra R R''] [Algebra S S''] [Algebra R S''] [IsScalarTower R R'' S''] [Algebra R' R''] [Algebra S' S''] [Algebra R' S''] [IsScalarTower R' R'' S''] [IsScalarTower R R' R''] [IsScalarTower S S' S''] {P : Extension R S} (f : P.Hom P') (g : P'.Hom P'') :
                      map (g.comp f) = S (map g) ∘ₗ map f
                      noncomputable def Algebra.Extension.H1Cotangent.equiv {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P₁ : Extension R S} {P₂ : Extension R S} (f₁ : P₁.Hom P₂) (f₂ : P₂.Hom P₁) :

                      Maps P₁ → P₂ and P₂ → P₁ between extensions induce an isomorphism between H¹(L_P₁) and H¹(L_P₂).

                      Equations
                      Instances For
                        @[simp]
                        theorem Algebra.Extension.H1Cotangent.equiv_apply {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {P₁ : Extension R S} {P₂ : Extension R S} (f₁ : P₁.Hom P₂) (f₂ : P₂.Hom P₁) (c : (LinearMap.ker P₁.cotangentComplex)) :
                        (equiv f₁ f₂) c = (Cotangent.map f₁) c,
                        noncomputable def Algebra.Generators.cotangentSpaceBasis {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {ι : Type w} (P : Generators R S ι) :

                        The canonical basis on the CotangentSpace.

                        Equations
                        Instances For
                          @[simp]
                          theorem Algebra.Generators.cotangentSpaceBasis_repr_tmul {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {ι : Type w} (P : Generators R S ι) (r : S) (x : P.Ring) (i : ι) :
                          noncomputable def Algebra.Generators.cotangentRestrict {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {ι : Type w} (P : Generators R S ι) {σ : Type u_1} {u : σι} (hu : Function.Injective u) :

                          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
                          Instances For
                            theorem Algebra.Generators.cotangentRestrict_mk {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {ι : Type w} (P : Generators R S ι) {σ : Type u_1} {u : σι} (hu : Function.Injective u) (x : P.ker) :
                            ((P.cotangentRestrict hu) (Extension.Cotangent.mk x)) = fun (j : σ) => (MvPolynomial.aeval P.val) ((MvPolynomial.pderiv (u j)) x)
                            @[simp]
                            theorem Algebra.Generators.repr_CotangentSpaceMap {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {ι : Type w} {P : Generators R S ι} {R' : Type u'} {S' : Type v'} {ι' : Type w'} [CommRing R'] [CommRing S'] [Algebra R' S'] {P' : Generators R' S' ι'} [Algebra R R'] [Algebra S S'] [Algebra R S'] [IsScalarTower R R' S'] [IsScalarTower R S S'] (f : P.Hom P') (i : ι) (j : ι') :
                            @[simp]
                            noncomputable def Algebra.Generators.H1Cotangent.equiv {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {ι : Type w} {ι' : Type u_1} (P : Generators R S ι) (P' : Generators R S ι') :

                            H¹(L_{S/R}) is independent of the presentation chosen.

                            Equations
                            Instances For
                              @[simp]
                              theorem Algebra.Generators.H1Cotangent.equiv_apply {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {ι : Type w} {ι' : Type u_1} (P : Generators R S ι) (P' : Generators R S ι') (c : (LinearMap.ker P.toExtension.cotangentComplex)) :
                              @[reducible, inline]
                              abbrev Algebra.H1Cotangent (R : Type u) (S : Type v) [CommRing R] [CommRing S] [Algebra R S] :
                              Type (max u v)

                              H¹(L_{S/R}), the first homology of the (naive) cotangent complex of S over R.

                              Equations
                              Instances For
                                noncomputable def Algebra.H1Cotangent.map (R : Type u) (S : Type v) [CommRing R] [CommRing S] [Algebra R S] (S' : Type u_2) [CommRing S'] [Algebra R S'] (T : Type w) [CommRing T] [Algebra R T] [Algebra S T] [IsScalarTower R S T] [Algebra S' T] [IsScalarTower R S' T] :

                                The induced map on the first homology of the (naive) cotangent complex of S over R.

                                Equations
                                Instances For
                                  noncomputable def Algebra.H1Cotangent.mapEquiv (R : Type u) (S : Type v) [CommRing R] [CommRing S] [Algebra R S] (S' : Type u_2) [CommRing S'] [Algebra R S'] (e : S ≃ₐ[R] S') :

                                  Isomorphic algebras induce isomorphic H¹(L_{S/R}).

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    @[reducible, inline]
                                    noncomputable abbrev Algebra.Generators.equivH1Cotangent {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {ι : Type w} (P : Generators R S ι) :

                                    H¹(L_{S/R}) is independent of the presentation chosen.

                                    Equations
                                    Instances For