Presentations of algebras #
A presentation of an R
-algebra S
is a distinguished family of generators and relations.
Main definition #
Algebra.Presentation
: A presentation of anR
-algebraS
is a family of generators withrels
: The type of relations.relation : relations → MvPolynomial vars R
: The assignment of each relation to a polynomial in the generators.
Algebra.Presentation.IsFinite
: A presentation is called finite, if both variables and relations are finite.Algebra.Presentation.dimension
: The dimension of a presentation is the number of generators minus the number of relations.
We also give constructors for localization, base change and composition.
TODO #
- Define
Hom
s of presentations.
Notes #
This contribution was created as part of the AIM workshop "Formalizing algebraic geometry" in June 2024.
A presentation of an R
-algebra S
is a family of
generators with
rels
: The type of relations.relation : relations → MvPolynomial vars R
: The assignment of each relation to a polynomial in the generators.
- val : self.vars → S
- σ' : S → MvPolynomial self.vars R
- aeval_val_σ' : ∀ (s : S), (MvPolynomial.aeval self.val) (self.σ' s) = s
- rels : Type t
The type of relations.
- relation : self.rels → self.Ring
The assignment of each relation to a polynomial in the generators.
- span_range_relation_eq_ker : Ideal.span (Set.range self.relation) = self.ker
The relations span the kernel of the canonical map.
Instances For
The polynomial algebra wrt a family of generators modulo a family of relations.
Instances For
P.Quotient
is P.Ring
-isomorphic to S
and in particular R
-isomorphic to S
.
Equations
- P.quotientEquiv = Ideal.quotientKerAlgEquivOfRightInverse ⋯
Instances For
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.
Instances For
A presentation is finite if there are only finitely-many relations and finitely-many relations.
Instances
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
Equations
- ⋯ = ⋯
The canonical R
-presentation of R
with no generators and no relations.
Instances For
Equations
- ⋯ = ⋯
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
Equations
- ⋯ = ⋯
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
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
.
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
Equations
- ⋯ = ⋯