Polynomial ring (#78)

This commit is contained in:
Patrick Stevens
2019-11-19 08:45:08 +00:00
committed by GitHub
parent 07ffda201a
commit f0790e4f52
7 changed files with 257 additions and 185 deletions

View File

@@ -47,6 +47,7 @@ open import Rings.Orders.Partial.Lemmas
open import Rings.IntegralDomains
open import Rings.DirectSum
open import Rings.Polynomial.Ring
open import Rings.Polynomial.Evaluation
open import Rings.Ideals.Definition
open import Rings.Isomorphisms.Definition

View File

@@ -2,6 +2,7 @@
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
@@ -25,37 +26,88 @@ open Group G
_+P_ : NaivePoly NaivePoly NaivePoly
_+P_ = listZip _+_ id id
PidentRight : {m : NaivePoly} polysEqual (m +P []) m
PidentRight {[]} = record {}
PidentRight {x :: m} = reflexive ,, identityOfIndiscernablesLeft polysEqual (Equivalence.reflexive (Setoid.eq naivePolySetoid)) (equalityCommutative (mapId m))
inverse' : NaivePoly NaivePoly
inverse' = map (Group.inverse G)
+PwellDefined : {m n x y : NaivePoly} polysEqual m x polysEqual n y polysEqual (m +P n) (x +P y)
+PwellDefined {[]} {[]} {[]} {[]} m=x n=y = record {}
+PwellDefined {[]} {[]} {[]} {x :: y} m=x (fst ,, snd) = fst ,, identityOfIndiscernablesRight polysEqual snd (equalityCommutative (mapId y))
+PwellDefined {[]} {[]} {x :: xs} {[]} (fst ,, snd) n=y = fst ,, identityOfIndiscernablesRight polysEqual snd (equalityCommutative (mapId xs))
+PwellDefined {[]} {[]} {x :: xs} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive (+WellDefined fst fst2) identRight ,, +PwellDefined {[]} {[]} {xs} {ys} snd snd2
+PwellDefined {[]} {n :: ns} {[]} {[]} m=x (fst ,, snd) = fst ,, identityOfIndiscernablesLeft polysEqual snd (equalityCommutative (mapId ns))
+PwellDefined {[]} {n :: ns} {[]} {y :: ys} m=x (fst ,, snd) = fst ,, +PwellDefined m=x snd
+PwellDefined {[]} {n :: ns} {x :: xs} {[]} (fst ,, snd) (fst2 ,, snd2) = transitive fst2 (symmetric fst) ,, ans
where
ans : polysEqual (map (λ z z) ns) (map (λ z z) xs)
ans rewrite mapId ns | mapId xs = Equivalence.transitive (Setoid.eq naivePolySetoid) snd2 snd
+PwellDefined {[]} {n :: ns} {x :: xs} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive (symmetric identLeft) (+WellDefined (symmetric fst) fst2) ,, (+PwellDefined snd snd2)
+PwellDefined {m :: ms} {[]} {[]} {[]} (fst ,, snd) _ = fst ,, identityOfIndiscernablesLeft polysEqual snd (equalityCommutative (mapId ms))
+PwellDefined {m :: ms} {[]} {[]} {x :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive fst (symmetric fst2) ,, ans
where
ans : polysEqual (map (λ z z) ms) (map (λ z z) ys)
ans rewrite mapId ms | mapId ys = Equivalence.transitive (Setoid.eq naivePolySetoid) snd snd2
+PwellDefined {m :: ms} {[]} {x :: xs} {[]} (fst ,, snd) n=y = fst ,, ans
where
ans : polysEqual (map (λ z z) ms) (map (λ z z) xs)
ans rewrite mapId ms | mapId xs = snd
+PwellDefined {m :: ms} {[]} {x :: xs} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive (symmetric identRight) (+WellDefined fst (symmetric fst2)) ,, identityOfIndiscernablesLeft polysEqual (Equivalence.transitive (Setoid.eq naivePolySetoid) (Equivalence.symmetric (Setoid.eq naivePolySetoid) PidentRight) (+PwellDefined snd snd2)) (equalityCommutative (mapId ms))
+PwellDefined {m :: ms} {n :: ns} {[]} {[]} (fst ,, snd) (fst2 ,, snd2) = transitive (+WellDefined fst fst2) identLeft ,, +PwellDefined snd snd2
+PwellDefined {m :: ms} {n :: ns} {[]} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive (+WellDefined fst fst2) identLeft ,, +PwellDefined snd snd2
+PwellDefined {m :: ms} {n :: ns} {x :: xs} {[]} (fst ,, snd) (fst2 ,, snd2) = transitive (+WellDefined fst fst2) identRight ,, identityOfIndiscernablesRight polysEqual (Equivalence.transitive (Setoid.eq naivePolySetoid) (+PwellDefined snd snd2) PidentRight) (equalityCommutative (mapId xs))
+PwellDefined {m :: ms} {n :: ns} {x :: xs} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = +WellDefined fst fst2 ,, +PwellDefined snd snd2
abstract
PidentRight : {m : NaivePoly} polysEqual (m +P []) m
PidentRight {[]} = record {}
PidentRight {x :: m} = reflexive ,, identityOfIndiscernablesLeft polysEqual (Equivalence.reflexive (Setoid.eq naivePolySetoid)) (equalityCommutative (mapId m))
PidentLeft : {m : NaivePoly} polysEqual ([] +P m) m
PidentLeft {[]} = record {}
PidentLeft {x :: m} = reflexive ,, identityOfIndiscernablesLeft polysEqual (Equivalence.reflexive (Setoid.eq naivePolySetoid)) (equalityCommutative (mapId m))
+PwellDefined : {m n x y : NaivePoly} polysEqual m x polysEqual n y polysEqual (m +P n) (x +P y)
+PwellDefined {[]} {[]} {[]} {[]} m=x n=y = record {}
+PwellDefined {[]} {[]} {[]} {x :: y} m=x (fst ,, snd) = fst ,, identityOfIndiscernablesRight polysEqual snd (equalityCommutative (mapId y))
+PwellDefined {[]} {[]} {x :: xs} {[]} (fst ,, snd) n=y = fst ,, identityOfIndiscernablesRight polysEqual snd (equalityCommutative (mapId xs))
+PwellDefined {[]} {[]} {x :: xs} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive (+WellDefined fst fst2) identRight ,, +PwellDefined {[]} {[]} {xs} {ys} snd snd2
+PwellDefined {[]} {n :: ns} {[]} {[]} m=x (fst ,, snd) = fst ,, identityOfIndiscernablesLeft polysEqual snd (equalityCommutative (mapId ns))
+PwellDefined {[]} {n :: ns} {[]} {y :: ys} m=x (fst ,, snd) = fst ,, +PwellDefined m=x snd
+PwellDefined {[]} {n :: ns} {x :: xs} {[]} (fst ,, snd) (fst2 ,, snd2) = transitive fst2 (symmetric fst) ,, ans
where
ans : polysEqual (map (λ z z) ns) (map (λ z z) xs)
ans rewrite mapId ns | mapId xs = Equivalence.transitive (Setoid.eq naivePolySetoid) snd2 snd
+PwellDefined {[]} {n :: ns} {x :: xs} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive (symmetric identLeft) (+WellDefined (symmetric fst) fst2) ,, (+PwellDefined snd snd2)
+PwellDefined {m :: ms} {[]} {[]} {[]} (fst ,, snd) _ = fst ,, identityOfIndiscernablesLeft polysEqual snd (equalityCommutative (mapId ms))
+PwellDefined {m :: ms} {[]} {[]} {x :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive fst (symmetric fst2) ,, ans
where
ans : polysEqual (map (λ z z) ms) (map (λ z z) ys)
ans rewrite mapId ms | mapId ys = Equivalence.transitive (Setoid.eq naivePolySetoid) snd snd2
+PwellDefined {m :: ms} {[]} {x :: xs} {[]} (fst ,, snd) n=y = fst ,, ans
where
ans : polysEqual (map (λ z z) ms) (map (λ z z) xs)
ans rewrite mapId ms | mapId xs = snd
+PwellDefined {m :: ms} {[]} {x :: xs} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive (symmetric identRight) (+WellDefined fst (symmetric fst2)) ,, identityOfIndiscernablesLeft polysEqual (Equivalence.transitive (Setoid.eq naivePolySetoid) (Equivalence.symmetric (Setoid.eq naivePolySetoid) PidentRight) (+PwellDefined snd snd2)) (equalityCommutative (mapId ms))
+PwellDefined {m :: ms} {n :: ns} {[]} {[]} (fst ,, snd) (fst2 ,, snd2) = transitive (+WellDefined fst fst2) identLeft ,, +PwellDefined snd snd2
+PwellDefined {m :: ms} {n :: ns} {[]} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = transitive (+WellDefined fst fst2) identLeft ,, +PwellDefined snd snd2
+PwellDefined {m :: ms} {n :: ns} {x :: xs} {[]} (fst ,, snd) (fst2 ,, snd2) = transitive (+WellDefined fst fst2) identRight ,, identityOfIndiscernablesRight polysEqual (Equivalence.transitive (Setoid.eq naivePolySetoid) (+PwellDefined snd snd2) PidentRight) (equalityCommutative (mapId xs))
+PwellDefined {m :: ms} {n :: ns} {x :: xs} {y :: ys} (fst ,, snd) (fst2 ,, snd2) = +WellDefined fst fst2 ,, +PwellDefined snd snd2
PidentLeft : {m : NaivePoly} polysEqual ([] +P m) m
PidentLeft {[]} = record {}
PidentLeft {x :: m} = reflexive ,, identityOfIndiscernablesLeft polysEqual (Equivalence.reflexive (Setoid.eq naivePolySetoid)) (equalityCommutative (mapId m))
invLeft' : {a : NaivePoly} polysEqual ((inverse' a) +P a) []
invLeft' {[]} = record {}
invLeft' {x :: a} = Group.invLeft G ,, invLeft' {a}
invRight' : {a : NaivePoly} polysEqual (a +P (inverse' a)) []
invRight' {[]} = record {}
invRight' {x :: a} = Group.invRight G ,, invRight' {a}
assoc : {a b c : NaivePoly} polysEqual (a +P (b +P c)) ((a +P b) +P c)
assoc {[]} {[]} {[]} = record {}
assoc {[]} {[]} {x :: c} = reflexive ,, ans
where
ans : polysEqual (map (λ z z) (map (λ z z) c)) (map (λ z z) c)
ans rewrite mapId c | mapId c = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {[]} {b :: bs} {[]} = reflexive ,, ans
where
ans : polysEqual (map (λ z z) (map (λ z z) bs)) (map (λ z z) (map (λ z z) bs))
ans rewrite mapId bs | mapId bs = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {[]} {b :: bs} {c :: cs} = reflexive ,, ans
where
ans : polysEqual (map (λ z z) (listZip _+_ (λ z z) (λ z z) bs cs)) (listZip _+_ (λ z z) (λ z z) (map (λ z z) bs) cs)
ans rewrite mapId bs | mapId (listZip _+_ id id bs cs) = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {a :: as} {[]} {[]} = reflexive ,, ans
where
ans : polysEqual (map (λ z z) as) (map (λ z z) (map (λ z z) as))
ans rewrite mapId as | mapId as = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {a :: as} {[]} {c :: cs} = reflexive ,, ans
where
ans : polysEqual (listZip _+_ (λ z z) (λ z z) as (map (λ z z) cs)) (listZip _+_ (λ z z) (λ z z) (map (λ z z) as) cs)
ans rewrite mapId cs | mapId as = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {a :: as} {b :: bs} {[]} = reflexive ,, ans
where
ans : polysEqual (listZip _+_ (λ z z) (λ z z) as (map (λ z z) bs)) (map (λ z z) (listZip _+_ (λ z z) (λ z z) as bs))
ans rewrite mapId (listZip _+_ id id as bs) | mapId bs = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {a :: as} {b :: bs} {c :: cs} = Group.+Associative G ,, assoc {as} {bs} {cs}
comm : AbelianGroup G {x y : NaivePoly} polysEqual (x +P y) (y +P x)
comm com {[]} {y} = Equivalence.transitive (Setoid.eq naivePolySetoid) PidentLeft (Equivalence.symmetric (Setoid.eq naivePolySetoid) PidentRight)
comm com {x :: xs} {[]} = reflexive ,, Equivalence.reflexive (Setoid.eq naivePolySetoid)
comm com {x :: xs} {y :: ys} = AbelianGroup.commutative com ,, comm com {xs} {ys}
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

