mathlib3 documentation

algebra.ring_quot

Quotients of non-commutative rings #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

Unfortunately, ideals have only been developed in the commutative case as ideal, and it's not immediately clear how one should formalise ideals in the non-commutative case.

In this file, we directly define the quotient of a semiring by any relation, by building a bigger relation that represents the ideal generated by that relation.

We prove the universal properties of the quotient, and recommend avoiding relying on the actual definition, which is made irreducible for this purpose.

Since everything runs in parallel for quotients of R-algebras, we do that case at the same time.

@[protected, instance]
def ring_con.quotient.algebra {S : Type u₂} [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] (c : ring_con A) :
Equations
@[simp, norm_cast]
theorem ring_con.coe_algebra_map {S : Type u₂} [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] (c : ring_con A) (s : S) :
inductive ring_quot.rel {R : Type u₁} [semiring R] (r : R R Prop) :
R R Prop

Given an arbitrary relation r on a ring, we strengthen it to a relation rel r, such that the equivalence relation generated by rel r has x ~ y if and only if x - y is in the ideal generated by elements a - b such that r a b.

theorem ring_quot.rel.add_right {R : Type u₁} [semiring R] {r : R R Prop} ⦃a b c : R⦄ (h : ring_quot.rel r b c) :
ring_quot.rel r (a + b) (a + c)
theorem ring_quot.rel.neg {R : Type u₁} [ring R] {r : R R Prop} ⦃a b : R⦄ (h : ring_quot.rel r a b) :
ring_quot.rel r (-a) (-b)
theorem ring_quot.rel.sub_left {R : Type u₁} [ring R] {r : R R Prop} ⦃a b c : R⦄ (h : ring_quot.rel r a b) :
ring_quot.rel r (a - c) (b - c)
theorem ring_quot.rel.sub_right {R : Type u₁} [ring R] {r : R R Prop} ⦃a b c : R⦄ (h : ring_quot.rel r b c) :
ring_quot.rel r (a - b) (a - c)
theorem ring_quot.rel.smul {S : Type u₂} [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] {r : A A Prop} (k : S) ⦃a b : A⦄ (h : ring_quot.rel r a b) :
ring_quot.rel r (k a) (k b)
def ring_quot.ring_con {R : Type u₁} [semiring R] (r : R R Prop) :

eqv_gen (ring_quot.rel r) is a ring congruence.

Equations
theorem ring_quot.eqv_gen_rel_eq {R : Type u₁} [semiring R] (r : R R Prop) :
@[protected, instance]
def ring_quot.has_zero {R : Type u₁} [semiring R] (r : R R Prop) :
Equations
@[protected, instance]
def ring_quot.has_one {R : Type u₁} [semiring R] (r : R R Prop) :
Equations
@[protected, instance]
def ring_quot.has_add {R : Type u₁} [semiring R] (r : R R Prop) :
Equations
@[protected, instance]
def ring_quot.has_mul {R : Type u₁} [semiring R] (r : R R Prop) :
Equations
@[protected, instance]
def ring_quot.nat.has_pow {R : Type u₁} [semiring R] (r : R R Prop) :
Equations
@[protected, instance]
def ring_quot.has_neg {R : Type u₁} [ring R] (r : R R Prop) :
Equations
@[protected, instance]
def ring_quot.has_sub {R : Type u₁} [ring R] (r : R R Prop) :
Equations
@[protected, instance]
def ring_quot.has_smul {R : Type u₁} [semiring R] {S : Type u₂} [comm_semiring S] (r : R R Prop) [algebra S R] :
Equations
theorem ring_quot.zero_quot {R : Type u₁} [semiring R] (r : R R Prop) :
{to_quot := quot.mk (ring_quot.rel r) 0} = 0
theorem ring_quot.one_quot {R : Type u₁} [semiring R] (r : R R Prop) :
{to_quot := quot.mk (ring_quot.rel r) 1} = 1
theorem ring_quot.add_quot {R : Type u₁} [semiring R] (r : R R Prop) {a b : R} :
{to_quot := quot.mk (ring_quot.rel r) a} + {to_quot := quot.mk (ring_quot.rel r) b} = {to_quot := quot.mk (ring_quot.rel r) (a + b)}
theorem ring_quot.mul_quot {R : Type u₁} [semiring R] (r : R R Prop) {a b : R} :
{to_quot := quot.mk (ring_quot.rel r) a} * {to_quot := quot.mk (ring_quot.rel r) b} = {to_quot := quot.mk (ring_quot.rel r) (a * b)}
theorem ring_quot.pow_quot {R : Type u₁} [semiring R] (r : R R Prop) {a : R} {n : } :
{to_quot := quot.mk (ring_quot.rel r) a} ^ n = {to_quot := quot.mk (ring_quot.rel r) (a ^ n)}
theorem ring_quot.neg_quot {R : Type u₁} [ring R] (r : R R Prop) {a : R} :
-{to_quot := quot.mk (ring_quot.rel r) a} = {to_quot := quot.mk (ring_quot.rel r) (-a)}
theorem ring_quot.sub_quot {R : Type u₁} [ring R] (r : R R Prop) {a b : R} :
{to_quot := quot.mk (ring_quot.rel r) a} - {to_quot := quot.mk (ring_quot.rel r) b} = {to_quot := quot.mk (ring_quot.rel r) (a - b)}
theorem ring_quot.smul_quot {R : Type u₁} [semiring R] {S : Type u₂} [comm_semiring S] (r : R R Prop) [algebra S R] {n : S} {a : R} :
n {to_quot := quot.mk (ring_quot.rel r) a} = {to_quot := quot.mk (ring_quot.rel r) (n a)}
@[protected, instance]
def ring_quot.semiring {R : Type u₁} [semiring R] (r : R R Prop) :
Equations
@[protected, instance]
def ring_quot.inhabited {R : Type u₁} [semiring R] (r : R R Prop) :
Equations
@[protected, instance]
def ring_quot.algebra {R : Type u₁} [semiring R] {S : Type u₂} [comm_semiring S] [algebra S R] (r : R R Prop) :
Equations
@[irreducible]
def ring_quot.mk_ring_hom {R : Type u₁} [semiring R] (r : R R Prop) :

