Documentation

Mathlib.Algebra.Category.Ring.Under.Limits

Limits in Under R for a commutative ring R #

We show that Under.pushout f is left-exact, i.e. preserves finite limits, if f : R ⟶ S is flat.

The canonical fan on P : ι → Under R given by ∀ i, P i.

Equations
  • One or more equations did not get rendered due to their size.
Instances For

    The canonical fan is limiting.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      def CommRingCat.Under.tensorProductFan {R : CommRingCat} (S : CommRingCat) [Algebra R S] {ι : Type u} (P : ιCategoryTheory.Under R) :
      CategoryTheory.Limits.Fan fun (i : ι) => S.mkUnder (TensorProduct R S (P i).right)

      The fan on i ↦ S ⊗[R] P i given by S ⊗[R] ∀ i, P i

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        def CommRingCat.Under.tensorProductFan' {R : CommRingCat} (S : CommRingCat) [Algebra R S] {ι : Type u} (P : ιCategoryTheory.Under R) :
        CategoryTheory.Limits.Fan fun (i : ι) => S.mkUnder (TensorProduct R S (P i).right)

        The fan on i ↦ S ⊗[R] P i given by ∀ i, S ⊗[R] P i

        Equations
        • One or more equations did not get rendered due to their size.
        Instances For

          The two fans on i ↦ S ⊗[R] P i agree if ι is finite.

          Equations
          Instances For

            The fan on i ↦ S ⊗[R] P i given by S ⊗[R] ∀ i, P i is limiting if ι is finite.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              noncomputable def CommRingCat.Under.piFanTensorProductIsLimit {R S : CommRingCat} [Algebra R S] {ι : Type u} (P : ιCategoryTheory.Under R) [Finite ι] :

              tensorProd R S preserves the limit of the canonical fan on P.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For

                The canonical fork on f g : A ⟶ B given by the equalizer.

                Equations
                Instances For
                  def CommRingCat.Under.equalizerFork' {R : CommRingCat} {A B : Type u} [CommRing A] [CommRing B] [Algebra (↑R) A] [Algebra (↑R) B] (f g : A →ₐ[R] B) :
                  CategoryTheory.Limits.Fork f.toUnder g.toUnder

                  Variant of Under.equalizerFork' for algebra maps. This is definitionally equal to Under.equalizerFork but this is costly in applications.

                  Equations
                  Instances For
                    @[simp]
                    theorem CommRingCat.Under.equalizerFork'_ι {R : CommRingCat} {A B : Type u} [CommRing A] [CommRing B] [Algebra (↑R) A] [Algebra (↑R) B] (f g : A →ₐ[R] B) :

                    The canonical fork on f g : A ⟶ B is limiting.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      def CommRingCat.Under.tensorProdEqualizer {R S : CommRingCat} [Algebra R S] {A B : CategoryTheory.Under R} (f g : A B) :
                      CategoryTheory.Limits.Fork ((R.tensorProd S).map f) ((R.tensorProd S).map g)

                      The fork on 𝟙 ⊗[R] f and 𝟙 ⊗[R] g given by S ⊗[R] eq(f, g).

                      Equations
                      Instances For
                        noncomputable def CommRingCat.Under.tensorProdMapEqualizerForkIsLimit {R S : CommRingCat} [Algebra R S] [Module.Flat R S] {A B : CategoryTheory.Under R} (f g : A B) :

                        If S is R-flat, tensorProd R S preserves the equalizer of f and g.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For

                          Under.pushout f preserves finite limits if f is flat.