The category of commutative rings has all colimits. #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file uses a "pre-automated" approach, just as for Mon/colimits.lean
.
It is a very uniform approach, that conceivably could be synthesised directly
by a tactic that analyses the shape of comm_ring
and ring_hom
.
We build the colimit of a diagram in CommRing
by constructing the
free commutative ring on the disjoint union of all the commutative rings in the diagram,
then taking the quotient by the commutative ring laws within each commutative ring,
and the identifications given by the morphisms in the diagram.
- of : Π {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (j : J), ↥(F.obj j) → CommRing.colimits.prequotient F
- zero : Π {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing}, CommRing.colimits.prequotient F
- one : Π {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing}, CommRing.colimits.prequotient F
- neg : Π {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing}, CommRing.colimits.prequotient F → CommRing.colimits.prequotient F
- add : Π {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing}, CommRing.colimits.prequotient F → CommRing.colimits.prequotient F → CommRing.colimits.prequotient F
- mul : Π {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing}, CommRing.colimits.prequotient F → CommRing.colimits.prequotient F → CommRing.colimits.prequotient F
An inductive type representing all commutative ring expressions (without relations)
on a collection of types indexed by the objects of J
.
Instances for CommRing.colimits.prequotient
- CommRing.colimits.prequotient.has_sizeof_inst
- CommRing.colimits.prequotient.inhabited
- CommRing.colimits.colimit_setoid
Equations
- refl : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x : CommRing.colimits.prequotient F), CommRing.colimits.relation F x x
- symm : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y : CommRing.colimits.prequotient F), CommRing.colimits.relation F x y → CommRing.colimits.relation F y x
- trans : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y z : CommRing.colimits.prequotient F), CommRing.colimits.relation F x y → CommRing.colimits.relation F y z → CommRing.colimits.relation F x z
- map : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (j j' : J) (f : j ⟶ j') (x : ↥(F.obj j)), CommRing.colimits.relation F (CommRing.colimits.prequotient.of j' (⇑(F.map f) x)) (CommRing.colimits.prequotient.of j x)
- zero : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (j : J), CommRing.colimits.relation F (CommRing.colimits.prequotient.of j 0) CommRing.colimits.prequotient.zero
- one : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (j : J), CommRing.colimits.relation F (CommRing.colimits.prequotient.of j 1) CommRing.colimits.prequotient.one
- neg : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (j : J) (x : ↥(F.obj j)), CommRing.colimits.relation F (CommRing.colimits.prequotient.of j (-x)) (CommRing.colimits.prequotient.of j x).neg
- add : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (j : J) (x y : ↥(F.obj j)), CommRing.colimits.relation F (CommRing.colimits.prequotient.of j (x + y)) ((CommRing.colimits.prequotient.of j x).add (CommRing.colimits.prequotient.of j y))
- mul : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (j : J) (x y : ↥(F.obj j)), CommRing.colimits.relation F (CommRing.colimits.prequotient.of j (x * y)) ((CommRing.colimits.prequotient.of j x).mul (CommRing.colimits.prequotient.of j y))
- neg_1 : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x x' : CommRing.colimits.prequotient F), CommRing.colimits.relation F x x' → CommRing.colimits.relation F x.neg x'.neg
- add_1 : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x x' y : CommRing.colimits.prequotient F), CommRing.colimits.relation F x x' → CommRing.colimits.relation F (x.add y) (x'.add y)
- add_2 : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y y' : CommRing.colimits.prequotient F), CommRing.colimits.relation F y y' → CommRing.colimits.relation F (x.add y) (x.add y')
- mul_1 : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x x' y : CommRing.colimits.prequotient F), CommRing.colimits.relation F x x' → CommRing.colimits.relation F (x.mul y) (x'.mul y)
- mul_2 : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y y' : CommRing.colimits.prequotient F), CommRing.colimits.relation F y y' → CommRing.colimits.relation F (x.mul y) (x.mul y')
- zero_add : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x : CommRing.colimits.prequotient F), CommRing.colimits.relation F (CommRing.colimits.prequotient.zero.add x) x
- add_zero : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x : CommRing.colimits.prequotient F), CommRing.colimits.relation F (x.add CommRing.colimits.prequotient.zero) x
- one_mul : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x : CommRing.colimits.prequotient F), CommRing.colimits.relation F (CommRing.colimits.prequotient.one.mul x) x
- mul_one : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x : CommRing.colimits.prequotient F), CommRing.colimits.relation F (x.mul CommRing.colimits.prequotient.one) x
- add_left_neg : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x : CommRing.colimits.prequotient F), CommRing.colimits.relation F (x.neg.add x) CommRing.colimits.prequotient.zero
- add_comm : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y : CommRing.colimits.prequotient F), CommRing.colimits.relation F (x.add y) (y.add x)
- mul_comm : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y : CommRing.colimits.prequotient F), CommRing.colimits.relation F (x.mul y) (y.mul x)
- add_assoc : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y z : CommRing.colimits.prequotient F), CommRing.colimits.relation F ((x.add y).add z) (x.add (y.add z))
- mul_assoc : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y z : CommRing.colimits.prequotient F), CommRing.colimits.relation F ((x.mul y).mul z) (x.mul (y.mul z))
- left_distrib : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y z : CommRing.colimits.prequotient F), CommRing.colimits.relation F (x.mul (y.add z)) ((x.mul y).add (x.mul z))
- right_distrib : ∀ {J : Type v} [_inst_1 : category_theory.small_category J] {F : J ⥤ CommRing} (x y z : CommRing.colimits.prequotient F), CommRing.colimits.relation F ((x.add y).mul z) ((x.mul z).add (y.mul z))
The relation on prequotient
saying when two expressions are equal
because of the commutative ring laws, or
because one element is mapped to another by a morphism in the diagram.
The setoid corresponding to commutative expressions modulo monoid relations and identifications.
Equations
- CommRing.colimits.colimit_setoid F = {r := CommRing.colimits.relation F, iseqv := _}
The underlying type of the colimit of a diagram in CommRing
.
Equations
Instances for CommRing.colimits.colimit_type
Equations
- CommRing.colimits.colimit_type.add_group F = {add := quot.lift (λ (x : CommRing.colimits.prequotient F), quot.lift (λ (y : CommRing.colimits.prequotient F), quot.mk setoid.r (x.add y)) _) _, add_assoc := _, zero := quot.mk setoid.r CommRing.colimits.prequotient.zero, zero_add := _, add_zero := _, nsmul := sub_neg_monoid.nsmul._default (quot.mk setoid.r CommRing.colimits.prequotient.zero) (quot.lift (λ (x : CommRing.colimits.prequotient F), quot.lift (λ (y : CommRing.colimits.prequotient F), quot.mk setoid.r (x.add y)) _) _) _ _, nsmul_zero' := _, nsmul_succ' := _, neg := quot.lift (λ (x : CommRing.colimits.prequotient F), quot.mk setoid.r x.neg) _, sub := sub_neg_monoid.sub._default (quot.lift (λ (x : CommRing.colimits.prequotient F), quot.lift (λ (y : CommRing.colimits.prequotient F), quot.mk setoid.r (x.add y)) _) _) _ (quot.mk setoid.r CommRing.colimits.prequotient.zero) _ _ (sub_neg_monoid.nsmul._default (quot.mk setoid.r CommRing.colimits.prequotient.zero) (quot.lift (λ (x : CommRing.colimits.prequotient F), quot.lift (λ (y : CommRing.colimits.prequotient F), quot.mk setoid.r (x.add y)) _) _) _ _) _ _ (quot.lift (λ (x : CommRing.colimits.prequotient F), quot.mk setoid.r x.neg) _), sub_eq_add_neg := _, zsmul := sub_neg_monoid.zsmul._default (quot.lift (λ (x : CommRing.colimits.prequotient F), quot.lift (λ (y : CommRing.colimits.prequotient F), quot.mk setoid.r (x.add y)) _) _) _ (quot.mk setoid.r CommRing.colimits.prequotient.zero) _ _ (sub_neg_monoid.nsmul._default (quot.mk setoid.r CommRing.colimits.prequotient.zero) (quot.lift (λ (x : CommRing.colimits.prequotient F), quot.lift (λ (y : CommRing.colimits.prequotient F), quot.mk setoid.r (x.add y)) _) _) _ _) _ _ (quot.lift (λ (x : CommRing.colimits.prequotient F), quot.mk setoid.r x.neg) _), zsmul_zero' := _, zsmul_succ' := _, zsmul_neg' := _, add_left_neg := _}
Equations
- CommRing.colimits.colimit_type.add_group_with_one F = {int_cast := int.cast_def (sub_neg_monoid.to_has_neg (CommRing.colimits.colimit_type F)), add := add_group.add (CommRing.colimits.colimit_type.add_group F), add_assoc := _, zero := add_group.zero (CommRing.colimits.colimit_type.add_group F), zero_add := _, add_zero := _, nsmul := add_group.nsmul (CommRing.colimits.colimit_type.add_group F), nsmul_zero' := _, nsmul_succ' := _, neg := add_group.neg (CommRing.colimits.colimit_type.add_group F), sub := add_group.sub (CommRing.colimits.colimit_type.add_group F), sub_eq_add_neg := _, zsmul := add_group.zsmul (CommRing.colimits.colimit_type.add_group F), zsmul_zero' := _, zsmul_succ' := _, zsmul_neg' := _, add_left_neg := _, nat_cast := add_monoid_with_one.nat_cast._default (quot.mk setoid.r CommRing.colimits.prequotient.one) add_group.add _ add_group.zero _ _ add_group.nsmul _ _, one := quot.mk setoid.r CommRing.colimits.prequotient.one, nat_cast_zero := _, nat_cast_succ := _, int_cast_of_nat := _, int_cast_neg_succ_of_nat := _}
Equations
- CommRing.colimits.colimit_type.comm_ring F = {add := add_group_with_one.add (CommRing.colimits.colimit_type.add_group_with_one F), add_assoc := _, zero := add_group_with_one.zero (CommRing.colimits.colimit_type.add_group_with_one F), zero_add := _, add_zero := _, nsmul := add_group_with_one.nsmul (CommRing.colimits.colimit_type.add_group_with_one F), nsmul_zero' := _, nsmul_succ' := _, neg := add_group_with_one.neg (CommRing.colimits.colimit_type.add_group_with_one F), sub := add_group_with_one.sub (CommRing.colimits.colimit_type.add_group_with_one F), sub_eq_add_neg := _, zsmul := add_group_with_one.zsmul (CommRing.colimits.colimit_type.add_group_with_one F), zsmul_zero' := _, zsmul_succ' := _, zsmul_neg' := _, add_left_neg := _, add_comm := _, int_cast := add_group_with_one.int_cast (CommRing.colimits.colimit_type.add_group_with_one F), nat_cast := add_group_with_one.nat_cast (CommRing.colimits.colimit_type.add_group_with_one F), one := add_group_with_one.one (CommRing.colimits.colimit_type.add_group_with_one F), nat_cast_zero := _, nat_cast_succ := _, int_cast_of_nat := _, int_cast_neg_succ_of_nat := _, mul := quot.lift (λ (x : CommRing.colimits.prequotient F), quot.lift (λ (y : CommRing.colimits.prequotient F), quot.mk setoid.r (x.mul y)) _) _, mul_assoc := _, one_mul := _, mul_one := _, npow := ring.npow._default add_group_with_one.one (quot.lift (λ (x : CommRing.colimits.prequotient F), quot.lift (λ (y : CommRing.colimits.prequotient F), quot.mk setoid.r (x.mul y)) _) _) _ _, npow_zero' := _, npow_succ' := _, left_distrib := _, right_distrib := _, mul_comm := _}
The bundled commutative ring giving the colimit of a diagram.
Equations
The function from a given commutative ring in the diagram to the colimit commutative ring.
Equations
- CommRing.colimits.cocone_fun F j x = quot.mk setoid.r (CommRing.colimits.prequotient.of j x)
The ring homomorphism from a given commutative ring in the diagram to the colimit commutative ring.
Equations
- CommRing.colimits.cocone_morphism F j = {to_fun := CommRing.colimits.cocone_fun F j, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}
The cocone over the proposed colimit commutative ring.
Equations
- CommRing.colimits.colimit_cocone F = {X := CommRing.colimits.colimit F, ι := {app := CommRing.colimits.cocone_morphism F, naturality' := _}}
The function from the free commutative ring on the diagram to the cone point of any other cocone.
Equations
- CommRing.colimits.desc_fun_lift F s (x.mul y) = CommRing.colimits.desc_fun_lift F s x * CommRing.colimits.desc_fun_lift F s y
- CommRing.colimits.desc_fun_lift F s (x.add y) = CommRing.colimits.desc_fun_lift F s x + CommRing.colimits.desc_fun_lift F s y
- CommRing.colimits.desc_fun_lift F s x.neg = -CommRing.colimits.desc_fun_lift F s x
- CommRing.colimits.desc_fun_lift F s CommRing.colimits.prequotient.one = 1
- CommRing.colimits.desc_fun_lift F s CommRing.colimits.prequotient.zero = 0
- CommRing.colimits.desc_fun_lift F s (CommRing.colimits.prequotient.of j x) = ⇑(s.ι.app j) x
The function from the colimit commutative ring to the cone point of any other cocone.
Equations
- CommRing.colimits.desc_fun F s = quot.lift (CommRing.colimits.desc_fun_lift F s) _
The ring homomorphism from the colimit commutative ring to the cone point of any other cocone.
Equations
- CommRing.colimits.desc_morphism F s = {to_fun := CommRing.colimits.desc_fun F s, map_one' := _, map_mul' := _, map_zero' := _, map_add' := _}
Evidence that the proposed colimit is the colimit.
Equations
- CommRing.colimits.colimit_is_colimit F = {desc := λ (s : category_theory.limits.cocone F), CommRing.colimits.desc_morphism F s, fac' := _, uniq' := _}