Documentation

Mathlib.RingTheory.Presentation

Presentations of algebras #

A presentation of an R-algebra S is a distinguished family of generators and relations.

Main definition #

We also give constructors for localization, base change and composition.

TODO #

Notes #

This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in June 2024.

structure Algebra.Presentation (R : Type u) (S : Type v) [CommRing R] [CommRing S] [Algebra R S] extends Algebra.Generators R S :
Type (max (max (max (t + 1) u) v) (w + 1))

A presentation of an R-algebra S is a family of generators with

  1. rels: The type of relations.
  2. relation : relations → MvPolynomial vars R: The assignment of each relation to a polynomial in the generators.
Instances For
    @[simp]
    theorem Algebra.Presentation.aeval_val_relation {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Algebra.Presentation R S) (i : P.rels) :
    (MvPolynomial.aeval P.val) (P.relation i) = 0
    @[reducible, inline]
    abbrev Algebra.Presentation.Quotient {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Algebra.Presentation R S) :
    Type (max w u)

    The polynomial algebra wrt a family of generators modulo a family of relations.

    Equations
    • P.Quotient = (P.Ring P.ker)
    Instances For
      def Algebra.Presentation.quotientEquiv {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Algebra.Presentation R S) :
      P.Quotient ≃ₐ[P.Ring] S

      P.Quotient is P.Ring-isomorphic to S and in particular R-isomorphic to S.

      Equations
      Instances For
        @[simp]
        theorem Algebra.Presentation.quotientEquiv_mk {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Algebra.Presentation R S) (p : P.Ring) :
        P.quotientEquiv ((Ideal.Quotient.mk P.ker) p) = (algebraMap P.Ring S) p
        @[simp]
        theorem Algebra.Presentation.quotientEquiv_symm {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Algebra.Presentation R S) (x : S) :
        P.quotientEquiv.symm x = (Ideal.Quotient.mk P.ker) (P x)
        noncomputable def Algebra.Presentation.dimension {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Algebra.Presentation R S) :

        Dimension of a presentation defined as the cardinality of the generators minus the cardinality of the relations.

        Note: this definition is completely non-sensical for non-finite presentations and even then for this to make sense, you should assume that the presentation is a complete intersection.

        Equations
        Instances For

          A presentation is finite if there are only finitely-many relations and finitely-many relations.

          Instances
            theorem Algebra.Presentation.ideal_fg_of_isFinite {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (P : Algebra.Presentation R S) [P.IsFinite] :
            P.ker.FG

            If a presentation is finite, the corresponding quotient is of finite presentation.

            Equations
            • =

            If algebraMap R S is bijective, the empty generators are a presentation with no relations.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              noncomputable def Algebra.Presentation.id (R : Type u) [CommRing R] :

              The canonical R-presentation of R with no generators and no relations.

              Equations
              Instances For
                Equations
                • =
                noncomputable def Algebra.Presentation.localizationAway {R : Type u} (S : Type v) [CommRing R] [CommRing S] [Algebra R S] (r : R) [IsLocalization.Away r S] :

                If S is the localization of R away from r, we can construct a natural presentation of S as R-algebra with a single generator X and the relation r * X - 1 = 0.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem Algebra.Presentation.localizationAway_relation {R : Type u} (S : Type v) [CommRing R] [CommRing S] [Algebra R S] (r : R) [IsLocalization.Away r S] (x✝ : Unit) :
                  (Algebra.Presentation.localizationAway S r).relation x✝ = MvPolynomial.C r * MvPolynomial.X () - 1
                  noncomputable def Algebra.Presentation.baseChange {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (T : Type u_1) [CommRing T] [Algebra R T] (P : Algebra.Presentation R S) :

                  If P is a presentation of S over R and T is an R-algebra, we obtain a natural presentation of T ⊗[R] S over T.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    @[simp]
                    theorem Algebra.Presentation.baseChange_relation {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (T : Type u_1) [CommRing T] [Algebra R T] (P : Algebra.Presentation R S) (i : P.rels) :
                    (Algebra.Presentation.baseChange T P).relation i = (MvPolynomial.map (algebraMap R T)) (P.relation i)
                    theorem Algebra.Presentation.baseChange_rels {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (T : Type u_1) [CommRing T] [Algebra R T] (P : Algebra.Presentation R S) :
                    instance Algebra.Presentation.baseChange_isFinite {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] (T : Type u_3) [CommRing T] [Algebra R T] (P : Algebra.Presentation R S) [P.IsFinite] :
                    Equations
                    • =

                    Composition of presentations #

                    Let S be an R-algebra with presentation P and T be an S-algebra with presentation Q. In this section we construct a presentation of T as an R-algebra.

                    For the underlying generators see Algebra.Generators.comp. The family of relations is indexed by Q.rels ⊕ P.rels.

                    We have two canonical maps: MvPolynomial P.vars R →ₐ[R] MvPolynomial (Q.vars ⊕ P.vars) R induced by Sum.inr and aux : MvPolynomial (Q.vars ⊕ P.vars) R →ₐ[R] MvPolynomial Q.vars S induced by the evaluation MvPolynomial P.vars R →ₐ[R] S (see below).

                    Now i : P.rels is mapped to the image of P.relation i under the first map and j : Q.rels is mapped to a pre-image under aux of Q.relation j (see comp_relation_aux for the construction of the pre-image and comp_relation_aux_map for a proof that it is indeed a pre-image).

                    The evaluation map factors as: MvPolynomial (Q.vars ⊕ P.vars) R →ₐ[R] MvPolynomial Q.vars S →ₐ[R] T, where the first map is aux. The goal is to compute that the kernel of this composition is spanned by the relations indexed by Q.rels ⊕ P.rels (span_range_relation_eq_ker_comp). One easily sees that this kernel is the pre-image under aux of the kernel of the evaluation of Q, where the latter is by assumption spanned by the relations Q.relation j.

                    Since aux is surjective (aux_surjective), the pre-image is the sum of the ideal spanned by the constructed pre-images of the Q.relation j and the kernel of aux. It hence remains to show that the kernel of aux is spanned by the image of the P.relation i under the canonical map MvPolynomial P.vars R →ₐ[R] MvPolynomial (Q.vars ⊕ P.vars) R. By assumption this span is the kernel of the evaluation map of P. For this, we use the isomorphism MvPolynomial (Q.vars ⊕ P.vars) R ≃ₐ[R] MvPolynomial Q.vars (MvPolynomial P.vars R) and MvPolynomial.ker_map.

                    noncomputable def Algebra.Presentation.comp {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {T : Type u_1} [CommRing T] [Algebra S T] (Q : Algebra.Presentation S T) (P : Algebra.Presentation R S) [Algebra R T] [IsScalarTower R S T] :

                    Given presentations of T over S and of S over R, we may construct a presentation of T over R.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      @[simp]
                      theorem Algebra.Presentation.comp_rels {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {T : Type u_1} [CommRing T] [Algebra S T] (Q : Algebra.Presentation S T) (P : Algebra.Presentation R S) [Algebra R T] [IsScalarTower R S T] :
                      (Q.comp P).rels = (Q.rels P.rels)
                      theorem Algebra.Presentation.comp_relation {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {T : Type u_1} [CommRing T] [Algebra S T] (Q : Algebra.Presentation S T) (P : Algebra.Presentation R S) [Algebra R T] [IsScalarTower R S T] (a✝ : Q.rels P.rels) :
                      (Q.comp P).relation a✝ = Sum.elim (Algebra.Presentation.comp_relation_aux Q P) (fun (rp : P.rels) => (MvPolynomial.rename Sum.inr) (P.relation rp)) a✝
                      @[simp]
                      theorem Algebra.Presentation.comp_relation_inr {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {T : Type u_4} [CommRing T] [Algebra S T] (Q : Algebra.Presentation S T) (P : Algebra.Presentation R S) [Algebra R T] [IsScalarTower R S T] (r : P.rels) :
                      (Q.comp P).relation (Sum.inr r) = (MvPolynomial.rename Sum.inr) (P.relation r)
                      theorem Algebra.Presentation.comp_aeval_relation_inl {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {T : Type u_3} [CommRing T] [Algebra S T] (Q : Algebra.Presentation S T) (P : Algebra.Presentation R S) [Algebra R T] [IsScalarTower R S T] (r : Q.rels) :
                      (MvPolynomial.aeval (Sum.elim MvPolynomial.X (MvPolynomial.C P.val))) ((Q.comp P).relation (Sum.inl r)) = Q.relation r
                      instance Algebra.Presentation.comp_isFinite {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] {T : Type u_5} [CommRing T] [Algebra S T] (Q : Algebra.Presentation S T) (P : Algebra.Presentation R S) [Algebra R T] [IsScalarTower R S T] [P.IsFinite] [Q.IsFinite] :
                      (Q.comp P).IsFinite
                      Equations
                      • =