Congruence relations and ring homomorphisms #
This file contains elementary definitions involving congruence relations and morphisms for rings and semirings
Main definitions #
RingCon.ker: the kernel of a monoid homomorphism as a congruence relationRingCon.lift,RingCon.liftₐ: the homomorphism / the algebra morphism on the quotient given that the congruence is in the kernelRingCon.map,RingCon.mapₐ: homomorphism / algebra morphism from a smaller to a larger quotientRingCon.quotientKerEquivRangeS,RingCon.quotientKerEquivRange,RingCon.quotientKerEquivRangeₐ: the first isomorphism theorem for semirings (usingRingHom.rangeS), rings (usingRingHom.range) and algebras (usingAlgHom.range).RingCon.comapQuotientEquivRangeS,RingCon.comapQuotientEquivRange,RingCon.comapQuotientEquivRangeₐ: the second isomorphism theorem for semirings (usingRingHom.rangeS), rings (usingRingHom.range) and algebras (usingAlgHom.range).RingCon.quotientQuotientEquivQuotient,RingCon.quotientQuotientEquivQuotientₐ: the third isomorphism theorem for semirings (or rings) and algebras
Tags #
congruence, congruence relation, quotient, quotient by congruence relation, ring, quotient ring
The kernel of a ring homomorphism as a ring congruence relation.
Equations
- RingCon.ker f = ⊥.comap f
Instances For
The definition of the ring congruence relation defined by a ring homomorphism's kernel.
The kernel of the quotient map induced by a ring congruence relation c equals c.
Makes an isomorphism of quotients by two ring congruence relations, given that the relations are equal.
Equations
- RingCon.congr h = { toEquiv := Quotient.congr (Equiv.refl M) ⋯, map_mul' := ⋯, map_add' := ⋯ }
Instances For
Given a function f, the smallest ring congruence relation containing the binary
relation on f's image defined by 'x ≈ y iff the elements of f⁻¹(x) are related to
the elements of f⁻¹(y) by a ring congruence relation c.'
Equations
- RingCon.mapGen f = ringConGen (Relation.Map (⇑c) f f)
Instances For
If c is a ring congruence on M, then the smallest ring
congruence relation on N deduced from c by a ring homomorphism
from M to N is the relation deduced from c.
Given a ring congruence relation c on a semiring M, the order-preserving
bijection between the set of ring congruence relations containing c and the
ring congruence relations on the quotient of M by c.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The homomorphism on the quotient of a ring by a congruence relation c
induced by a homomorphism constant on the equivalence classes of c.
Equations
Instances For
The diagram describing the universal property for quotients of ring commutes.
The diagram describing the universal property for quotients of rings commutes.
The diagram describing the universal property for quotients of rings commutes.
Given a homomorphism f from the quotient of a ring by a ring congruence
relation, f equals the homomorphism on the quotient induced by f composed
with the natural map from the ring to the quotient.
Homomorphisms on the quotient of a ring by a ring congruence relation are equal if they are equal on elements that are coercions from the ring.
Surjective ring homomorphisms constant on a the equivalence classes of a ring congruence relation induce a surjective homomorphism on the quotient.
Surjective ring homomorphisms constant on a the equivalence classes of a ring congruence relation induce a surjective homomorphism on the quotient.
Given a ring homomorphism f from M to P whose kernel contains c,
the lift of Mto N is injective iff ker f = c.
Given a ring homomorphism f from M to P, the kernel of f is the
unique ring congruence relation on M whose induced map from the quotient of
M to P is injective.
The homomorphism induced on the quotient of a ring by the kernel of a ring homomorphism.
Equations
- RingCon.kerLift f = (RingCon.ker f).lift f ⋯
Instances For
The diagram described by the universal property for quotients of rings, when the ring congruence relation is the kernel of the homomorphism, commutes.
A ring homomorphism f induces an injective homomorphism on the quotient by f's kernel.
Given ring congruence relations c, d on a ring such that d contains c,
d's quotient map induces a homomorphism from the quotient by c to the
quotient by d.
Instances For
Given ring congruence relations c, d on a ring such that d contains c,
the definition of the homomorphism from the quotient by c to the quotient by
d induced by d's quotient map.
Given a congruence relation c on a semiring and a homomorphism
f constant on the equivalence classes of c, f has the same image
as the homomorphism that f induces on the quotient.
Given a ring homomorphism f, the induced homomorphism
on the quotient by f's kernel has the same image as f.
The first isomorphism theorem for semirings.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The first isomorphism theorem for semirings in the case of a homomorphism with right inverse.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The first isomorphism theorem for rings in the case of a surjective homomorphism.
For a computable version, see RingCon.quotientKerEquivOfRightInverse.
Equations
Instances For
A surjective ring homomorphismsf : M →+* N induces
a ring equivalence d.Quotient ≃+* c.Quotient,
whenever c : RingCon M and d : RingCon N are such that d = c.comap f.
Equations
- c.comapQuotientEquivOfSurj f hf hcd = (RingCon.congr ⋯).trans (RingCon.quotientKerEquivOfSurjective (c.mk'.comp f) ⋯)
Instances For
This version infers the surjectivity of the function from a RingEquiv function
The second isomorphism theorem for semirings.
Equations
- c.comapQuotientEquivRangeS f hcd = (RingCon.congr ⋯).trans (RingCon.quotientKerEquivRangeS (c.mk'.comp f))
Instances For
The third isomorphism theorem for (semi-)rings.
Equations
- c.quotientQuotientEquivQuotient d h = { toEquiv := c.quotientQuotientEquivQuotient d.toSetoid h, map_mul' := ⋯, map_add' := ⋯ }
Instances For
Given a congruence relation c on a ring and a homomorphism f
constant on the equivalence classes of c, f has the same image
as the homomorphism that f induces on the quotient.
The second isomorphism theorem for rings.
Equations
- c.comapQuotientEquivRange f hcd = c.comapQuotientEquivRangeS f hcd
Instances For
Makes an algebra isomorphism of quotients by two ring congruence relations, given that the relations are equal.
Equations
- RingCon.congrₐ R h = { toEquiv := (RingCon.congr h).toEquiv, map_mul' := ⋯, map_add' := ⋯, commutes' := ⋯ }
Instances For
The algebra homomorphism on the quotient of an algebra by a
congruence relation c induced by an algebra homomorphism
constant on the equivalence classes of c.
Instances For
Given a congruence relation c on an algebra and a homomorphism f
constant on the equivalence classes of c, f has the same image
as the homomorphism that f induces on the quotient.
The homomorphism induced on the quotient of a ring by the kernel of a ring homomorphism.
Equations
- RingCon.kerLiftₐ f = (RingCon.ker f.toRingHom).liftₐ f ⋯
Instances For
The diagram described by the universal property for quotients of rings, when the ring congruence relation is the kernel of the homomorphism, commutes.
Given ring congruence relations c, d on an algebra such that d
contains c, d's quotient map induces an algebra homomorphism from
the quotient by c to the quotient by d.
Equations
- RingCon.factorₐ R h = c.liftₐ (RingCon.mkₐ R d) ⋯
Instances For
Given ring congruence relations c, d on a ring such that d contains c,
the definition of the homomorphism from the quotient by c to the quotient by
d induced by d's quotient map.
Given ring congruence relations c, d on a ring such that d contains c,
the definition of the homomorphism from the quotient by c to the quotient by
d induced by d's quotient map.
The first isomorphism theorem for algebra morphisms.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The second isomorphism theorem for algebras.
Equations
- c.comapQuotientEquivRangeₐ f h = (RingCon.congrₐ R ⋯).trans (RingCon.quotientKerEquivRangeₐ ((RingCon.mkₐ R c).comp f))
Instances For
The third isomorphism theorem for algebras.
Equations
- RingCon.quotientQuotientEquivQuotientₐ R h = { toEquiv := (c.quotientQuotientEquivQuotient d h).toEquiv, map_mul' := ⋯, map_add' := ⋯, commutes' := ⋯ }