Documentation

Mathlib.RingTheory.Ideal.Cotangent

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.

def Ideal.Cotangent {R : Type u} [CommRing R] (I : Ideal R) :

I ⧸ I ^ 2 as a quotient of I.

Equations
Instances For
    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) :
    def Ideal.toCotangent {R : Type u} [CommRing R] (I : Ideal R) :

    The quotient map from I to I ⧸ I ^ 2.

    Equations
    Instances For
      theorem Ideal.toCotangent_apply {R : Type u} [CommRing R] (I : Ideal R) (a✝ : I) :
      theorem Ideal.mem_toCotangent_ker {R : Type u} [CommRing R] (I : Ideal R) {x : I} :
      theorem Ideal.toCotangent_eq {R : Type u} [CommRing R] (I : Ideal R) {x y : I} :
      I.toCotangent x = I.toCotangent y x - y I ^ 2
      theorem Ideal.toCotangent_eq_zero {R : Type u} [CommRing R] (I : Ideal R) (x : I) :
      I.toCotangent x = 0 x I ^ 2

      The inclusion map I ⧸ I ^ 2 to R ⧸ I ^ 2.

      Equations
      Instances For
        @[simp]
        theorem Ideal.Cotangent.smul_eq_zero_of_mem {R : Type u} [CommRing R] {I : Ideal R} {x : R} (hx : x I) (m : I.Cotangent) :
        x m = 0
        def Ideal.cotangentIdeal {R : Type u} [CommRing R] (I : Ideal R) :
        Ideal (R I ^ 2)

        I ⧸ I ^ 2 as an ideal of R ⧸ I ^ 2.

        Equations
        Instances For
          theorem Ideal.mk_mem_cotangentIdeal {R : Type u} [CommRing R] {I : Ideal R} {x : R} :
          @[deprecated Ideal.range_cotangentToQuotientSquare (since := "2025-01-04")]

          Alias of Ideal.range_cotangentToQuotientSquare.

          noncomputable def Ideal.cotangentEquivIdeal {R : Type u} [CommRing R] (I : Ideal R) :

          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) :

            The lift of f : A →ₐ[R] B to A ⧸ J ^ 2 →ₐ[R] B with J being the kernel of f.

            Equations
            Instances For
              def Ideal.quotCotangent {R : Type u} [CommRing R] (I : Ideal R) :

              The quotient ring of I ⧸ I ^ 2 is R ⧸ I.

              Equations
              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₁ comap f I₂) :

                The map I/I² → J/J² if I ≤ f⁻¹(J).

                Equations
                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₁ 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
                  Instances For
                    @[deprecated IsLocalRing.CotangentSpace (since := "2024-11-11")]

                    Alias of IsLocalRing.CotangentSpace.


                    The A ⧸ I-vector space I ⧸ I ^ 2.

                    Equations
                    Instances For
                      @[deprecated IsLocalRing.subsingleton_cotangentSpace_iff (since := "2024-11-11")]

                      Alias of IsLocalRing.subsingleton_cotangentSpace_iff.

                      @[deprecated IsLocalRing.CotangentSpace.map_eq_top_iff (since := "2024-11-11")]

                      Alias of IsLocalRing.CotangentSpace.map_eq_top_iff.

                      @[deprecated IsLocalRing.CotangentSpace.span_image_eq_top_iff (since := "2024-11-11")]

                      Alias of IsLocalRing.CotangentSpace.span_image_eq_top_iff.

                      @[deprecated IsLocalRing.finrank_cotangentSpace_eq_zero_iff (since := "2024-11-11")]

                      Alias of IsLocalRing.finrank_cotangentSpace_eq_zero_iff.

                      @[deprecated IsLocalRing.finrank_cotangentSpace_eq_zero (since := "2024-11-11")]

                      Alias of IsLocalRing.finrank_cotangentSpace_eq_zero.