Polynomial ring (#76)

This commit is contained in:
Patrick Stevens
2019-11-17 17:37:10 +00:00
committed by GitHub
parent c55dd5f63e
commit 8377c23613
23 changed files with 984 additions and 341 deletions

View File

@@ -14,49 +14,26 @@ open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
-- Following Part IB's course Groups, Rings, and Modules, we take rings to be commutative with one.
module Rings.Definition where
record Ring {n m} {A : Set n} (S : Setoid {n} {m} A) (_+_ : A A A) (_*_ : A A A) : Set (lsuc n m) where
field
additiveGroup : Group S _+_
open Group additiveGroup
open Setoid S
open Equivalence eq
0R : A
0R = 0G
_-R_ : A A A
a -R b = a + (inverse b)
field
*WellDefined : {r s t u : A} (r t) (s u) r * s t * u
1R : A
groupIsAbelian : {a b : A} a + b b + a
*Associative : {a b c : A} (a * (b * c)) (a * b) * c
*Commutative : {a b : A} a * b b * a
*DistributesOver+ : {a b c : A} a * (b + c) (a * b) + (a * c)
identIsIdent : {a : A} 1R * a a
timesZero : {a : A} a * 0R 0R
timesZero {a} = symmetric (transitive (transitive (symmetric invLeft) (+WellDefined reflexive (transitive (*WellDefined {a} {a} reflexive (symmetric identRight)) *DistributesOver+))) (transitive +Associative (transitive (+WellDefined invLeft reflexive) identLeft)))
--directSumRing : {m n : _} {A : Set m} {B : Set n} (r : Ring A) (s : Ring B) Ring (A && B)
--Ring.additiveGroup (directSumRing r s) = directSumGroup (Ring.additiveGroup r) (Ring.additiveGroup s)
--Ring._*_ (directSumRing r s) (a ,, b) (c ,, d) = (Ring._*_ r a c) ,, Ring._*_ s b d
--Ring.multWellDefined (directSumRing r s) (a ,, b) (c ,, d) = Ring.multWellDefined r a c ,, Ring.multWellDefined s b d
--Ring.1R (directSumRing r s) = Ring.1R r ,, Ring.1R s
--Ring.groupIsAbelian (directSumRing r s) = Ring.groupIsAbelian r ,, Ring.groupIsAbelian s
--Ring.assoc (directSumRing r s) = Ring.assoc r ,, Ring.assoc s
--Ring.multCommutative (directSumRing r s) = Ring.multCommutative r ,, Ring.multCommutative s
--Ring.multDistributes (directSumRing r s) = Ring.multDistributes r ,, Ring.multDistributes s
--Ring.identIsIdent (directSumRing r s) = Ring.identIsIdent r ,, Ring.identIsIdent s
record RingHom {m n o p : _} {A : Set m} {B : Set n} {SA : Setoid {m} {o} A} {SB : Setoid {n} {p} B} {_+A_ : A A A} {_*A_ : A A A} (R : Ring SA _+A_ _*A_) {_+B_ : B B B} {_*B_ : B B B} (S : Ring SB _+B_ _*B_) (f : A B) : Set (m n o p) where
open Ring S
open Group additiveGroup
open Setoid SB
field
preserves1 : f (Ring.1R R) 1R
ringHom : {r s : A} f (r *A s) (f r) *B (f s)
groupHom : GroupHom (Ring.additiveGroup R) additiveGroup f
--record RingIso {m n : _} {A : Set m} {B : Set n} (R : Ring A) (S : Ring B) (f : A → B) : Set (m ⊔ n) where
-- field
-- ringHom : RingHom R S f
-- bijective : Bijection f
record Ring {n m} {A : Set n} (S : Setoid {n} {m} A) (_+_ : A A A) (_*_ : A A A) : Set (lsuc n m) where
field
additiveGroup : Group S _+_
open Group additiveGroup
open Setoid S
open Equivalence eq
0R : A
0R = 0G
_-R_ : A A A
a -R b = a + (inverse b)
field
*WellDefined : {r s t u : A} (r t) (s u) r * s t * u
1R : A
groupIsAbelian : {a b : A} a + b b + a
*Associative : {a b c : A} (a * (b * c)) (a * b) * c
*Commutative : {a b : A} a * b b * a
*DistributesOver+ : {a b c : A} a * (b + c) (a * b) + (a * c)
identIsIdent : {a : A} 1R * a a
timesZero : {a : A} a * 0R 0R
timesZero {a} = symmetric (transitive (transitive (symmetric invLeft) (+WellDefined reflexive (transitive (*WellDefined {a} {a} reflexive (symmetric identRight)) *DistributesOver+))) (transitive +Associative (transitive (+WellDefined invLeft reflexive) identLeft)))
*DistributesOver+' : {a b c : A} (a + b) * c (a * c) + (b * c)
*DistributesOver+' = transitive *Commutative (transitive *DistributesOver+ (+WellDefined *Commutative *Commutative))

View File

@@ -0,0 +1,26 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.Homomorphisms.Definition
open import Groups.Definition
open import Numbers.Naturals.Naturals
open import Setoids.Orders
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
-- Following Part IB's course Groups, Rings, and Modules, we take rings to be commutative with one.
module Rings.Homomorphisms.Definition where
record RingHom {m n o p : _} {A : Set m} {B : Set n} {SA : Setoid {m} {o} A} {SB : Setoid {n} {p} B} {_+A_ : A A A} {_*A_ : A A A} (R : Ring SA _+A_ _*A_) {_+B_ : B B B} {_*B_ : B B B} (S : Ring SB _+B_ _*B_) (f : A B) : Set (m n o p) where
open Ring S
open Group additiveGroup
open Setoid SB
field
preserves1 : f (Ring.1R R) 1R
ringHom : {r s : A} f (r *A s) (f r) *B (f s)
groupHom : GroupHom (Ring.additiveGroup R) additiveGroup f

View File

@@ -0,0 +1,20 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.Homomorphisms.Definition
open import Groups.Definition
open import Numbers.Naturals.Naturals
open import Setoids.Orders
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Rings.Homomorphisms.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Ideals.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
ringKernel : {c d : _} {C : Set c} {T : Setoid {c} {d} C} {_+2_ _*2_ : C C C} (R2 : Ring T _+2_ _*2_) {f : A C} (fHom : RingHom R R2 f) Set (a d)
ringKernel {T = T} R2 {f} fHom = Sg A (λ a Setoid.__ T (f a) (Ring.0R R2))

View File

@@ -0,0 +1,22 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.Homomorphisms.Definition
open import Groups.Definition
open import Numbers.Naturals.Naturals
open import Setoids.Orders
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Rings.Homomorphisms.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Isomorphisms.Definition {a b c d : _} {A : Set a} {S : Setoid {a} {b} A} {_+1_ _*1_ : A A A} (R1 : Ring S _+1_ _*1_) {B : Set c} {T : Setoid {c} {d} B} {_+2_ _*2_ : B B B} (R2 : Ring T _+2_ _*2_) where
record RingIso (f : A B) : Set (a b c d) where
field
ringHom : RingHom R1 R2 f
bijective : Bijection f

View File

@@ -0,0 +1,33 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.Homomorphisms.Definition
open import Groups.Definition
open import Numbers.Naturals.Definition
open import Setoids.Orders
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Vectors
open import Lists.Lists
open import Maybe
open import Rings.Homomorphisms.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Polynomial.Definition {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
open Setoid S
open Equivalence eq
open Ring R
open import Groups.Polynomials.Definition additiveGroup
1P : NaivePoly
1P = 1R :: []
inducedFunction : NaivePoly A A
inducedFunction [] a = 0R
inducedFunction (x :: p) a = x + (a * inducedFunction p a)

View File

@@ -0,0 +1,123 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.Abelian.Definition
open import Groups.Homomorphisms.Definition
open import Groups.Definition
open import Numbers.Naturals.Definition
open import Setoids.Orders
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Vectors
open import Lists.Lists
open import Maybe
open import Rings.Homomorphisms.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Polynomial.Multiplication {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
open Ring R
open Group additiveGroup
open import Groups.Polynomials.Definition additiveGroup
open import Groups.Polynomials.Group additiveGroup
open Setoid S
open Equivalence eq
_*P_ : NaivePoly NaivePoly NaivePoly
[] *P b = []
(x :: a) *P [] = []
(x :: a) *P (y :: b) = (x * y) :: ((map (x *_) b) +P (map (y *_) a))
p*Commutative : {a b : NaivePoly} polysEqual (a *P b) (b *P a)
p*Commutative {[]} {[]} = record {}
p*Commutative {[]} {x :: b} = record {}
p*Commutative {x :: a} {[]} = record {}
p*Commutative {x :: xs} {y :: ys} = *Commutative ,, AbelianGroup.commutative (abelian (record { commutative = Ring.groupIsAbelian R })) {map (_*_ x) ys} {map (_*_ y) xs}
zeroTimes1 : {a : NaivePoly} (c : A) (c 0G) polysEqual (map (_*_ c) a) []
zeroTimes1 {[]} c c=0 = record {}
zeroTimes1 {x :: a} c c=0 = transitive (transitive *Commutative (*WellDefined reflexive c=0)) (timesZero {x}) ,, zeroTimes1 {a} c c=0
zeroTimes2 : {a : NaivePoly} (c : A) polysEqual a [] polysEqual (map (_*_ c) a) []
zeroTimes2 {[]} c a=0 = record {}
zeroTimes2 {x :: a} c (fst ,, snd) = transitive (*WellDefined reflexive fst) (timesZero {c}) ,, zeroTimes2 {a} c snd
mapWellDefined : (a c : A) (bs : NaivePoly) (a c) polysEqual (map (_*_ a) bs) (map (_*_ c) bs)
mapWellDefined a c [] a=c = record {}
mapWellDefined a c (x :: bs) a=c = *WellDefined a=c reflexive ,, mapWellDefined a c bs a=c
mapWellDefined' : (a : A) (bs cs : NaivePoly) polysEqual bs cs polysEqual (map (_*_ a) bs) (map (_*_ a) cs)
mapWellDefined' a [] [] bs=cs = record {}
mapWellDefined' a [] (x :: cs) (fst ,, snd) = transitive (*WellDefined reflexive fst) (timesZero {a}) ,, Equivalence.symmetric (Setoid.eq naivePolySetoid) (zeroTimes2 {cs} a (Equivalence.symmetric (Setoid.eq naivePolySetoid) snd))
mapWellDefined' a (x :: bs) [] (fst ,, snd) = transitive (*WellDefined reflexive fst) (timesZero {a}) ,, zeroTimes2 {bs} a snd
mapWellDefined' a (b :: bs) (c :: cs) (fst ,, snd) = *WellDefined reflexive fst ,, mapWellDefined' a bs cs snd
*PwellDefinedL : {a b c : NaivePoly} polysEqual a c polysEqual (a *P b) (c *P b)
*PwellDefinedL {[]} {[]} {[]} a=c = record {}
*PwellDefinedL {[]} {[]} {x :: c} a=c = record {}
*PwellDefinedL {[]} {x :: b} {[]} a=c = record {}
*PwellDefinedL {[]} {b :: bs} {c :: cs} (fst ,, snd) = transitive (transitive *Commutative (*WellDefined reflexive fst)) (timesZero {b}) ,, Equivalence.symmetric (Setoid.eq naivePolySetoid) (Equivalence.transitive (Setoid.eq naivePolySetoid) (+PwellDefined (zeroTimes1 {bs} c fst) (zeroTimes2 {cs} b (Equivalence.symmetric (Setoid.eq naivePolySetoid) snd))) (Group.identLeft polyGroup))
*PwellDefinedL {a :: as} {[]} {[]} a=c = record {}
*PwellDefinedL {a :: as} {[]} {x :: c} a=c = record {}
*PwellDefinedL {a :: as} {b :: bs} {[]} (fst ,, snd) = transitive (transitive *Commutative (*WellDefined reflexive fst)) (timesZero {b}) ,, Equivalence.transitive (Setoid.eq naivePolySetoid) (+PwellDefined (zeroTimes1 {bs} a fst) (zeroTimes2 {as} b snd)) (Group.identLeft polyGroup)
*PwellDefinedL {a :: as} {b :: bs} {c :: cs} (fst ,, snd) = *WellDefined fst reflexive ,, +PwellDefined (mapWellDefined a c bs fst) (mapWellDefined' b as cs snd)
*PwellDefinedR : {a b c : NaivePoly} polysEqual b c polysEqual (a *P b) (a *P c)
*PwellDefinedR {a} {b} {c} b=c = Equivalence.transitive (Setoid.eq naivePolySetoid) (p*Commutative {a} {b}) (Equivalence.transitive (Setoid.eq naivePolySetoid) (*PwellDefinedL b=c) (p*Commutative {c} {a}))
*PwellDefined : {a b c d : NaivePoly} polysEqual a c polysEqual b d polysEqual (a *P b) (c *P d)
*PwellDefined {a}{b}{c}{d} a=c b=d = Equivalence.transitive (Setoid.eq naivePolySetoid) (*PwellDefinedL a=c) (*PwellDefinedR b=d)
private
*1 : (a : NaivePoly) polysEqual (map (_*_ 1R) a) a
*1 [] = record {}
*1 (x :: a) = Ring.identIsIdent R ,, *1 a
*Pident : {a : NaivePoly} polysEqual ((1R :: []) *P a) a
*Pident {[]} = record {}
*Pident {x :: a} = Ring.identIsIdent R ,, (Equivalence.transitive (Setoid.eq naivePolySetoid) (Group.identRight polyGroup {map (_*_ 1R) a}) (*1 a))
private
mapMap' : (f g : A A) (xs : NaivePoly) map f (map g xs) map (λ x f (g x)) xs
mapMap' f g [] = refl
mapMap' f g (x :: xs) rewrite mapMap' f g xs = refl
mapDist : (f : A A) (fDist : {x y : A} f (x + y) (f x) + (f y)) (xs ys : NaivePoly) polysEqual (map f (xs +P ys)) ((map f xs) +P (map f ys))
mapDist f fDist [] [] = record {}
mapDist f fDist [] (x :: ys) rewrite mapId ys | mapId (map f ys) = reflexive ,, Equivalence.reflexive (Setoid.eq naivePolySetoid)
mapDist f fDist (x :: xs) [] rewrite mapId xs | mapId (map f xs) = reflexive ,, Equivalence.reflexive (Setoid.eq naivePolySetoid)
mapDist f fDist (x :: xs) (y :: ys) = fDist {x} {y} ,, mapDist f fDist xs ys
mapWd : (f g : A A) (xs : NaivePoly) ((x : A) (f x) (g x)) polysEqual (map f xs) (map g xs)
mapWd f g [] ext = record {}
mapWd f g (x :: xs) ext = ext x ,, mapWd f g xs ext
mapDist' : (b c : A) (as : NaivePoly) polysEqual (map (_*_ (b + c)) as) (map (_*_ c) as +P map (_*_ b) as)
mapDist' b c [] = record {}
mapDist' b c (x :: as) = transitive (Ring.*DistributesOver+' R {b} {c} {x}) groupIsAbelian ,, mapDist' b c as
*Passoc : {a b c : NaivePoly} polysEqual (a *P (b *P c)) ((a *P b) *P c)
*Passoc {[]} {b} {c} = record {}
*Passoc {a :: as} {[]} {c} = record {}
*Passoc {a :: as} {b :: bs} {[]} = record {}
*Passoc {a :: as} {b :: bs} {c :: cs} = *Associative ,, Equivalence.transitive (Setoid.eq naivePolySetoid) (+PwellDefined (mapDist (_*_ a) *DistributesOver+ (map (_*_ b) cs) (map (_*_ c) bs)) (Equivalence.reflexive (Setoid.eq naivePolySetoid))) (Equivalence.transitive (Setoid.eq naivePolySetoid) ans (+PwellDefined (Equivalence.reflexive (Setoid.eq naivePolySetoid) {map (_*_ (a * b)) cs}) (Equivalence.symmetric (Setoid.eq naivePolySetoid) (mapDist (_*_ c) *DistributesOver+ (map (_*_ a) bs) (map (_*_ b) as)))))
where
ans : polysEqual (listZip _+_ (λ z z) (λ z z) (listZip _+_ (λ z z) (λ z z) (map (_*_ a) (map (_*_ b) cs)) (map (_*_ a) (map (_*_ c) bs))) (map (_*_ (b * c)) as)) (listZip _+_ (λ z z) (λ z z) (map (_*_ (a * b)) cs) (listZip _+_ (λ z z) (λ z z) (map (_*_ c) (map (_*_ a) bs)) (map (_*_ c) (map (_*_ b) as))))
ans rewrite mapMap' (_*_ a) (_*_ c) bs | mapMap' (_*_ a) (_*_ b) cs | mapMap' (_*_ c) (_*_ a) bs | mapMap' (_*_ c) (_*_ b) as = Equivalence.transitive (Setoid.eq naivePolySetoid) (Equivalence.symmetric (Setoid.eq naivePolySetoid) (Group.+Associative polyGroup {map (λ x a * (b * x)) cs} {map (λ x a * (c * x)) bs} {map (_*_ (b * c)) as})) (+PwellDefined {map (λ x a * (b * x)) cs} {(map (λ x a * (c * x)) bs) +P map (_*_ (b * c)) as} {(map (_*_ (a * b)) cs)} (mapWd (λ x a * (b * x)) (_*_ (a * b)) cs λ x *Associative) (+PwellDefined {map (λ x a * (c * x)) bs} {map (_*_ (b * c)) as} {map (λ x c * (a * x)) bs} {map (λ x c * (b * x)) as} (mapWd (λ x a * (c * x)) (λ x c * (a * x)) bs λ x transitive *Commutative (transitive (symmetric *Associative) (*WellDefined reflexive *Commutative))) (mapWd (_*_ (b * c)) (λ x c * (b * x)) as λ x transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))))
*Pdistrib : {a b c : NaivePoly} polysEqual (a *P (b +P c)) ((a *P b) +P (a *P c))
*Pdistrib {[]} {b} {c} = record {}
*Pdistrib {a :: as} {[]} {[]} = record {}
*Pdistrib {a :: as} {[]} {c :: cs} = reflexive ,, ans
where
ans : polysEqual (listZip _+_ (λ z z) (λ z z) (map (_*_ a) (map (λ z z) cs)) (map (_*_ c) as)) (map (λ z z) (listZip _+_ (λ z z) (λ z z) (map (_*_ a) cs) (map (_*_ c) as)))
ans rewrite mapId (listZip _+_ id id (map (_*_ a) cs) (map (_*_ c) as)) | mapId cs = Equivalence.reflexive (Setoid.eq naivePolySetoid)
*Pdistrib {a :: as} {b :: bs} {[]} = reflexive ,, ans
where
ans : polysEqual (listZip _+_ (λ z z) (λ z z) (map (_*_ a) (map (λ z z) bs)) (map (_*_ b) as)) (map (λ z z) (listZip _+_ (λ z z) (λ z z) (map (_*_ a) bs) (map (_*_ b) as)))
ans rewrite mapId (listZip _+_ id id (map (_*_ a) bs) (map (_*_ b) as)) | mapId bs = Equivalence.reflexive (Setoid.eq naivePolySetoid)
*Pdistrib {a :: as} {b :: bs} {c :: cs} = *DistributesOver+ ,, Equivalence.transitive (Setoid.eq naivePolySetoid) (Equivalence.transitive (Setoid.eq naivePolySetoid) (Equivalence.transitive (Setoid.eq naivePolySetoid) (Equivalence.transitive (Setoid.eq naivePolySetoid) (Equivalence.transitive (Setoid.eq naivePolySetoid) (+PwellDefined {map (_*_ a) (bs +P cs)} {map (_*_ (b + c)) as} {(map (_*_ a) bs) +P (map (_*_ a) cs)} (mapDist (_*_ a) *DistributesOver+ bs cs) (mapDist' b c as)) (Group.+Associative polyGroup {(map (_*_ a) bs) +P (map (_*_ a) cs)} {map (_*_ c) as} {map (_*_ b) as})) (+PwellDefined (Equivalence.symmetric (Setoid.eq naivePolySetoid) (Group.+Associative polyGroup {map (_*_ a) bs} {map (_*_ a) cs} {map (_*_ c) as})) (Equivalence.reflexive (Setoid.eq naivePolySetoid) {map (_*_ b) as}))) (Equivalence.symmetric (Setoid.eq naivePolySetoid) (Group.+Associative polyGroup {map (_*_ a) bs} {(map (_*_ a) cs) +P (map (_*_ c) as)} {map (_*_ b) as}))) (+PwellDefined (Equivalence.reflexive (Setoid.eq naivePolySetoid) {map (_*_ a) bs}) (AbelianGroup.commutative (abelian (record { commutative = groupIsAbelian })) {map (_*_ a) cs +P map (_*_ c) as} {map (_*_ b) as}))) (Group.+Associative polyGroup {map (_*_ a) bs} {map (_*_ b) as})

View File

@@ -0,0 +1,45 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Groups.Groups
open import Groups.Abelian.Definition
open import Groups.Homomorphisms.Definition
open import Groups.Definition
open import Numbers.Naturals.Definition
open import Setoids.Orders
open import Setoids.Setoids
open import Functions
open import Sets.EquivalenceRelations
open import Rings.Definition
open import Vectors
open import Lists.Lists
open import Maybe
open import Rings.Homomorphisms.Definition
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
module Rings.Polynomial.Ring {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
open Ring R
open Setoid S
open Equivalence eq
open import Groups.Polynomials.Group additiveGroup
open import Groups.Polynomials.Definition additiveGroup
open import Rings.Polynomial.Definition R
open import Rings.Polynomial.Multiplication R
polyRing : Ring naivePolySetoid _+P_ _*P_
Ring.additiveGroup polyRing = polyGroup
Ring.*WellDefined polyRing {a} {b} {c} {d} = *PwellDefined {a} {b} {c} {d}
Ring.1R polyRing = 1R :: []
Ring.groupIsAbelian polyRing {x} {y} = AbelianGroup.commutative (abelian (record { commutative = Ring.groupIsAbelian R })) {x} {y}
Ring.*Associative polyRing {a} {b} {c} = *Passoc {a} {b} {c}
Ring.*Commutative polyRing {a} {b} = p*Commutative {a} {b}
Ring.*DistributesOver+ polyRing {a} {b} {c} = *Pdistrib {a} {b} {c}
Ring.identIsIdent polyRing {a} = *Pident {a}
polyInjectionIsHom : RingHom R polyRing polyInjection
RingHom.preserves1 polyInjectionIsHom = reflexive ,, record {}
RingHom.ringHom polyInjectionIsHom = reflexive ,, record {}
GroupHom.groupHom (RingHom.groupHom polyInjectionIsHom) = reflexive ,, record {}
GroupHom.wellDefined (RingHom.groupHom polyInjectionIsHom) = SetoidInjection.wellDefined polyInjectionIsInj