Documentation

Mathlib.Algebra.Category.AlgCat.Basic

Category instance for algebras over a commutative ring #

We introduce the bundled category AlgCat of algebras over a fixed commutative ring R along with the forgetful functors to RingCat and ModuleCat. We furthermore show that the functor associating to a type the free R-algebra on that type is left adjoint to the forgetful functor.

structure AlgCat (R : Type u) [CommRing R] :
Type (max u (v + 1))

The category of R-algebras and their morphisms.

  • carrier : Type v

    The underlying type.

  • isRing : Ring self
  • isAlgebra : Algebra R self
Instances For
    @[reducible, inline]
    abbrev AlgCat.of (R : Type u) [CommRing R] (X : Type v) [Ring X] [Algebra R X] :

    The object in the category of R-algebras associated to a type equipped with the appropriate typeclasses. This is the preferred way to construct a term of AlgCat R.

    Equations
    • AlgCat.of R X = { carrier := X, isRing := inst✝¹, isAlgebra := inst✝ }
    Instances For
      theorem AlgCat.coe_of (R : Type u) [CommRing R] (X : Type v) [Ring X] [Algebra R X] :
      (of R X) = X
      structure AlgCat.Hom {R : Type u} [CommRing R] (A B : AlgCat R) :

      The type of morphisms in AlgCat R.

      • hom' : A →ₐ[R] B

        The underlying algebra map.

      Instances For
        theorem AlgCat.Hom.ext {R : Type u} {inst✝ : CommRing R} {A B : AlgCat R} {x y : A.Hom B} (hom' : x.hom' = y.hom') :
        x = y
        theorem AlgCat.Hom.ext_iff {R : Type u} {inst✝ : CommRing R} {A B : AlgCat R} {x y : A.Hom B} :
        x = y x.hom' = y.hom'
        Equations
        • One or more equations did not get rendered due to their size.
        Equations
        • One or more equations did not get rendered due to their size.
        @[reducible, inline]
        abbrev AlgCat.Hom.hom {R : Type u} [CommRing R] {A B : AlgCat R} (f : A.Hom B) :
        A →ₐ[R] B

        Turn a morphism in AlgCat back into an AlgHom.

        Equations
        Instances For
          @[reducible, inline]
          abbrev AlgCat.ofHom {R : Type u} [CommRing R] {A B : Type v} [Ring A] [Ring B] [Algebra R A] [Algebra R B] (f : A →ₐ[R] B) :
          of R A of R B

          Typecheck an AlgHom as a morphism in AlgCat.

          Equations
          Instances For
            def AlgCat.Hom.Simps.hom {R : Type u} [CommRing R] (A B : AlgCat R) (f : A.Hom B) :
            A →ₐ[R] B

            Use the ConcreteCategory.hom projection for @[simps] lemmas.

            Equations
            Instances For

              The results below duplicate the ConcreteCategory simp lemmas, but we can keep them for dsimp.

              @[simp]
              theorem AlgCat.hom_comp (R : Type u) [CommRing R] {A B C : AlgCat R} (f : A B) (g : B C) :
              theorem AlgCat.hom_ext (R : Type u) [CommRing R] {A B : AlgCat R} {f g : A B} (hf : Hom.hom f = Hom.hom g) :
              f = g
              theorem AlgCat.hom_ext_iff {R : Type u} [CommRing R] {A B : AlgCat R} {f g : A B} :
              @[simp]
              theorem AlgCat.hom_ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] (f : X →ₐ[R] Y) :
              @[simp]
              theorem AlgCat.ofHom_hom (R : Type u) [CommRing R] {A B : AlgCat R} (f : A B) :
              @[simp]
              theorem AlgCat.ofHom_id (R : Type u) [CommRing R] {X : Type v} [Ring X] [Algebra R X] :
              @[simp]
              theorem AlgCat.ofHom_comp (R : Type u) [CommRing R] {X Y Z : Type v} [Ring X] [Ring Y] [Ring Z] [Algebra R X] [Algebra R Y] [Algebra R Z] (f : X →ₐ[R] Y) (g : Y →ₐ[R] Z) :
              theorem AlgCat.ofHom_apply {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] (f : X →ₐ[R] Y) (x : X) :
              Equations
              theorem AlgCat.forget_obj (R : Type u) [CommRing R] {A : AlgCat R} :
              Equations
              • One or more equations did not get rendered due to their size.
              Equations
              • One or more equations did not get rendered due to their size.
              @[deprecated CategoryTheory.Iso.refl (since := "2025-05-15")]
              def AlgCat.ofSelfIso {R : Type u} [CommRing R] (M : AlgCat R) :
              of R M M

              Forgetting to the underlying type and then building the bundled object returns the original algebra.

              Equations
              Instances For

                The "free algebra" functor, sending a type S to the free algebra on S.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  @[simp]
                  theorem AlgCat.free_map (R : Type u) [CommRing R] {X✝ Y✝ : Type u} (f : X✝ Y✝) :
                  @[simp]
                  theorem AlgCat.free_obj (R : Type u) [CommRing R] (S : Type u) :
                  (free R).obj S = of R (FreeAlgebra R S)

                  The free/forget adjunction for R-algebras.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    def AlgEquiv.toAlgebraIso {R : Type u} [CommRing R] {X₁ X₂ : Type u} {g₁ : Ring X₁} {g₂ : Ring X₂} {m₁ : Algebra R X₁} {m₂ : Algebra R X₂} (e : X₁ ≃ₐ[R] X₂) :
                    AlgCat.of R X₁ AlgCat.of R X₂

                    Build an isomorphism in the category AlgCat R from a AlgEquiv between Algebras.

                    Equations
                    Instances For
                      @[simp]
                      theorem AlgEquiv.toAlgebraIso_inv {R : Type u} [CommRing R] {X₁ X₂ : Type u} {g₁ : Ring X₁} {g₂ : Ring X₂} {m₁ : Algebra R X₁} {m₂ : Algebra R X₂} (e : X₁ ≃ₐ[R] X₂) :
                      @[simp]
                      theorem AlgEquiv.toAlgebraIso_hom {R : Type u} [CommRing R] {X₁ X₂ : Type u} {g₁ : Ring X₁} {g₂ : Ring X₂} {m₁ : Algebra R X₁} {m₂ : Algebra R X₂} (e : X₁ ≃ₐ[R] X₂) :
                      def CategoryTheory.Iso.toAlgEquiv {R : Type u} [CommRing R] {X Y : AlgCat R} (i : X Y) :
                      X ≃ₐ[R] Y

                      Build a AlgEquiv from an isomorphism in the category AlgCat R.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        @[simp]
                        theorem CategoryTheory.Iso.toAlgEquiv_symm_apply {R : Type u} [CommRing R] {X Y : AlgCat R} (i : X Y) (a : Y) :
                        @[simp]
                        theorem CategoryTheory.Iso.toAlgEquiv_apply {R : Type u} [CommRing R] {X Y : AlgCat R} (i : X Y) (a : X) :
                        def algEquivIsoAlgebraIso {R : Type u} [CommRing R] {X Y : Type u} [Ring X] [Ring Y] [Algebra R X] [Algebra R Y] :

                        Algebra equivalences between Algebras are the same as (isomorphic to) isomorphisms in AlgCat.

                        Equations
                        Instances For
                          @[simp]
                          theorem algEquivIsoAlgebraIso_inv {R : Type u} [CommRing R] {X Y : Type u} [Ring X] [Ring Y] [Algebra R X] [Algebra R Y] (i : AlgCat.of R X AlgCat.of R Y) :
                          @[simp]
                          theorem algEquivIsoAlgebraIso_hom {R : Type u} [CommRing R] {X Y : Type u} [Ring X] [Ring Y] [Algebra R X] [Algebra R Y] (e : X ≃ₐ[R] Y) :