View File

@@ -23,59 +23,15 @@ open import Groups.Polynomials.Addition G public
open Setoid S
open Equivalence eq
inverse : NaivePoly NaivePoly
inverse = map (Group.inverse G)
invLeft : {a : NaivePoly} polysEqual ((inverse a) +P a) []
invLeft {[]} = record {}
invLeft {x :: a} = Group.invLeft G ,, invLeft {a}
invRight : {a : NaivePoly} polysEqual (a +P (inverse a)) []
invRight {[]} = record {}
invRight {x :: a} = Group.invRight G ,, invRight {a}
assoc : {a b c : NaivePoly} polysEqual (a +P (b +P c)) ((a +P b) +P c)
assoc {[]} {[]} {[]} = record {}
assoc {[]} {[]} {x :: c} = reflexive ,, ans
where
ans : polysEqual (map (λ z z) (map (λ z z) c)) (map (λ z z) c)
ans rewrite mapId c | mapId c = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {[]} {b :: bs} {[]} = reflexive ,, ans
where
ans : polysEqual (map (λ z z) (map (λ z z) bs)) (map (λ z z) (map (λ z z) bs))
ans rewrite mapId bs | mapId bs = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {[]} {b :: bs} {c :: cs} = reflexive ,, ans
where
ans : polysEqual (map (λ z z) (listZip _+_ (λ z z) (λ z z) bs cs)) (listZip _+_ (λ z z) (λ z z) (map (λ z z) bs) cs)
ans rewrite mapId bs | mapId (listZip _+_ id id bs cs) = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {a :: as} {[]} {[]} = reflexive ,, ans
where
ans : polysEqual (map (λ z z) as) (map (λ z z) (map (λ z z) as))
ans rewrite mapId as | mapId as = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {a :: as} {[]} {c :: cs} = reflexive ,, ans
where
ans : polysEqual (listZip _+_ (λ z z) (λ z z) as (map (λ z z) cs)) (listZip _+_ (λ z z) (λ z z) (map (λ z z) as) cs)
ans rewrite mapId cs | mapId as = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {a :: as} {b :: bs} {[]} = reflexive ,, ans
where
ans : polysEqual (listZip _+_ (λ z z) (λ z z) as (map (λ z z) bs)) (map (λ z z) (listZip _+_ (λ z z) (λ z z) as bs))
ans rewrite mapId (listZip _+_ id id as bs) | mapId bs = Equivalence.reflexive (Setoid.eq naivePolySetoid)
assoc {a :: as} {b :: bs} {c :: cs} = Group.+Associative G ,, assoc {as} {bs} {cs}
polyGroup : Group naivePolySetoid _+P_
Group.+WellDefined polyGroup = +PwellDefined
Group.0G polyGroup = []
Group.inverse polyGroup = inverse
Group.inverse polyGroup = inverse'
Group.+Associative polyGroup {a} {b} {c} = assoc {a} {b} {c}
Group.identRight polyGroup = PidentRight
Group.identLeft polyGroup = PidentLeft
Group.invLeft polyGroup {a} = invLeft {a}
Group.invRight polyGroup {a} = invRight {a}
comm : AbelianGroup G {x y : NaivePoly} polysEqual (x +P y) (y +P x)
comm com {[]} {y} = Equivalence.transitive (Setoid.eq naivePolySetoid) (Group.identLeft polyGroup) (Equivalence.symmetric (Setoid.eq naivePolySetoid) (Group.identRight polyGroup))
comm com {x :: xs} {[]} = reflexive ,, Equivalence.reflexive (Setoid.eq naivePolySetoid)
comm com {x :: xs} {y :: ys} = AbelianGroup.commutative com ,, comm com {xs} {ys}
Group.invLeft polyGroup {a} = invLeft' {a}
Group.invRight polyGroup {a} = invRight' {a}
abelian : AbelianGroup G AbelianGroup polyGroup
AbelianGroup.commutative (abelian ab) {x} {y} = comm ab {x} {y}

