Adjunctions in CommRingCat
#
Main results #
CommRingCat.adj
:σ ↦ ℤ[σ]
is left adjoint to the forgetful functorCommRingCat ⥤ Type
.CommRingCat.coyonedaAdj
:Fun(-, R)
is left adjoint toHom_{CRing}(R, -)
.CommRingCat.monoidAlgebraAdj
:G ↦ R[G]
asCommGrp ⥤ R-Alg
is left adjoint toS ↦ Sˣ
.CommRingCat.unitsAdj
:G ↦ ℤ[G]
is left adjoint toS ↦ Sˣ
.
The free functor Type u ⥤ CommRingCat
sending a type X
to the multivariable (commutative)
polynomials with variables x : X
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The free-forgetful adjunction for commutative rings.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Fun(-, -)
as a functor Type vᵒᵖ ⥤ CommRingCat ⥤ CommRingCat
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
@[simp]
theorem
CommRingCat.coyoneda_map_app
{m n : Type vᵒᵖ}
(f : m ⟶ n)
(R : CommRingCat)
:
(coyoneda.map f).app R = ofHom (Pi.ringHom fun (x : Opposite.unop n) => Pi.evalRingHom (fun (a : Opposite.unop m) => ↑R) (f.unop x))
@[simp]
theorem
CommRingCat.coyoneda_obj_map
(n : Type vᵒᵖ)
{R S : CommRingCat}
(φ : R ⟶ S)
:
(coyoneda.obj n).map φ = ofHom (Pi.ringHom fun (x : Opposite.unop n) => (Hom.hom φ).comp (Pi.evalRingHom (fun (a : Opposite.unop n) => ↑R) x))
The adjunction Hom_{CRing}(Fun(n, R), S) ≃ Fun(n, Hom_{CRing}(R, S))
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
If n
is a singleton, Hom(n, -)
is the identity in CommRingCat
.
Equations
- CommRingCat.coyonedaUnique = CategoryTheory.NatIso.ofComponents (fun (X : CommRingCat) => (RingEquiv.piUnique fun (i : Opposite.unop (Opposite.op n)) => ↑X).toCommRingCatIso) ⋯
Instances For
@[simp]
theorem
CommRingCat.coyonedaUnique_inv_app_hom_apply
{n : Type v}
[Unique n]
(X : CommRingCat)
(a✝ : ↑X)
(a✝¹ : Opposite.unop (Opposite.op n))
:
@[simp]
theorem
CommRingCat.coyonedaUnique_hom_app_hom_apply
{n : Type v}
[Unique n]
(X : CommRingCat)
(a✝ : Opposite.unop (Opposite.op n) → ↑X)
:
The monoid algebra functor CommGrp ⥤ R-Alg
given by G ↦ R[G]
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[simp]
@[simp]
theorem
CommRingCat.monoidAlgebra_map
(R : CommRingCat)
{X✝ Y✝ : CommMonCat}
(f : X✝ ⟶ Y✝)
:
R.monoidAlgebra.map f = CategoryTheory.Under.homMk (ofHom (MonoidAlgebra.mapDomainRingHom (↑R) (CommMonCat.Hom.hom f))) ⋯
Equations
- One or more equations did not get rendered due to their size.
@[simp]
@[simp]
The adjunction G ↦ R[G]
and S ↦ S
between CommGrp
and R-Alg
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The adjunction G ↦ ℤ[G]
and R ↦ Rˣ
between CommGrp
and CommRing
.