summaryrefslogtreecommitdiff
path: root/src/Soat/FirstOrder/Algebra/Coproduct.idr
diff options
context:
space:
mode:
Diffstat (limited to 'src/Soat/FirstOrder/Algebra/Coproduct.idr')
-rw-r--r--src/Soat/FirstOrder/Algebra/Coproduct.idr150
1 files changed, 73 insertions, 77 deletions
diff --git a/src/Soat/FirstOrder/Algebra/Coproduct.idr b/src/Soat/FirstOrder/Algebra/Coproduct.idr
index 5cb4c45..a8e4c0e 100644
--- a/src/Soat/FirstOrder/Algebra/Coproduct.idr
+++ b/src/Soat/FirstOrder/Algebra/Coproduct.idr
@@ -1,6 +1,5 @@
module Soat.FirstOrder.Algebra.Coproduct
-import Data.Morphism.Indexed
import Data.Setoid.Indexed
import Data.Setoid.Either
@@ -17,7 +16,9 @@ CoproductSet : (0 sig : Signature) -> (0 x, y : sig.T -> Type) -> sig.T -> Type
CoproductSet sig x y = Term sig (\i => Either (x i) (y i))
public export
-data (~=~) : (0 x, y : RawSetoidAlgebra sig) -> IRel (CoproductSet sig x.raw.U y.raw.U) where
+data (~=~) : (0 x, y : RawSetoidAlgebra sig) -> (t : sig.T)
+ -> Rel (CoproductSet sig x.raw.U y.raw.U t)
+ where
DoneL : {0 x, y : RawSetoidAlgebra sig}
-> {t : sig.T} -> {i, j : x.raw.U t} -> x.relation t i j
-> (~=~) x y t (Done $ Left i) (Done $ Left j)
@@ -36,31 +37,29 @@ data (~=~) : (0 x, y : RawSetoidAlgebra sig) -> IRel (CoproductSet sig x.raw.U y
-> (~=~) x y t tm tm' -> (~=~) x y t tm' tm'' -> (~=~) x y t tm tm''
parameters {0 x, y : RawSetoidAlgebra sig}
- coprodRelRefl : IReflexive x.raw.U x.relation -> IReflexive y.raw.U y.relation
+ coprodRelRefl : (xRefl : (t : sig.T) -> (a : x.raw.U t) -> x.relation t a a)
+ -> (yRefl : (t : sig.T) -> (a : y.raw.U t) -> y.relation t a a)
-> {t : _} -> (tm : _) -> (~=~) x y t tm tm
- coprodsRelRefl : IReflexive x.raw.U x.relation -> IReflexive y.raw.U y.relation
+ coprodsRelRefl : (xRefl : (t : sig.T) -> (a : x.raw.U t) -> x.relation t a a)
+ -> (yRefl : (t : sig.T) -> (a : y.raw.U t) -> y.relation t a a)
-> {ts : _} -> (tms : _ ^ ts) -> Pointwise ((~=~) x y) tms tms
- coprodRelRefl x y (Done (Left i)) = DoneL $ reflexive @{x _}
- coprodRelRefl x y (Done (Right i)) = DoneR $ reflexive @{y _}
+ coprodRelRefl x y (Done (Left i)) = DoneL $ x _ i
+ coprodRelRefl x y (Done (Right i)) = DoneR $ y _ i
coprodRelRefl x y (Call op ts) = Call op $ coprodsRelRefl x y ts
coprodsRelRefl x y [] = []
coprodsRelRefl x y (t :: ts) = coprodRelRefl x y t :: coprodsRelRefl x y ts
- coprodsRelReflexive : IReflexive x.raw.U x.relation -> IReflexive y.raw.U y.relation
- -> {ts : _} -> {tms, tms' : _ ^ ts} -> tms = tms' -> Pointwise ((~=~) x y) tms tms'
- coprodsRelReflexive x y Refl = coprodsRelRefl x y _
-
- tmRelImpliesCoprodRel : Term.Rel.(~=~) (\i => Either (x.relation i) (y.relation i)) t tm tm'
+ tmRelImpliesCoprodRel : Term.Rel.(~=~) (\i => x.relation i `Or` y.relation i) t tm tm'
-> (~=~) x y t tm tm'
tmRelsImpliesCoprodRels : {0 tms, tms' : _ ^ ts}
- -> Pointwise (Term.Rel.(~=~) (\i => Either (x.relation i) (y.relation i))) tms tms'
+ -> Pointwise (Term.Rel.(~=~) (\i => x.relation i `Or` y.relation i)) tms tms'
-> Pointwise ((~=~) x y) tms tms'
- tmRelImpliesCoprodRel (Done' (Left eq)) = DoneL eq
- tmRelImpliesCoprodRel (Done' (Right eq)) = DoneR eq
- tmRelImpliesCoprodRel (Call' op eqs) = Call op $ tmRelsImpliesCoprodRels eqs
+ tmRelImpliesCoprodRel (Done (Left eq)) = DoneL eq
+ tmRelImpliesCoprodRel (Done (Right eq)) = DoneR eq
+ tmRelImpliesCoprodRel (Call op eqs) = Call op $ tmRelsImpliesCoprodRels eqs
tmRelsImpliesCoprodRels [] = []
tmRelsImpliesCoprodRels (eq :: eqs) = tmRelImpliesCoprodRel eq :: tmRelsImpliesCoprodRels eqs
@@ -76,17 +75,17 @@ public export
CoproductIsAlgebra : IsAlgebra sig x -> IsAlgebra sig y
-> IsAlgebra sig (Coproduct x y)
CoproductIsAlgebra xIsAlgebra yIsAlgebra = MkIsAlgebra
- { relation =
- (~=~)
- (MkRawSetoidAlgebra x xIsAlgebra.relation)
- (MkRawSetoidAlgebra y yIsAlgebra.relation)
- , equivalence = \t => MkEquivalence
- { refl = MkReflexive $ coprodRelRefl
- (\t => (xIsAlgebra.equivalence t).refl)
- (\t => (yIsAlgebra.equivalence t).refl)
- _
- , sym = MkSymmetric $ Sym
- , trans = MkTransitive $ Trans
+ { equivalence = MkIndexedEquivalence
+ { relation =
+ (~=~)
+ (MkRawSetoidAlgebra x xIsAlgebra.equivalence.relation)
+ (MkRawSetoidAlgebra y yIsAlgebra.equivalence.relation)
+ , reflexive = \_ =>
+ coprodRelRefl
+ xIsAlgebra.equivalence.reflexive
+ yIsAlgebra.equivalence.reflexive
+ , symmetric = \_, _, _ => Sym
+ , transitive = \_, _, _, _ => Trans
}
, semCong = Call
}
@@ -98,87 +97,84 @@ CoproductAlgebra x y = MkAlgebra (Coproduct x.raw y.raw) $ CoproductIsAlgebra x.
public export
injectLHomo : x ~> CoproductAlgebra x y
injectLHomo = MkHomomorphism
- { func = \_ => Done . Left
- , cong = \_ => DoneL
+ { func = MkIndexedSetoidHomomorphism (\_ => Done . Left) (\_, _, _ => DoneL)
, semHomo = InjectL
}
public export
injectRHomo : y ~> CoproductAlgebra x y
injectRHomo = MkHomomorphism
- { func = \_ => Done . Right
- , cong = \_ => DoneR
+ { func = MkIndexedSetoidHomomorphism (\_ => Done . Right) (\_, _, _ => DoneR)
, semHomo = InjectR
}
public export
-coproduct : {z : RawAlgebra sig} -> IFunc x z.U -> IFunc y z.U -> IFunc (CoproductSet sig x y) z.U
+coproduct : {z : RawAlgebra sig}
+ -> (f : (t : sig.T) -> x t -> z.U t) -> (g : (t : sig.T) -> y t -> z.U t)
+ -> (t : sig.T) -> CoproductSet sig x y t -> z.U t
coproduct f g _ = bindTerm (\i => either (f i) (g i))
public export
-coproducts : {z : RawAlgebra sig} -> IFunc x z.U -> IFunc y z.U -> (ts : _)
- -> CoproductSet sig x y ^ ts -> z.U ^ ts
+coproducts : {z : RawAlgebra sig}
+ -> (f : (t : sig.T) -> x t -> z.U t) -> (g : (t : sig.T) -> y t -> z.U t)
+ -> (ts : List sig.T) -> CoproductSet sig x y ^ ts -> z.U ^ ts
coproducts f g _ = bindTerms (\i => either (f i) (g i))
public export
coproductCong : {x, y, z : Algebra sig} -> (f : x ~> z) -> (g : y ~> z)
- -> (t : sig.T) -> {tm, tm' : _} -> (~=~) x.rawSetoid y.rawSetoid t tm tm'
- -> (z.relation t `on` coproduct {z = z.raw} f.func g.func t) tm tm'
+ -> {t : sig.T} -> {tm, tm' : _} -> (~=~) x.rawSetoid y.rawSetoid t tm tm'
+ -> (z.relation t `on` coproduct {z = z.raw} f.func.H g.func.H t) tm tm'
public export
coproductsCong : {x, y, z : Algebra sig} -> (f : x ~> z) -> (g : y ~> z)
- -> (ts : List sig.T) -> {tms, tms' : _ ^ ts}
+ -> {ts : List sig.T} -> {tms, tms' : _ ^ ts}
-> Pointwise ((~=~) x.rawSetoid y.rawSetoid) tms tms'
- -> (Pointwise z.relation `on` (coproducts {z = z.raw} f.func g.func ts)) tms tms'
-
-coproductCong f g _ (DoneL eq) = f.cong _ eq
-coproductCong f g _ (DoneR eq) = g.cong _ eq
-coproductCong f g _ (Call op eqs) =
- z.algebra.semCong op $
- coproductsCong f g _ eqs
-coproductCong f g _ (InjectL op ts) = CalcWith (z.setoid.index _) $
- |~ f.func _ (x.raw.sem op ts)
- ~~ z.raw.sem op (map f.func ts) ...(f.semHomo op ts)
- ~~ z.raw.sem op (map (coproduct f.func g.func) (map (\_ => Done . Left) ts)) .=.(cong (z.raw.sem op) $ mapComp ts)
- ~~ z.raw.sem op (coproducts f.func g.func _ (map (\_ => Done . Left) ts)) .=<(cong (z.raw.sem op) $ bindTermsIsMap {a = z.raw} _ _)
-coproductCong f g _ (InjectR op ts) = CalcWith (z.setoid.index _) $
- |~ g.func _ (y.raw.sem op ts)
- ~~ z.raw.sem op (map g.func ts) ...(g.semHomo op ts)
- ~~ z.raw.sem op (map (coproduct f.func g.func) (map (\_ => Done . Right) ts)) .=.(cong (z.raw.sem op) $ mapComp ts)
- ~~ z.raw.sem op (coproducts f.func g.func _ (map (\_ => Done . Right) ts)) .=<(cong (z.raw.sem op) $ bindTermsIsMap {a = z.raw} _ _)
-coproductCong f g _ (Sym eq) = (z.algebra.equivalence _).symmetric $ coproductCong f g _ eq
-coproductCong f g _ (Trans eq eq') = (z.algebra.equivalence _).transitive
- (coproductCong f g _ eq)
- (coproductCong f g _ eq')
-
-coproductsCong f g _ [] = []
-coproductsCong f g _ (eq :: eqs) =
- coproductCong f g _ eq :: coproductsCong f g _ eqs
+ -> (Pointwise z.relation `on` (coproducts {z = z.raw} f.func.H g.func.H ts)) tms tms'
+
+coproductCong f g (DoneL eq) = f.func.homomorphic _ _ _ eq
+coproductCong f g (DoneR eq) = g.func.homomorphic _ _ _ eq
+coproductCong f g (Call op eqs) = z.algebra.semCong op $ coproductsCong f g eqs
+coproductCong f g (InjectL op ts) = CalcWith (index z.setoid _) $
+ |~ f.func.H _ (x.raw.sem op ts)
+ ~~ z.raw.sem op (map f.func.H ts) ...(f.semHomo op ts)
+ ~~ z.raw.sem op (map (coproduct f.func.H g.func.H) (map (\_ => Done . Left) ts)) .=.(cong (z.raw.sem op) $ mapComp ts)
+ ~~ z.raw.sem op (coproducts f.func.H g.func.H _ (map (\_ => Done . Left) ts)) .=<(cong (z.raw.sem op) $ bindTermsIsMap {a = z.raw} _ _)
+coproductCong f g (InjectR op ts) = CalcWith (index z.setoid _) $
+ |~ g.func.H _ (y.raw.sem op ts)
+ ~~ z.raw.sem op (map g.func.H ts) ...(g.semHomo op ts)
+ ~~ z.raw.sem op (map (coproduct f.func.H g.func.H) (map (\_ => Done . Right) ts)) .=.(cong (z.raw.sem op) $ mapComp ts)
+ ~~ z.raw.sem op (coproducts f.func.H g.func.H _ (map (\_ => Done . Right) ts)) .=<(cong (z.raw.sem op) $ bindTermsIsMap {a = z.raw} _ _)
+coproductCong f g (Sym eq) = z.algebra.equivalence.symmetric _ _ _ $ coproductCong f g eq
+coproductCong f g (Trans eq eq') = z.algebra.equivalence.transitive _ _ _ _
+ (coproductCong f g eq)
+ (coproductCong f g eq')
+
+coproductsCong f g [] = []
+coproductsCong f g (eq :: eqs) = coproductCong f g eq :: coproductsCong f g eqs
public export
coproductHomo : {x, y, z : Algebra sig} -> x ~> z -> y ~> z
-> CoproductAlgebra x y ~> z
coproductHomo f g = MkHomomorphism
- { func = coproduct {z = z.raw} f.func g.func
- , cong = coproductCong f g
+ { func = MkIndexedSetoidHomomorphism
+ { H = coproduct {z = z.raw} f.func.H g.func.H
+ , homomorphic = \_, _, _ => coproductCong f g
+ }
, semHomo = \op, tms =>
- (z.algebra.equivalence _).equalImpliesEq $
+ reflect (index z.setoid _) $
cong (z.raw.sem op) $
bindTermsIsMap {a = z.raw} _ tms
}
public export
termToCoproduct : (x, y : Algebra sig)
- -> FreeAlgebra (fromIndexed (\i => EitherSetoid (x.setoid.index i) (y.setoid.index i))) ~>
+ -> FreeAlgebra (bundle (\t => Either (index x.setoid t) (index y.setoid t))) ~>
CoproductAlgebra x y
termToCoproduct x y = MkHomomorphism
- { func = \_ => id
- , cong = \t => tmRelImpliesCoprodRel
+ { func = MkIndexedSetoidHomomorphism (\_ => id) (\t, _, _ => tmRelImpliesCoprodRel)
, semHomo = \op, tms =>
Call op $
- coprodsRelReflexive
- (\i => (x.algebra.equivalence i).refl)
- (\i => (y.algebra.equivalence i).refl) $
+ reflect (index (pwSetoid (CoproductAlgebra x y).setoid) _) $
sym $
mapId tms
}
@@ -188,13 +184,13 @@ coproductUnique' : {x, y, z : Algebra sig}
-> (f : CoproductAlgebra x y ~> z)
-> (g : CoproductAlgebra x y ~> z)
-> (congL : {t : sig.T} -> (i : x.raw.U t)
- -> z.relation t (f.func t (Done (Left i))) (g.func t (Done (Left i))))
+ -> z.relation t (f.func.H t (Done (Left i))) (g.func.H t (Done (Left i))))
-> (congR : {t : sig.T} -> (i : y.raw.U t)
- -> z.relation t (f.func t (Done (Right i))) (g.func t (Done (Right i))))
+ -> z.relation t (f.func.H t (Done (Right i))) (g.func.H t (Done (Right i))))
-> {t : sig.T} -> (tm : _)
- -> z.relation t (f.func t tm) (g.func t tm)
+ -> z.relation t (f.func.H t tm) (g.func.H t tm)
coproductUnique' f g congL congR tm = bindUnique'
- { v = fromIndexed (\i => EitherSetoid (x.setoid.index i) (y.setoid.index i))
+ { v = bundle (\t => Either (index x.setoid t) (index y.setoid t))
, f = f . termToCoproduct x y
, g = g . termToCoproduct x y
, cong = \x => case x of
@@ -207,9 +203,9 @@ public export
coproductUnique : {x, y, z : Algebra sig} -> (f : x ~> z) -> (g : y ~> z)
-> (h : CoproductAlgebra x y ~> z)
-> (congL : {t : sig.T} -> (i : x.raw.U t)
- -> z.relation t (h.func t (Done (Left i))) (f.func t i))
+ -> z.relation t (h.func.H t (Done (Left i))) (f.func.H t i))
-> (congR : {t : sig.T} -> (i : y.raw.U t)
- -> z.relation t (h.func t (Done (Right i))) (g.func t i))
+ -> z.relation t (h.func.H t (Done (Right i))) (g.func.H t i))
-> {t : sig.T} -> (tm : _)
- -> z.relation t (h.func t tm) (coproduct {z = z.raw} f.func g.func t tm)
+ -> z.relation t (h.func.H t tm) (coproduct {z = z.raw} f.func.H g.func.H t tm)
coproductUnique f g h = coproductUnique' h (coproductHomo f g)