The quotient map from a ring to its quotient, as a homomorphism of rings.

Equations
theorem ring_quot.mk_ring_hom_rel {R : Type u₁} [semiring R] {r : R R Prop} {x y : R} (w : r x y) :
@[ext]
theorem ring_quot.ring_quot_ext {R : Type u₁} [semiring R] {T : Type u₄} [semiring T] {r : R R Prop} (f g : ring_quot r →+* T) (w : f.comp (ring_quot.mk_ring_hom r) = g.comp (ring_quot.mk_ring_hom r)) :
f = g
@[irreducible]
def ring_quot.lift {R : Type u₁} [semiring R] {T : Type u₄} [semiring T] {r : R R Prop} :
{f // ⦃x y : R⦄, r x y f x = f y} (ring_quot r →+* T)

Any ring homomorphism f : R →+* T which respects a relation r : R → R → Prop factors uniquely through a morphism ring_quot r →+* T.

Equations
@[simp]
theorem ring_quot.lift_mk_ring_hom_apply {R : Type u₁} [semiring R] {T : Type u₄} [semiring T] (f : R →+* T) {r : R R Prop} (w : ⦃x y : R⦄, r x y f x = f y) (x : R) :
theorem ring_quot.lift_unique {R : Type u₁} [semiring R] {T : Type u₄} [semiring T] (f : R →+* T) {r : R R Prop} (w : ⦃x y : R⦄, r x y f x = f y) (g : ring_quot r →+* T) (h : g.comp (ring_quot.mk_ring_hom r) = f) :
theorem ring_quot.eq_lift_comp_mk_ring_hom {R : Type u₁} [semiring R] {T : Type u₄} [semiring T] {r : R R Prop} (f : ring_quot r →+* T) :

We now verify that in the case of a commutative ring, the ring_quot construction agrees with the quotient by the appropriate ideal.

The universal ring homomorphism from ring_quot r to B ⧸ ideal.of_rel r.

Equations
theorem ring_quot.rel.star {R : Type u₁} [semiring R] (r : R R Prop) [star_ring R] (hr : (a b : R), r a b r (has_star.star a) (has_star.star b)) ⦃a b : R⦄ (h : ring_quot.rel r a b) :
theorem ring_quot.star'_quot {R : Type u₁} [semiring R] (r : R R Prop) [star_ring R] (hr : (a b : R), r a b r (has_star.star a) (has_star.star b)) {a : R} :
star' r hr {to_quot := quot.mk (ring_quot.rel r) a} = {to_quot := quot.mk (ring_quot.rel r) (has_star.star a)}
def ring_quot.star_ring {R : Type u₁} [semiring R] [star_ring R] (r : R R Prop) (hr : (a b : R), r a b r (has_star.star a) (has_star.star b)) :

Transfer a star_ring instance through a quotient, if the quotient is invariant to star

Equations
@[irreducible]
def ring_quot.mk_alg_hom (S : Type u₂) [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] (s : A A Prop) :

The quotient map from an S-algebra to its quotient, as a homomorphism of S-algebras.

Equations
@[simp]
theorem ring_quot.mk_alg_hom_coe (S : Type u₂) [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] (s : A A Prop) :
theorem ring_quot.mk_alg_hom_rel (S : Type u₂) [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] {s : A A Prop} {x y : A} (w : s x y) :
@[ext]
theorem ring_quot.ring_quot_ext' (S : Type u₂) [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] {B : Type u₄} [semiring B] [algebra S B] {s : A A Prop} (f g : ring_quot s →ₐ[S] B) (w : f.comp (ring_quot.mk_alg_hom S s) = g.comp (ring_quot.mk_alg_hom S s)) :
f = g
@[irreducible]
def ring_quot.lift_alg_hom (S : Type u₂) [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] {B : Type u₄} [semiring B] [algebra S B] {s : A A Prop} :
{f // ⦃x y : A⦄, s x y f x = f y} (ring_quot s →ₐ[S] B)

Any S-algebra homomorphism f : A →ₐ[S] B which respects a relation s : A → A → Prop factors uniquely through a morphism ring_quot s →ₐ[S] B.

Equations
@[simp]
theorem ring_quot.lift_alg_hom_mk_alg_hom_apply (S : Type u₂) [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] {B : Type u₄} [semiring B] [algebra S B] (f : A →ₐ[S] B) {s : A A Prop} (w : ⦃x y : A⦄, s x y f x = f y) (x : A) :
theorem ring_quot.lift_alg_hom_unique (S : Type u₂) [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] {B : Type u₄} [semiring B] [algebra S B] (f : A →ₐ[S] B) {s : A A Prop} (w : ⦃x y : A⦄, s x y f x = f y) (g : ring_quot s →ₐ[S] B) (h : g.comp (ring_quot.mk_alg_hom S s) = f) :
theorem ring_quot.eq_lift_alg_hom_comp_mk_alg_hom (S : Type u₂) [comm_semiring S] {A : Type u₃} [semiring A] [algebra S A] {B : Type u₄} [semiring B] [algebra S B] {s : A A Prop} (f : ring_quot s →ₐ[S] B) :