Define the free product (#104)

This commit is contained in:
Patrick Stevens
2020-03-28 21:34:14 +00:00
committed by GitHub
parent 162a1c7a40
commit a27375db4e
21 changed files with 1326 additions and 18 deletions

View File

@@ -1,6 +1,9 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import Sets.FinSet.Definition
open import LogicalFormulae
open import Groups.Abelian.Definition
open import Groups.Vector
open import Groups.Definition
open import Groups.Abelian.Definition
open import Numbers.Integers.Integers
@@ -9,7 +12,20 @@ open import Sets.EquivalenceRelations
open import Modules.Definition
open import Groups.Cyclic.Definition
open import Groups.Cyclic.DefinitionLemmas
open import Rings.Polynomial.Multiplication
open import Rings.Polynomial.Ring
open import Rings.Definition
open import Rings.Lemmas
open import Groups.Polynomials.Definition
open import Numbers.Naturals.Semiring
open import Vectors
open import Modules.Span
open import Numbers.Integers.Definition
open import Numbers.Integers.Addition
open import Rings.IntegralDomains.Definition
open import Numbers.Naturals.Order
open import Numbers.Modulo.Definition
open import Numbers.Modulo.Group
module Modules.Examples where
@@ -24,3 +40,107 @@ Module.dotDistributesRight (abGrpModule {S = S} {G' = G'} G) {r} {s} {x} = symme
open Equivalence (Setoid.eq S)
Module.dotAssociative (abGrpModule {G' = G'} G) {r} {s} {x} = elementPowerMultiplies G' r s x
Module.dotIdentity (abGrpModule {G' = G'} G) = Group.identRight G'
zAndZ : Module Ring AbGrp _*Z_
Module.dotWellDefined zAndZ refl refl = refl
Module.dotDistributesLeft zAndZ = Ring.*DistributesOver+ Ring
Module.dotDistributesRight zAndZ {x} {y} {z} = Ring.*DistributesOver+' Ring {x} {y} {z}
Module.dotAssociative zAndZ {x} {y} {z} = equalityCommutative (Ring.*Associative Ring {x} {y} {z})
Module.dotIdentity zAndZ = Ring.identIsIdent Ring
twoOrAnother : Bool
twoOrAnother _ BoolTrue = nonneg 2
twoOrAnother n BoolFalse = n
23Spans : Spans zAndZ (twoOrAnother (nonneg 3))
23Spans m = 2 , (((BoolTrue ,- (BoolFalse ,- [])) ,, (Group.inverse Group m ,- (m ,- []))) , t m)
where
open Group Group
open Ring Ring
t : (m : ) inverse m *Z nonneg 2 +Z (m *Z nonneg 3 +Z nonneg 0) m
t m rewrite identRight {m *Z nonneg 3} | *DistributesOver+ {m} {nonneg 2} {nonneg 1} | +Associative {inverse m *Z nonneg 2} {m *Z nonneg 2} {m *Z nonneg 1} = transitivity (transitivity (+WellDefined {inverse m *Z nonneg 2 +Z m *Z nonneg 2} {m *Z nonneg 1} {nonneg 0} {_} (transitivity (equalityCommutative (*DistributesOver+' {inverse m} {m} {nonneg 2})) (*WellDefined (invLeft {m}) refl)) refl) *Commutative) identIsIdent
evenNotOdd : (i : ) nonneg 0 i *Z nonneg 2 +Z nonneg 1 False
evenNotOdd (nonneg x) pr rewrite equalityCommutative (+Zinherits (x *N 2) 1) = evenNotOdd' x (nonnegInjective pr)
where
evenNotOdd' : (i : ) 0 i *N 2 +N 1 False
evenNotOdd' zero ()
evenNotOdd' (succ i) ()
evenNotOdd (negSucc zero) ()
evenNotOdd (negSucc (succ zero)) ()
evenNotOdd (negSucc (succ (succ x))) ()
2NotSpan : Spans zAndZ {_} {True} (λ _ nonneg 2) False
2NotSpan f with f (nonneg 1)
2NotSpan f | n , (trues , b) = bad (_&&_.fst trues) (_&&_.snd trues) (nonneg 0) b
where
open Group Group
open Ring Ring
bad : {n : } (t : Vec True n) (u : Vec n) (i : ) dot zAndZ (vecMap (λ _ nonneg 2) t) u ((i *Z nonneg 2) +Z nonneg 1) False
bad [] [] i x = evenNotOdd i x
bad (record {} ,- t) (u ,- us) i x = bad t us (i +Z inverse u) (transitivity (+WellDefined (equalityCommutative (invLeft {u *Z nonneg 2})) refl) (transitivity (equalityCommutative (+Associative {inverse (u *Z nonneg 2)} {u *Z nonneg 2})) (transitivity (+WellDefined (refl {x = inverse (u *Z nonneg 2)}) x) (transitivity (+Associative {inverse (u *Z nonneg 2)} {i *Z nonneg 2}) (+WellDefined {inverse (u *Z nonneg 2) +Z i *Z nonneg 2} {_} {(i +Z inverse u) *Z nonneg 2} (transitivity (transitivity (groupIsAbelian {inverse (u *Z nonneg 2)}) (applyEquality (i *Z nonneg 2 +Z_) (equalityCommutative (ringMinusExtracts' Ring)))) (equalityCommutative (*DistributesOver+' {i} {inverse u} {nonneg 2}))) (refl {x = nonneg 1}))))))
3NotSpan : Spans zAndZ {_} {True} (λ _ nonneg 3) False
3NotSpan f with f (nonneg 1)
... | n , ((trues ,, coeffs) , b) = bad trues coeffs (nonneg 0) b
where
open Group Group
open Ring Ring
bad : {n : } (t : Vec True n) (u : Vec n) (i : ) dot zAndZ (vecMap (λ _ nonneg 3) t) u ((i *Z nonneg 3) +Z nonneg 1) False
bad [] [] (nonneg zero) ()
bad [] [] (nonneg (succ i)) ()
bad [] [] (negSucc zero) ()
bad [] [] (negSucc (succ i)) ()
bad (record {} ,- ts) (u ,- us) i x = bad ts us (i +Z inverse u) (transitivity (+WellDefined (equalityCommutative (invLeft {u *Z nonneg 3})) (refl {x = dot zAndZ (vecMap (λ _ nonneg 3) ts) us})) (transitivity (equalityCommutative (+Associative {inverse (u *Z nonneg 3)} {u *Z nonneg 3})) (transitivity (+WellDefined (equalityCommutative (ringMinusExtracts' Ring {u} {nonneg 3})) x) (transitivity (+Associative {inverse u *Z nonneg 3} {i *Z nonneg 3} {nonneg 1}) (+WellDefined (transitivity {x = (inverse u *Z nonneg 3 +Z i *Z nonneg 3)} (equalityCommutative (*DistributesOver+' {inverse u} {i})) (*WellDefined (groupIsAbelian {inverse u} {i}) (refl {x = nonneg 3}))) (refl {x = nonneg 1}))))))
nothingNotSpan : Spans zAndZ {_} {False} (λ ()) False
nothingNotSpan f with f (nonneg 1)
nothingNotSpan f | zero , (([] ,, []) , ())
nothingNotSpan f | succ n , (((() ,- fst) ,, snd) , b)
1Basis : Basis zAndZ {_} {True} (λ _ nonneg 1)
_&&_.fst 1Basis m = 1 , (((record {} ,- []) ,, (m ,- [])) , transitivity (Group.+WellDefined Group (transitivity (Ring.*Commutative Ring {m}) (Ring.identIsIdent Ring)) refl) (Group.identRight Group))
_&&_.snd 1Basis {zero} r x [] x₁ = record {}
_&&_.snd 1Basis {succ zero} (record {} ,- []) x (b ,- []) pr = equalityCommutative (transitivity (equalityCommutative (transitivity (Group.+WellDefined Group (transitivity (Ring.*Commutative Ring {b}) (Ring.identIsIdent Ring)) refl) (Group.identRight Group))) pr) ,, record {}
_&&_.snd 1Basis {succ (succ n)} (record {} ,- (record {} ,- r)) x b pr with x {0} {1} (succIsPositive _) (succPreservesInequality (succIsPositive _)) refl
_&&_.snd 1Basis {succ (succ n)} (record {} ,- (record {} ,- r)) x b pr | ()
2Independent : Independent zAndZ {_} {True} (λ _ nonneg 2)
2Independent {zero} r x [] x₁ = record {}
2Independent {succ zero} (record {} ,- []) x (coeff ,- []) pr rewrite Group.identRight Group {coeff *Z nonneg 2} = equalityCommutative (IntegralDomain.intDom IntDom (transitivity (Ring.*Commutative Ring) pr) λ ()) ,, record {}
2Independent {succ (succ n)} (record {} ,- (record {} ,- rest)) pr b x = exFalso (naughtE t)
where
t : 0 1
t = pr {0} {1} (succIsPositive (succ n)) (succPreservesInequality (succIsPositive n)) refl
2AndExtraNotIndependent : (n : ) Independent zAndZ (twoOrAnother n) False
2AndExtraNotIndependent n indep = exFalso (naughtE (nonnegInjective ohNo))
where
r : {a b : } (a<n : a <N 2) (b<n : b <N 2) vecIndex (BoolTrue ,- (BoolFalse ,- [])) a a<n vecIndex (BoolTrue ,- (BoolFalse ,- [])) b b<n a b
r {zero} {zero} a<n b<n x = refl
r {zero} {succ (succ b)} a<n (le (succ zero) ()) x
r {zero} {succ (succ b)} a<n (le (succ (succ y)) ()) x
r {succ zero} {succ zero} a<n b<n x = refl
r {succ zero} {succ (succ b)} a<n (le (succ zero) ()) x
r {succ zero} {succ (succ b)} a<n (le (succ (succ i)) ()) x
r {succ (succ a)} {b} (le (succ zero) ()) b<n x
r {succ (succ a)} {b} (le (succ (succ i)) ()) b<n x
s : dot zAndZ (vecMap (twoOrAnother n) (BoolTrue ,- (BoolFalse ,- []))) (Group.inverse Group n ,- (nonneg 2 ,- [])) nonneg 0
s rewrite Group.identRight Group {nonneg 2 *Z n} | Ring.*Commutative Ring {nonneg 2} {n} | equalityCommutative (Ring.*DistributesOver+' Ring {additiveInverse n} {n} {nonneg 2}) | Group.invLeft Group {n} = refl
t : _=V_ zAndZ (Ring.additiveGroup Ring) (nonneg 0 ,- (nonneg 0 ,- [])) (Group.inverse Group n ,- (nonneg 2 ,- []))
t = indep (BoolTrue ,- (BoolFalse ,- [])) r (Group.inverse Group n ,- (nonneg 2 ,- [])) s
ohNo : nonneg 0 nonneg 2
ohNo with t
ohNo | ()
oneLengthAlwaysInj : {a b : } (a<n : a <N 1) (b<n : b <N 1) a b
oneLengthAlwaysInj {zero} {zero} a<n b<n = refl
oneLengthAlwaysInj {zero} {succ b} a<n (le zero ())
oneLengthAlwaysInj {zero} {succ b} a<n (le (succ x) ())
oneLengthAlwaysInj {succ a} {b} (le zero ()) b<n
oneLengthAlwaysInj {succ a} {b} (le (succ x) ()) b<n
noBasisExample : Independent (abGrpModule (nAbGroup 5 (le 4 refl))) (λ (t : True) record { x = 1 ; xLess = le 3 refl }) False
noBasisExample ind with ind (record {} ,- []) (λ a<n b<n _ oneLengthAlwaysInj a<n b<n) (nonneg 5 ,- []) refl
noBasisExample ind | ()

View File

@@ -0,0 +1,104 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import Sets.EquivalenceRelations
open import LogicalFormulae
open import Setoids.Setoids
open import Rings.Definition
open import Modules.Definition
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Lists.Lists
open import Vectors
open import Groups.Definition
open import Orders.Total.Definition
open import Groups.Homomorphisms.Definition
open import Groups.Homomorphisms.Lemmas
open import Rings.Homomorphisms.Definition
module Modules.PolynomialModule {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
open Ring R
open import Groups.Polynomials.Definition additiveGroup
open import Groups.Polynomials.Group additiveGroup
open import Rings.Polynomial.Multiplication R
open import Rings.Polynomial.Evaluation R
open import Rings.Polynomial.Ring R
open import Rings.Lemmas polyRing
polynomialRingModule : Module R abelianUnderlyingGroup (λ a b _*P_ (polyInjection a) b)
Module.dotWellDefined (polynomialRingModule) r=s t=u = Ring.*WellDefined polyRing (r=s ,, record {}) t=u
Module.dotDistributesLeft (polynomialRingModule) {r} {s} {t} = Ring.*DistributesOver+ polyRing {polyInjection r} {s} {t}
Module.dotDistributesRight (polynomialRingModule) {r} {s} {t} = Ring.*DistributesOver+' polyRing {polyInjection r} {polyInjection s} {t}
Module.dotAssociative (polynomialRingModule) {r} {s} {t} = Equivalence.transitive (Setoid.eq naivePolySetoid) (Ring.*WellDefined polyRing m (Equivalence.reflexive (Setoid.eq naivePolySetoid) {t})) (Equivalence.symmetric (Setoid.eq naivePolySetoid) (Ring.*Associative polyRing {polyInjection r} {polyInjection s} {t}))
where
m : Setoid.__ naivePolySetoid ((r * s) :: []) ((r * s) :: (([] +P []) +P (0R :: [])))
m = Equivalence.reflexive (Setoid.eq S) ,, (Equivalence.reflexive (Setoid.eq S) ,, record {})
Module.dotIdentity (polynomialRingModule) = Ring.identIsIdent polyRing
open import Modules.Span polynomialRingModule
polynomialBasis : NaivePoly
polynomialBasis zero = polyInjection 1R
polynomialBasis (succ a) = 0R :: polynomialBasis a
count : (n : ) Vec n
count zero = []
count (succ n) = 0 ,- vecMap succ (count n)
lemma : {d : _} {D : Set d} (f : D List A) {n : } (m : Vec D n) (r : Vec A n) Setoid.__ naivePolySetoid (dot (vecMap (λ i 0R :: f i) m) r) (0R :: dot (vecMap f m) r)
lemma f [] [] = reflexive ,, record {}
where
open Setoid S
open Equivalence eq
lemma f (m ,- ms) (r ,- rs) rewrite refl {x = 0} = transitive (+WellDefined {(r * 0R) :: (((map (_*_ r) (f m)) +P []) +P (0R :: []))} {_} {(r * 0R) :: (((map (_*_ r) (f m)) +P []) +P (0R :: []))} (reflexive {(r * 0R) :: (((map (_*_ r) (f m)) +P []) +P (0R :: []))}) (lemma f ms rs)) (Equivalence.transitive (Setoid.eq S) (Group.identRight additiveGroup) (Ring.timesZero R) ,, +WellDefined {((map (_*_ r) (f m)) +P []) +P (0R :: [])} {dot (vecMap f ms) rs} {(r :: []) *P f m} {dot (vecMap f ms) rs} t (reflexive {dot (vecMap f ms) rs}))
where
open Setoid naivePolySetoid
open Equivalence eq
open Group polyGroup
lemm : (v : List A) polysEqual (map (_*_ r) v) ((r :: []) *P v)
lemm [] = record {}
lemm (x :: v) = Equivalence.reflexive (Setoid.eq S) ,, symmetric (transitive (+WellDefined {map (_*_ r) v +P []} {0R :: []} {_} {[]} reflexive (Equivalence.reflexive (Setoid.eq S) ,, record {})) (transitive (identRight {(map (_*_ r) v) +P []}) (identRight {map (_*_ r) v})))
t : ((map (_*_ r) (f m) +P []) +P (0R :: [])) ((r :: []) *P f m)
t = transitive (+WellDefined {map (_*_ r) (f m) +P []} {0R :: []} {_} {[]} reflexive (Equivalence.reflexive (Setoid.eq S) ,, record {})) (transitive (identRight {map (_*_ r) (f m) +P []}) (transitive (identRight {map (_*_ r) (f m)}) (lemm (f m))))
identityMap : (v : List A) Setoid.__ naivePolySetoid (dot (vecMap polynomialBasis (count (length v))) (listToVec v)) v
identityMap [] = record {}
identityMap (x :: v) rewrite vecMapCompose succ polynomialBasis (count (length v)) = transitive (+WellDefined {(x * 1R) :: Group.0G additiveGroup :: []} {dot (vecMap (λ i Group.0G additiveGroup :: polynomialBasis i) (count (length v))) (listToVec v)} {(x * 1R) :: 0R :: []} {0R :: dot (vecMap polynomialBasis (count (length v))) (listToVec v)} (reflexive {(x * 1R) :: 0R :: []}) (lemma polynomialBasis (count (length v)) (listToVec v))) (Equivalence.transitive (Setoid.eq S) (Group.identRight additiveGroup) (Equivalence.transitive (Setoid.eq S) *Commutative identIsIdent) ,, transitive (+WellDefined {0R :: []} {dot (vecMap polynomialBasis (count (length v))) (listToVec v)} {[]} (Equivalence.reflexive (Setoid.eq S) ,, record {}) reflexive) (transitive identLeft (identityMap v)))
where
open Group polyGroup
open Setoid naivePolySetoid
open Equivalence eq
polynomialBasisSpans : Spans polynomialBasis
polynomialBasisSpans m = length m , ((count (length m) ,, listToVec m) , identityMap m)
{-
private
indepWithZero : {n : } (rs : Vec n) (indicesDistinct : {a b : } → (a<n : a <N succ n) (b<n : b <N succ n) → vecIndex (0 ,- rs) a a<n ≡ vecIndex (0 ,- rs) b b<n → a ≡ b) (b : A) (bs : Vec A n) (dotZero : Setoid.__ naivePolySetoid (((b * 1R) :: 0R :: []) +P (dot (vecMap polynomialBasis rs) bs)) []) → (nonzero : {a : } → (a<n : a <N n) → 0 <N vecIndex rs a a<n) → Setoid.__ S 0R b && _=V_ additiveGroup (vecPure 0R) bs
indepWithZero rs indicesDistinct b bs dotZero indicesNonzero = symmetric b=0 ,, {!!}
where
open Setoid S
open Equivalence eq
open Group additiveGroup
t : (inducedFunction (((b * 1R) :: 0R :: []) +P (dot (vecMap polynomialBasis rs) bs)) 0R) 0R
t = inducedFunctionWellDefined dotZero 0R
u : ((inducedFunction ((b * 1R) :: 0R :: []) 0R) + inducedFunction (dot (vecMap polynomialBasis rs) bs) 0R) 0R
u = transitive (symmetric (GroupHom.groupHom (RingHom.groupHom (inducedFunctionIsHom 0R)) {((b * 1R) :: 0R :: [])} {dot (vecMap polynomialBasis rs) bs})) t
b=0 : b 0R
b=0 = transitive (symmetric (transitive (+WellDefined (transitive (transitive (transitive (+WellDefined reflexive (transitive *Commutative timesZero)) identRight) *Commutative) identIsIdent) {!imageOfIdentityIsIdentity (RingHom.groupHom (inducedFunctionIsHom 0R))!}) (identRight {b}))) u
polynomialBasisIndependent : Independent polynomialBasis
polynomialBasisIndependent [] indicesDistinct [] dotZero = record {}
polynomialBasisIndependent {succ n} (zero ,- rs) indicesDistinct (b ,- bs) dotZero = indepWithZero rs indicesDistinct b bs dotZero t
where
t : {a : } (a<n : a <N n) → 0 <N vecIndex rs a a<n
t {a} a<n with TotalOrder.totality TotalOrder 0 (vecIndex rs a a<n)
t {a} a<n | inl (inl x) = x
t {a} a<n | inr x with indicesDistinct {succ a} {0} (succPreservesInequality a<n) (succIsPositive n) (equalityCommutative x)
... | ()
polynomialBasisIndependent (succ r ,- rs) indicesDistinct (b ,- bs) dotZero = {!!}
where
rearr : {!!}
rearr = {!!}
-}

54
Modules/Span.agda Normal file
View File

@@ -0,0 +1,54 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Functions
open import Groups.Definition
open import Groups.Abelian.Definition
open import Setoids.Setoids
open import Rings.Definition
open import Sets.FinSet.Definition
open import Vectors
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Sets.EquivalenceRelations
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Modules.Definition
module Modules.Span {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+R_ : A A A} {_*_ : A A A} {R : Ring S _+R_ _*_} {m n : _} {M : Set m} {T : Setoid {m} {n} M} {_+_ : M M M} {G' : Group T _+_} {G : AbelianGroup G'} {_·_ : A M M} (mod : Module R G _·_) where
open Group G'
open Setoid T
open Equivalence eq
_=V_ : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ : A A A} (G : Group S _+_) {n1 : } Rel {a} {b n} (Vec A n1)
_=V_ G [] [] = True'
_=V_ {S = S} G (x ,- xs) (y ,- ys) = (Setoid.__ S x y) && (_=V_ G xs ys)
dot : {n : } (Vec M n) (Vec A n) M
dot [] [] = 0G
dot (v ,- vs) (x ,- xs) = (x · v) + (dot vs xs)
Spans : {c : _} {C : Set c} (f : C M) Set (a m n c)
Spans {C = C} f = (m : M) Sg (λ n Sg ((Vec C n) && (Vec A n)) (λ t (dot (vecMap f (_&&_.fst t)) (_&&_.snd t)) m))
Independent : {c : _} {C : Set c} (f : C M) Set (a b n c)
Independent {C = C} f = {n : } (r : Vec C n) ({a b : } (a<n : a <N n) (b<n : b <N n) vecIndex r a a<n vecIndex r b b<n a b) (b : Vec A n) (dot (vecMap f r) b) 0G _=V_ (Ring.additiveGroup R) (vecPure (Group.0G (Ring.additiveGroup R))) b
independentSubset : {c : _} {C : Set c} (f : C M) {d : _} {D : Set d} {inj : D C} (isInj : Injection inj) Independent f Independent (f inj)
independentSubset f {inj = inj} isInj indp {n = n} r coeffInj coeffs dotZero = indp {n = n} (vecMap inj r) inj' coeffs (transitive (identityOfIndiscernablesRight __ reflexive (applyEquality (λ i dot i coeffs) (vecMapCompose inj f r))) dotZero)
where
inj' : {a b : } (a<n : a <N n) (b<n : b <N n) vecIndex (vecMap inj r) a a<n vecIndex (vecMap inj r) b b<n a b
inj' a<n b<n x rewrite vecMapAndIndex r inj a<n | vecMapAndIndex r inj b<n = coeffInj a<n b<n (isInj x)
spanSuperset : {c : _} {C : Set c} (f : C M) {d : _} {D : Set d} {surj : D C} (isSurj : Surjection surj) Spans f Spans (f surj)
spanSuperset f {surj = surj} isSurj spans m with spans m
spanSuperset {C = C} f {surj = surj} isSurj spans m | n , ((coeffs ,, basis) , b) = n , ((vecMap (λ c underlying (isSurj c)) coeffs ,, basis) , transitive (identityOfIndiscernablesLeft __ reflexive (applyEquality (λ i dot i basis) (equalityCommutative {x = vecMap (λ i f (surj i)) (vecMap (λ c underlying (isSurj c)) coeffs)} {vecMap f coeffs} (transitivity (vecMapCompose (λ i underlying (isSurj i)) (λ z f (surj z)) coeffs) (t coeffs))))) b)
where
t : {n : } (coeffs : Vec C n) vecMap (λ i f (surj (underlying (isSurj i)))) coeffs vecMap f coeffs
t [] = refl
t (x ,- coeffs) with isSurj x
... | img , pr rewrite pr | t coeffs = refl
Basis : {c : _} {C : Set c} (f : C M) Set (a b m n c)
Basis v = Spans v && Independent v

73
Modules/VectorModule.agda Normal file
View File

@@ -0,0 +1,73 @@
{-# OPTIONS --safe --warning=error --without-K #-}
open import LogicalFormulae
open import Setoids.Setoids
open import Rings.Definition
open import Numbers.Naturals.Semiring
open import Numbers.Naturals.Order
open import Modules.Definition
open import Vectors
open import Sets.EquivalenceRelations
open import Sets.FinSet.Definition
open import Modules.Span
open import Groups.Definition
module Modules.VectorModule {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) where
open import Rings.Definition
open import Rings.Lemmas R
open import Groups.Vector (Ring.additiveGroup R)
open Ring R
open Group additiveGroup
open Setoid S
open Equivalence eq
ringModule : {n : } Module R (abelianVectorGroup {n = n} abelianUnderlyingGroup) λ a v vecMap (_* a) v
Module.dotWellDefined (ringModule) {r} {s} {[]} {[]} r=s t=u = record {}
Module.dotWellDefined (ringModule) {r} {s} {t ,- ts} {u ,- us} r=s (t=u ,, ts=us) = Ring.*WellDefined R t=u r=s ,, Module.dotWellDefined (ringModule) {r} {s} {ts} {us} r=s ts=us
Module.dotDistributesLeft (ringModule) {a} {[]} {[]} = record {}
Module.dotDistributesLeft (ringModule) {a} {y ,- ys} {z ,- zs} = Ring.*DistributesOver+' R ,, Module.dotDistributesLeft (ringModule) {a} {ys} {zs}
Module.dotDistributesRight (ringModule) {a} {b} {[]} = record {}
Module.dotDistributesRight (ringModule) {a} {b} {z ,- zs} = Ring.*DistributesOver+ R ,, Module.dotDistributesRight (ringModule) {a} {b} {zs}
Module.dotAssociative (ringModule) {r} {s} {[]} = record {}
Module.dotAssociative (ringModule) {r} {s} {x ,- xs} = transitive (Ring.*WellDefined R reflexive (Ring.*Commutative R)) (Ring.*Associative R) ,, Module.dotAssociative (ringModule) {r} {s} {xs}
Module.dotIdentity (ringModule) {[]} = record {}
Module.dotIdentity (ringModule) {x ,- xs} = transitive (Ring.*Commutative R) (Ring.identIsIdent R) ,, Module.dotIdentity ringModule {xs}
vecModuleBasis : {n : } FinSet n Vec A n
vecModuleBasis {succ n} fzero = (Ring.1R R) ,- vecPure (Ring.0R R)
vecModuleBasis {succ n} (fsucc i) = (Ring.0R R) ,- vecModuleBasis i
vecTimesZero : {n : } (a : A) vecEquiv {n} (vecMap (_* a) (vecPure 0G)) (vecPure 0G)
vecTimesZero {zero} a = record {}
vecTimesZero {succ n} a = transitive *Commutative timesZero ,, vecTimesZero a
--private
-- lemma2 : {n : } → (f : _ → _) → (j : _) → vecEquiv (dot ringModule (vecMap (λ i → 0G ,- f i)) j) ?
-- lemma2 = ?
allEntries : (n : ) Vec (FinSet n) n
allEntries zero = []
allEntries (succ n) = fzero ,- vecMap fsucc (allEntries n)
extractZero : {n m : } (y : Vec A m) (xs : Vec (FinSet n) m) (f : FinSet n Vec A n) vecEquiv (dot ringModule (vecMap (λ i 0G ,- f i) xs) y) (0G ,- dot ringModule (vecMap f xs) y)
extractZero {zero} {zero} [] [] f = Equivalence.reflexive (Setoid.eq (vectorSetoid _))
extractZero {zero} {succ m} (y ,- ys) (() ,- xs) f
extractZero {succ n} {zero} [] [] f = reflexive ,, (reflexive ,, Equivalence.reflexive (Setoid.eq (vectorSetoid _)))
extractZero {succ n} {succ m} (y ,- ys) (x ,- xs) f = Equivalence.transitive (Setoid.eq (vectorSetoid _)) (Group.+WellDefined vectorGroup (transitive *Commutative timesZero ,, Equivalence.reflexive (Setoid.eq (vectorSetoid _))) (extractZero ys xs f)) (identLeft ,, Equivalence.reflexive (Setoid.eq (vectorSetoid _)))
identityMatrixProduct : {n : } (b : Vec A n) vecEquiv (dot ringModule (vecMap vecModuleBasis (allEntries n)) b) b
identityMatrixProduct [] = record {}
identityMatrixProduct {succ n} (x ,- b) rewrite vecMapCompose fsucc vecModuleBasis (allEntries n) = Equivalence.transitive (Setoid.eq (vectorSetoid _)) (Group.+WellDefined vectorGroup (identIsIdent ,, vecTimesZero x) (extractZero b (allEntries n) vecModuleBasis)) (identRight ,, Equivalence.transitive (Setoid.eq (vectorSetoid _)) (Group.identLeft vectorGroup) (identityMatrixProduct b))
rearrangeBasis : {n : } (coeffs : Vec A n) (r : Vec (FinSet (succ n)) n) (eq : vecEquiv (dot ringModule (vecMap vecModuleBasis r) coeffs) (vecPure 0G)) (uniq : {a : } {b : } (a<n : a <N n) (b<n : b <N n) vecIndex r a a<n vecIndex r b b<n a b) {!!}
rearrangeBasis coeffs r eq uniq = {!!}
vecModuleBasisSpans : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) {n : } Spans (ringModule {n}) (vecModuleBasis {n})
vecModuleBasisSpans R {n} m = n , ((allEntries n ,, m) , identityMatrixProduct m)
vecModuleBasisIndependent : {a b : _} {A : Set a} {S : Setoid {a} {b} A} {_+_ _*_ : A A A} (R : Ring S _+_ _*_) {n : } Independent (ringModule {n}) (vecModuleBasis {n})
vecModuleBasisIndependent R {zero} [] _ [] x = record {}
vecModuleBasisIndependent R {succ n} r _ [] x = record {}
vecModuleBasisIndependent R {succ n} r uniq coeffs x = {!!}