View File

@@ -1,33 +0,0 @@
{-# 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,60 @@
{-# 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.Evaluation {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 Group additiveGroup
open import Groups.Polynomials.Definition additiveGroup
open import Groups.Polynomials.Addition additiveGroup
open import Rings.Polynomial.Ring R
open import Rings.Polynomial.Multiplication R
inducedFunction : NaivePoly A A
inducedFunction [] a = 0R
inducedFunction (x :: p) a = x + (a * inducedFunction p a)
inducedFunctionMult : (as : NaivePoly) (b c : A) inducedFunction (map (_*_ b) as) c (inducedFunction as c) * b
inducedFunctionMult [] b c = symmetric (transitive *Commutative timesZero)
inducedFunctionMult (x :: as) b c = transitive (transitive (+WellDefined reflexive (transitive (transitive (*WellDefined reflexive (inducedFunctionMult as b c)) *Associative) *Commutative)) (symmetric *DistributesOver+)) *Commutative
inducedFunctionWellDefined : {a b : NaivePoly} polysEqual a b (c : A) inducedFunction a c inducedFunction b c
inducedFunctionWellDefined {[]} {[]} a=b c = reflexive
inducedFunctionWellDefined {[]} {x :: b} (fst ,, snd) c = symmetric (transitive (+WellDefined fst (transitive (*WellDefined reflexive (symmetric (inducedFunctionWellDefined {[]} {b} snd c))) (timesZero {c}))) identRight)
inducedFunctionWellDefined {a :: as} {[]} (fst ,, snd) c = transitive (+WellDefined fst reflexive) (transitive identLeft (transitive (*WellDefined reflexive (inducedFunctionWellDefined {as} {[]} snd c)) (timesZero {c})))
inducedFunctionWellDefined {a :: as} {b :: bs} (fst ,, snd) c = +WellDefined fst (*WellDefined reflexive (inducedFunctionWellDefined {as} {bs} snd c))
inducedFunctionGroupHom : {x y : NaivePoly} (a : A) inducedFunction (x +P y) a (inducedFunction x a + inducedFunction y a)
inducedFunctionGroupHom {[]} {[]} a = symmetric identLeft
inducedFunctionGroupHom {[]} {x :: y} a rewrite mapId y = symmetric identLeft
inducedFunctionGroupHom {x :: xs} {[]} a rewrite mapId xs = symmetric identRight
inducedFunctionGroupHom {x :: xs} {y :: ys} a = transitive (symmetric +Associative) (transitive (+WellDefined reflexive (transitive (transitive (+WellDefined reflexive (transitive (*WellDefined reflexive (transitive (inducedFunctionGroupHom {xs} {ys} a) groupIsAbelian)) *DistributesOver+)) +Associative) groupIsAbelian)) +Associative)
inducedFunctionRingHom : (r s : NaivePoly) (a : A) inducedFunction (r *P s) a (inducedFunction r a * inducedFunction s a)
inducedFunctionRingHom [] s a = symmetric (transitive *Commutative timesZero)
inducedFunctionRingHom (x :: xs) [] a = symmetric timesZero
inducedFunctionRingHom (b :: bs) (c :: cs) a = transitive (+WellDefined reflexive (*WellDefined reflexive (inducedFunctionGroupHom {map (_*_ b) cs +P map (_*_ c) bs} {0G :: (bs *P cs)} a))) (transitive (+WellDefined reflexive (*WellDefined reflexive (+WellDefined (inducedFunctionGroupHom {map (_*_ b) cs} {map (_*_ c) bs} a) identLeft))) (transitive (transitive (transitive (+WellDefined reflexive (transitive (transitive (transitive (*WellDefined reflexive (transitive (+WellDefined (transitive groupIsAbelian (+WellDefined (inducedFunctionMult bs c a) (inducedFunctionMult cs b a))) (transitive (transitive (*WellDefined reflexive (transitive (inducedFunctionRingHom bs cs a) *Commutative)) *Associative) *Commutative)) (symmetric +Associative))) *DistributesOver+) (+WellDefined reflexive *DistributesOver+)) (+WellDefined *Associative (+WellDefined (transitive *Associative *Commutative) *Associative)))) +Associative) (+WellDefined (symmetric *DistributesOver+') (symmetric *DistributesOver+'))) (symmetric *DistributesOver+)))
inducedFunctionIsHom : (a : A) RingHom polyRing R (λ p inducedFunction p a)
RingHom.preserves1 (inducedFunctionIsHom a) = transitive (+WellDefined reflexive (timesZero {a})) identRight
RingHom.ringHom (inducedFunctionIsHom a) {r} {s} = inducedFunctionRingHom r s a
GroupHom.groupHom (RingHom.groupHom (inducedFunctionIsHom a)) {x} {y} = inducedFunctionGroupHom {x} {y} a
GroupHom.wellDefined (RingHom.groupHom (inducedFunctionIsHom a)) x=y = inducedFunctionWellDefined x=y a

View File

@@ -30,94 +30,131 @@ 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))
(x :: a) *P (y :: b) = (x * y) :: (((map (x *_) b) +P (map (y *_) a)) +P (0R :: (a *P b)))
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}
abstract
+PCommutative : {x y : NaivePoly} polysEqual (x +P y) (y +P x)
+PCommutative {x} {y} = comm (record { commutative = groupIsAbelian }) {x} {y}
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
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 ,, +PwellDefined (+PCommutative {map (_*_ x) ys} {map (_*_ y) xs}) (reflexive ,, p*Commutative {xs} {ys})
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
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
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
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 : 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
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
*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)
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
*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}))
*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.transitive (Setoid.eq naivePolySetoid) (Equivalence.transitive (Setoid.eq naivePolySetoid) {_} {(map (_*_ c) bs) +P (0G :: (cs *P bs))} (Equivalence.transitive (Setoid.eq naivePolySetoid) {_} {[] +P (0G :: (cs *P bs))} (reflexive ,, ans) (+PwellDefined (Equivalence.symmetric (Setoid.eq naivePolySetoid) (zeroTimes1 {bs} c fst)) (Equivalence.reflexive (Setoid.eq naivePolySetoid) {0G :: (cs *P bs)}))) (+PwellDefined (Equivalence.symmetric (Setoid.eq naivePolySetoid) (Group.identRight polyGroup {map (_*_ c) bs})) (Equivalence.reflexive (Setoid.eq naivePolySetoid) {0G :: (cs *P bs)}))) (+PwellDefined {(map (_*_ c) bs +P [])} {0G :: (cs *P bs)} {(map (_*_ c) bs) +P map (_*_ b) cs} (+PwellDefined (Equivalence.reflexive (Setoid.eq naivePolySetoid) {map (_*_ c) bs}) (Equivalence.symmetric (Setoid.eq naivePolySetoid) (zeroTimes2 {cs} b (Equivalence.symmetric (Setoid.eq naivePolySetoid) snd)))) (Equivalence.reflexive (Setoid.eq naivePolySetoid) {0G :: (cs *P bs)}))
where
ans : polysEqual [] (map id (cs *P bs))
ans rewrite mapId (cs *P bs) = *PwellDefinedL snd
*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 {(map (_*_ a) bs +P map (_*_ b) as)} {0G :: (as *P bs)} {[] +P []} {[]} (+PwellDefined {map (_*_ a) bs} {map (_*_ b) as} {[]} {[]} (zeroTimes1 {bs} a fst) (zeroTimes2 {as} b snd)) (reflexive ,, *PwellDefinedL {as} {bs} {[]} snd)) (record {})
*PwellDefinedL {a :: as} {b :: bs} {c :: cs} (fst ,, snd) = *WellDefined fst reflexive ,, +PwellDefined {(map (_*_ a) bs) +P (map (_*_ b) as)} {0G :: (as *P bs)} {map (_*_ c) bs +P map (_*_ b) cs} {0G :: (cs *P bs)} (+PwellDefined {map (_*_ a) bs} {map (_*_ b) as} {map (_*_ c) bs} {map (_*_ b) cs} (mapWellDefined a c bs fst) (mapWellDefined' b as cs snd)) (reflexive ,, *PwellDefinedL {as} {bs} {cs} snd)
*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)
*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}))
private
*1 : (a : NaivePoly) polysEqual (map (_*_ 1R) a) a
*1 [] = record {}
*1 (x :: a) = Ring.identIsIdent R ,, *1 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)
*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
*1 : (a : NaivePoly) polysEqual (map (_*_ 1R) a) a
*1 [] = record {}
*1 (x :: a) = Ring.identIsIdent R ,, *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
*Pident : {a : NaivePoly} polysEqual ((1R :: []) *P a) a
*Pident {[]} = record {}
*Pident {x :: a} = Ring.identIsIdent R ,, Equivalence.transitive (Setoid.eq naivePolySetoid) (+PwellDefined {map (_*_ 1R) a +P []} {0G :: []} {map (_*_ 1R) a} {[]} (Group.identRight polyGroup) (reflexive ,, record {})) (Equivalence.transitive (Setoid.eq naivePolySetoid) (Group.identRight polyGroup {map (_*_ 1R) a}) (*1 a))
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
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
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
mapMap'' : (f g : A A) (xs : NaivePoly) Setoid.__ naivePolySetoid (map f (map g xs)) (map (λ x f (g x)) xs)
mapMap'' f g xs rewrite mapMap' f g xs = Equivalence.reflexive (Setoid.eq naivePolySetoid)
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
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
mapTimes : (a b : NaivePoly) (c : A) polysEqual (a *P (map (_*_ c) b)) (map (_*_ c) (a *P b))
mapTimes [] b c = record {}
mapTimes (a :: as) [] c = record {}
mapTimes (a :: as) (x :: b) c = transitive *Associative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)) ,, trans (+PwellDefined {(map (_*_ a) (map (_*_ c) b)) +P map (_*_ (c * x)) as} {0G :: (as *P map (_*_ c) b)} {map (_*_ c) (map (_*_ a) b +P map (_*_ x) as)} (trans (+PwellDefined {map (_*_ a) (map (_*_ c) b)} {map (_*_ (c * x)) as} {map (_*_ c) (map (_*_ a) b)} (trans (mapMap'' (_*_ a) (_*_ c) b) (trans (mapWd (λ x a * (c * x)) (λ x c * (a * x)) b (λ x transitive *Associative (transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)))) (sym (mapMap'' (_*_ c) (_*_ a) b)))) (trans (mapWd (_*_ (c * x)) (λ y c * (x * y)) as λ y symmetric *Associative) (sym (mapMap'' (_*_ c) (_*_ x) as)))) (sym (mapDist (_*_ c) *DistributesOver+ (map (_*_ a) b) (map (_*_ x) as)))) (symmetric timesZero ,, mapTimes as b c)) (sym (mapDist (_*_ c) *DistributesOver+ (map (_*_ a) b +P map (_*_ x) as) (0G :: (as *P b))))
where
open Equivalence (Setoid.eq naivePolySetoid) renaming (transitive to trans ; symmetric to sym ; reflexive to ref)
mapTimes' : (a b : NaivePoly) (c : A) polysEqual (map (_*_ c) (a *P b)) ((map (_*_ c) a) *P b)
mapTimes' a b c = trans (trans (mapWellDefined' c (a *P b) (b *P a) (p*Commutative {a} {b})) (sym (mapTimes b a c))) (p*Commutative {b} {map (_*_ c) a})
where
open Equivalence (Setoid.eq naivePolySetoid) renaming (transitive to trans ; symmetric to sym ; reflexive to ref)
bumpTimes' : (a bs : NaivePoly) (b : A) polysEqual (a *P (b :: bs)) (map (_*_ b) a +P (0G :: (a *P bs)))
bumpTimes' [] bs b = reflexive ,, record {}
bumpTimes' (x :: a) bs b = transitive *Commutative (symmetric identRight) ,, trans (+PwellDefined {map (_*_ x) bs +P map (_*_ b) a} {0G :: (a *P bs)} {_} {0G :: (a *P bs)} (+PCommutative {map (_*_ x) bs} {map (_*_ b) a}) (ref {0G :: (a *P bs)})) (trans (sym (Group.+Associative polyGroup {map (_*_ b) a})) (+PwellDefined {map (_*_ b) a} {_} {map (_*_ b) a} {_} (ref {map (_*_ b) a}) (trans (trans (+PwellDefined {map (_*_ x) bs} {_} {map (_*_ x) bs} (ref {map (_*_ x) bs}) (reflexive ,, p*Commutative {a})) (sym (bumpTimes' bs a x))) (p*Commutative {bs}))))
where
open Equivalence (Setoid.eq naivePolySetoid) renaming (transitive to trans ; symmetric to sym ; reflexive to ref)
bumpTimes : {a b : NaivePoly} polysEqual (a *P (0G :: b)) (0G :: (a *P b))
bumpTimes {[]} {b} = reflexive ,, record {}
bumpTimes {x :: a} {[]} = timesZero ,, trans (sym (Group.+Associative polyGroup {[]} {map (_*_ 0G) a})) ans
where
open Equivalence (Setoid.eq naivePolySetoid) renaming (transitive to trans ; symmetric to sym ; reflexive to ref)
ans : polysEqual (map id (map (_*_ 0G) a +P (0G :: (a *P [])))) []
ans rewrite mapId (map (_*_ 0G) a +P (0G :: (a *P []))) = (+PwellDefined {map (_*_ 0G) a} {0G :: (a *P [])} {[]} {[]} (zeroTimes1 0G reflexive) (reflexive ,, p*Commutative {a}))
bumpTimes {x :: a} {b :: bs} = timesZero ,, trans (+PwellDefined {((x * b) :: map (_*_ x) bs) +P map (_*_ 0G) a} {0G :: (a *P (b :: bs))} {(x * b) :: map (_*_ x) bs} {0G :: (a *P (b :: bs))} (trans (+PwellDefined {(x * b) :: map (_*_ x) bs} {map (_*_ 0G) a} {(x * b) :: map (_*_ x) bs} {[]} (ref {(x * b) :: map (_*_ x) bs}) (zeroTimes1 0G reflexive)) (Group.identRight polyGroup {(x * b) :: map (_*_ x) bs})) (ref {0G :: (a *P (b :: bs))})) (identRight ,, trans (+PwellDefined {map (_*_ x) bs} {_} {map (_*_ x) bs} ref (bumpTimes' a bs b)) (Group.+Associative polyGroup {map (_*_ x) bs} {map (_*_ b) a} {0G :: (a *P bs)}))
where
open Equivalence (Setoid.eq naivePolySetoid) renaming (transitive to trans ; symmetric to sym ; reflexive to ref)
*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} rewrite mapId cs | mapId ((map (_*_ a) cs +P map (_*_ c) as) +P (0G :: (as *P cs))) = reflexive ,, +PwellDefined {(map (_*_ a) cs) +P (map (_*_ c) as)} {_} {(map (_*_ a) cs) +P (map (_*_ c) as)} (Equivalence.reflexive (Setoid.eq naivePolySetoid)) (Equivalence.reflexive (Setoid.eq naivePolySetoid) {0G :: (as *P cs)})
*Pdistrib {a :: as} {b :: bs} {[]} rewrite mapId bs | mapId ((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs))) = reflexive ,, +PwellDefined {(map (_*_ a) bs) +P (map (_*_ b) as)} {_} {(map (_*_ a) bs) +P (map (_*_ b) as)} (Equivalence.reflexive (Setoid.eq naivePolySetoid)) (Equivalence.reflexive (Setoid.eq naivePolySetoid) {0G :: (as *P bs)})
*Pdistrib {a :: as} {b :: bs} {c :: cs} = *DistributesOver+ ,, trans (+PwellDefined {(map (_*_ a) (bs +P cs)) +P (map (_*_ (b + c)) as)} {0G :: (as *P (bs +P cs))} {(map (_*_ a) bs +P map (_*_ b) as) +P ((map (_*_ a) cs) +P (map (_*_ c) as))} {0G :: ((as *P bs) +P (as *P cs))} (trans (+PwellDefined (mapDist (_*_ a) *DistributesOver+ bs cs) (mapDist' b c as)) (trans (sym (+Assoc {map (_*_ a) bs} {map (_*_ a) cs})) (trans (+PwellDefined (ref {map (_*_ a) bs}) (trans (trans (+Assoc {map (_*_ a) cs} {map (_*_ c) as}) ref) (+PCommutative {map (_*_ a) cs +P map (_*_ c) as} {map (_*_ b) as}))) (+Assoc {map (_*_ a) bs} {map (_*_ b) as})))) (reflexive ,, *Pdistrib {as} {bs} {cs})) (trans (sym (+Assoc {(map (_*_ a) bs +P map (_*_ b) as)} {(map (_*_ a) cs +P map (_*_ c) as)} {0G :: ((as *P bs) +P (as *P cs))})) (trans (+PwellDefined (ref {map (_*_ a) bs +P map (_*_ b) as}) (trans (trans (+PwellDefined (ref {map (_*_ a) cs +P map (_*_ c) as}) (symmetric identLeft ,, +PCommutative {as *P bs})) (+Assoc {map (_*_ a) cs +P (map (_*_ c) as)} {0G :: (as *P cs)} {0G :: (as *P bs)})) (+PCommutative {(map (_*_ a) cs +P map (_*_ c) as) +P (0G :: (as *P cs))} {0G :: (as *P bs)}))) (+Assoc {map (_*_ a) bs +P map (_*_ b) as} {0G :: (as *P bs)})))
where
open Equivalence (Setoid.eq naivePolySetoid) renaming (transitive to trans ; symmetric to sym ; reflexive to ref)
open Group polyGroup renaming (+Associative to +Assoc ; 0G to 0P) hiding (identLeft)
*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)))))
*Passoc {a :: as} {b :: bs} {c :: cs} = *Associative ,, trans (+PwellDefined {(map (_*_ a) ((map (_*_ b) cs +P map (_*_ c) bs) +P (0G :: (bs *P cs))) +P map (_*_ (b * c)) as)} {0G :: (as *P ((map (_*_ b) cs +P map (_*_ c) bs) +P (0G :: (bs *P cs))))} {(((map (_*_ (a * b)) cs) +P (map (_*_ (a * c)) bs)) +P (map (_*_ a) (0G :: (bs *P cs)))) +P (map (_*_ (b * c)) as)} {0G :: (((map (_*_ b) (as *P cs)) +P (map (_*_ c) (as *P bs))) +P (as *P (0G :: (bs *P cs))))} (+PwellDefined {map (_*_ a) ((map (_*_ b) cs +P map (_*_ c) bs) +P (0G :: (bs *P cs)))} {map (_*_ (b * c)) as} {((map (_*_ (a * b)) cs +P map (_*_ (a * c)) bs) +P ((a * 0G) :: map (_*_ a) (bs *P cs)))} {map (_*_ (b * c)) as} (trans (mapDist (_*_ a) *DistributesOver+ (map (_*_ b) cs +P map (_*_ c) bs) (0G :: (bs *P cs))) (+PwellDefined {map (_*_ a) (map (_*_ b) cs +P map (_*_ c) bs)} {(a * 0G) :: map (_*_ a) (bs *P cs)} {map (_*_ (a * b)) cs +P map (_*_ (a * c)) bs} (trans (mapDist (_*_ a) *DistributesOver+ (map (_*_ b) cs) (map (_*_ c) bs)) (+PwellDefined {map (_*_ a) (map (_*_ b) cs)} {map (_*_ a) (map (_*_ c) bs)} {map (_*_ (a * b)) cs} (trans (mapMap'' (_*_ a) (_*_ b) cs) (mapWd (λ x a * (b * x)) (_*_ (a * b)) cs (λ x *Associative))) (trans (mapMap'' (_*_ a) (_*_ c) bs) (mapWd (λ x a * (c * x)) (_*_ (a * c)) bs λ x *Associative)))) (ref {(a * 0G) :: map (_*_ a) (bs *P cs)}))) (ref {map (_*_ (b * c)) as})) (reflexive ,, trans (*Pdistrib {as} {(map (_*_ b) cs +P map (_*_ c) bs)} {0G :: (bs *P cs)}) (+PwellDefined {(as *P (map (_*_ b) cs +P map (_*_ c) bs))} {as *P (0G :: (bs *P cs))} {(map (_*_ b) (as *P cs) +P map (_*_ c) (as *P bs))} (trans (*Pdistrib {as} {map (_*_ b) cs} {map (_*_ c) bs}) (+PwellDefined {as *P map (_*_ b) cs} {as *P map (_*_ c) bs} {map (_*_ b) (as *P cs)} (mapTimes as cs b) (mapTimes as bs c))) (ref {as *P (0G :: (bs *P cs))})))) (trans (trans (sym (Group.+Associative polyGroup {((map (_*_ (a * b)) cs) +P (map (_*_ (a * c)) bs)) +P ((a * 0G) :: map (_*_ a) (bs *P cs))})) (trans (sym (Group.+Associative polyGroup {(map (_*_ (a * b)) cs) +P (map (_*_ (a * c)) bs)})) (trans (sym (Group.+Associative polyGroup {map (_*_ (a * b)) cs})) (+PwellDefined {map (_*_ (a * b)) cs} (ref {map (_*_ (a * b)) cs}) (trans (trans ans (Group.+Associative polyGroup {map (_*_ c) (map (_*_ a) bs +P map (_*_ b) as)})) (+PwellDefined {_} {0G :: (((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs))) *P cs)} {map (_*_ c) ((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs)))} (sym (mapDist (_*_ c) *DistributesOver+ ((map (_*_ a) bs +P map (_*_ b) as)) (0G :: (as *P bs)))) (ref {0G :: (((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs))) *P cs)}))))))) (Group.+Associative polyGroup {map (_*_ (a * b)) cs} {map (_*_ c) ((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs)))}))
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})
open Equivalence (Setoid.eq naivePolySetoid) renaming (transitive to trans ; symmetric to sym ; reflexive to ref)
open Group polyGroup renaming (+Associative to +Assoc ; 0G to 0P) hiding (identLeft ; identRight)
ans2 : polysEqual ((map (_*_ a) (bs *P cs) +P (map (_*_ b) (as *P cs) +P map (_*_ c) (as *P bs))) +P (as *P (0G :: (bs *P cs)))) (map (_*_ c) (as *P bs) +P (((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs))) *P cs))
ans2 = trans (+PwellDefined {_} {as *P (0G :: (bs *P cs))} {_} {as *P (0G :: (bs *P cs))} (trans (+Assoc {map (_*_ a) (bs *P cs)} {map (_*_ b) (as *P cs)} {map (_*_ c) (as *P bs)}) (+PCommutative {map (_*_ a) (bs *P cs) +P map (_*_ b) (as *P cs)} {map (_*_ c) (as *P bs)})) ref) (trans (sym (+Assoc {map (_*_ c) (as *P bs)} {_} {as *P (0G :: (bs *P cs))}) ) (+PwellDefined {map (_*_ c) (as *P bs)} {_} {map (_*_ c) (as *P bs)} ref (trans (+PwellDefined {map (_*_ a) (bs *P cs) +P map (_*_ b) (as *P cs)} {as *P (0G :: (bs *P cs))} {((map (_*_ a) bs) *P cs) +P ((map (_*_ b) as) *P cs)} {0G :: (as *P (bs *P cs))} (+PwellDefined {map (_*_ a) (bs *P cs)} {map (_*_ b) (as *P cs)} {map (_*_ a) bs *P cs} {map (_*_ b) as *P cs} (mapTimes' bs cs a) (mapTimes' as cs b)) (bumpTimes {as} {bs *P cs})) (trans (trans (+PwellDefined {(map (_*_ a) bs *P cs) +P (map (_*_ b) as *P cs)} {0G :: (as *P (bs *P cs))} {cs *P (map (_*_ a) bs +P map (_*_ b) as)} {cs *P (0G :: (as *P bs))} (trans (+PwellDefined {map (_*_ a) bs *P cs} {map (_*_ b) as *P cs} {cs *P map (_*_ a) bs} {cs *P map (_*_ b) as} (p*Commutative {_} {cs}) (p*Commutative {_} {cs})) (sym (*Pdistrib {cs} {map (_*_ a) bs} {map (_*_ b) as}))) (trans (reflexive ,, trans (*Passoc {as} {bs} {cs}) (p*Commutative {_} {cs})) (sym (bumpTimes {cs} {as *P bs})))) (sym (*Pdistrib {cs} {(map (_*_ a) bs) +P (map (_*_ b) as)} {0G :: (as *P bs)}))) (p*Commutative {cs} {(map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs))})))))
ans :
polysEqual
((map (_*_ (a * c)) bs) +P (((a * 0G) :: map (_*_ a) (bs *P cs)) +P (map (_*_ (b * c)) as +P (0G :: ((map (_*_ b) (as *P cs) +P map (_*_ c) (as *P bs)) +P (as *P (0G :: (bs *P cs))))))))
((map (_*_ c) ((map (_*_ a) bs) +P (map (_*_ b) as))) +P (((c * 0G) :: map (_*_ c) (as *P bs)) +P (0G :: (((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs))) *P cs))))
ans = trans (trans (+PwellDefined {map (_*_ (a * c)) bs} {_} {map (_*_ c) (map (_*_ a) bs)} (trans (mapWd (_*_ (a * c)) (λ x c * (a * x)) bs (λ x transitive (*WellDefined *Commutative reflexive) (symmetric *Associative))) (sym (mapMap'' (_*_ c) (_*_ a) bs))) (trans (+PCommutative {(a * 0G) :: map (_*_ a) (bs *P cs)} {map (_*_ (b * c)) as +P (0G :: ((map (_*_ b) (as *P cs) +P map (_*_ c) (as *P bs)) +P (as *P (0G :: bs *P cs))))}) (trans (sym (+Assoc {map (_*_ (b * c)) as})) (+PwellDefined {map (_*_ (b * c)) as} {_} {map (_*_ c) (map (_*_ b) as)} (trans (mapWd (_*_ (b * c)) (λ x c * (b * x)) as λ x transitive (*WellDefined *Commutative reflexive) (symmetric *Associative)) (sym (mapMap'' (_*_ c) (_*_ b) as))) (transitive identLeft (transitive (transitive timesZero (symmetric timesZero)) (symmetric identRight)) ,, trans (+PCommutative {(map (_*_ b) (as *P cs) +P map (_*_ c) (as *P bs)) +P (as *P (0G :: (bs *P cs)))} {map (_*_ a) (bs *P cs)}) (trans (+Assoc {map (_*_ a) (bs *P cs)} {(map (_*_ b) (as *P cs)) +P (map (_*_ c) (as *P bs))} {as *P (0G :: (bs *P cs))}) ans2)))))) (+Assoc {map (_*_ c) (map (_*_ a) bs)})) (+PwellDefined {(map (_*_ c) (map (_*_ a) bs)) +P (map (_*_ c) (map (_*_ b) as))} {(((c * 0G) :: map (_*_ c) (as *P bs)) +P (0G :: (((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs))) *P cs)))} {map (_*_ c) ((map (_*_ a) bs) +P (map (_*_ b) as))} {(((c * 0G) :: map (_*_ c) (as *P bs)) +P (0G :: (((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs))) *P cs)))} (sym (mapDist (_*_ c) (*DistributesOver+) (map (_*_ a) bs) (map (_*_ b) as))) (ref {(((c * 0G) :: map (_*_ c) (as *P bs)) +P (0G :: (((map (_*_ a) bs +P map (_*_ b) as) +P (0G :: (as *P bs))) *P cs)))}))

View File

@@ -25,7 +25,6 @@ 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_
@@ -40,6 +39,6 @@ Ring.identIsIdent polyRing {a} = *Pident {a}
polyInjectionIsHom : RingHom R polyRing polyInjection
RingHom.preserves1 polyInjectionIsHom = reflexive ,, record {}
RingHom.ringHom polyInjectionIsHom = reflexive ,, record {}
RingHom.ringHom polyInjectionIsHom = reflexive ,, (reflexive ,, record {})
GroupHom.groupHom (RingHom.groupHom polyInjectionIsHom) = reflexive ,, record {}
GroupHom.wellDefined (RingHom.groupHom polyInjectionIsHom) = SetoidInjection.wellDefined polyInjectionIsInj