2021-03-03 16:49:32 +03:00
|
|
|
import Syntax.PreorderReasoning
|
|
|
|
|
|
|
|
-------- some notation ----------
|
|
|
|
infixr 6 .+.,:+:
|
|
|
|
infixr 7 .*.,:*:
|
|
|
|
|
|
|
|
interface Additive a where
|
|
|
|
constructor MkAdditive
|
|
|
|
(.+.) : a -> a -> a
|
|
|
|
O : a
|
|
|
|
|
|
|
|
interface Additive2 a where
|
|
|
|
constructor MkAdditive2
|
|
|
|
(:+:) : a -> a -> a
|
2021-03-03 17:36:33 +03:00
|
|
|
O2 : a
|
|
|
|
|
2021-03-03 16:49:32 +03:00
|
|
|
interface Multiplicative a where
|
|
|
|
constructor MkMultiplicative
|
|
|
|
(.*.) : a -> a -> a
|
|
|
|
I : a
|
|
|
|
|
2021-03-03 17:36:33 +03:00
|
|
|
record MonoidOver U where
|
2021-03-03 16:49:32 +03:00
|
|
|
constructor WithStruct
|
|
|
|
Unit : U
|
|
|
|
Mult : U -> U -> U
|
|
|
|
lftUnit : (x : U) -> Unit `Mult` x = x
|
|
|
|
rgtUnit : (x : U) -> x `Mult` Unit = x
|
2021-03-03 17:36:33 +03:00
|
|
|
assoc : (x,y,z : U)
|
2021-03-03 16:49:32 +03:00
|
|
|
-> x `Mult` (y `Mult` z) = (x `Mult` y) `Mult` z
|
|
|
|
|
|
|
|
record Monoid where
|
|
|
|
constructor MkMonoid
|
|
|
|
U : Type
|
|
|
|
Struct : MonoidOver U
|
|
|
|
|
|
|
|
AdditiveStruct : (m : MonoidOver a) -> Additive a
|
2021-03-03 17:36:33 +03:00
|
|
|
AdditiveStruct m = MkAdditive (Mult $ m)
|
2021-03-03 16:49:32 +03:00
|
|
|
(Unit $ m)
|
|
|
|
AdditiveStruct2 : (m : MonoidOver a) -> Additive2 a
|
2021-03-03 17:36:33 +03:00
|
|
|
AdditiveStruct2 m = MkAdditive2 (Mult $ m)
|
2021-03-03 16:49:32 +03:00
|
|
|
(Unit $ m)
|
2021-03-03 17:36:33 +03:00
|
|
|
AdditiveStructs : (ma : MonoidOver a) -> (mb : MonoidOver b)
|
2021-03-03 16:49:32 +03:00
|
|
|
-> (Additive a, Additive2 b)
|
|
|
|
AdditiveStructs ma mb = (AdditiveStruct ma, AdditiveStruct2 mb)
|
|
|
|
|
|
|
|
getAdditive : (m : Monoid) -> Additive (U m)
|
|
|
|
getAdditive m = AdditiveStruct (Struct m)
|
|
|
|
|
|
|
|
MultiplicativeStruct : (m : Monoid) -> Multiplicative (U m)
|
|
|
|
MultiplicativeStruct m = MkMultiplicative (Mult $ Struct m)
|
|
|
|
(Unit $ Struct m)
|
2021-03-03 17:36:33 +03:00
|
|
|
-----------------------------------------------------
|
2021-03-03 16:49:32 +03:00
|
|
|
|
2021-03-09 20:23:05 +03:00
|
|
|
0 Commutative : MonoidOver a -> Type
|
2021-03-03 16:49:32 +03:00
|
|
|
Commutative m = (x,y : a) -> let _ = AdditiveStruct m in
|
|
|
|
x .+. y = y .+. x
|
2021-03-03 17:36:33 +03:00
|
|
|
|
2021-03-09 20:23:05 +03:00
|
|
|
0 Commute : (Additive a, Additive2 a) => Type
|
2021-03-03 16:49:32 +03:00
|
|
|
Commute =
|
2021-03-03 17:36:33 +03:00
|
|
|
(x11,x12,x21,x22 : a) ->
|
2021-03-03 16:49:32 +03:00
|
|
|
((x11 :+: x12) .+. (x21 :+: x22))
|
|
|
|
=
|
|
|
|
((x11 .+. x21) :+: (x12 .+. x22))
|
|
|
|
|
2021-03-03 17:36:33 +03:00
|
|
|
product : (ma, mb : Monoid) -> Monoid
|
|
|
|
product ma mb =
|
|
|
|
let
|
2021-03-03 16:49:32 +03:00
|
|
|
%hint aStruct : Additive (U ma)
|
|
|
|
aStruct = getAdditive ma
|
|
|
|
%hint bStruct : Additive (U mb)
|
|
|
|
bStruct = getAdditive mb
|
2021-03-03 17:36:33 +03:00
|
|
|
|
2021-03-03 16:49:32 +03:00
|
|
|
in MkMonoid (U ma, U mb) $ WithStruct
|
|
|
|
(O, O)
|
|
|
|
(\(x1,y1), (x2, y2) => (x1 .+. x2, y1 .+. y2))
|
2021-03-03 17:36:33 +03:00
|
|
|
(\(x,y) => rewrite lftUnit (Struct ma) x in
|
2021-03-03 16:49:32 +03:00
|
|
|
rewrite lftUnit (Struct mb) y in
|
|
|
|
Refl)
|
|
|
|
(\(x,y) => rewrite rgtUnit (Struct ma) x in
|
|
|
|
rewrite rgtUnit (Struct mb) y in
|
|
|
|
Refl)
|
2021-03-03 17:36:33 +03:00
|
|
|
(\(x1,y1),(x2,y2),(x3,y3) =>
|
2021-03-03 16:49:32 +03:00
|
|
|
rewrite assoc (Struct ma) x1 x2 x3 in
|
|
|
|
rewrite assoc (Struct mb) y1 y2 y3 in
|
|
|
|
Refl)
|
|
|
|
|
2021-03-03 17:36:33 +03:00
|
|
|
EckmannHilton : forall a . (ma,mb : MonoidOver a)
|
2021-03-03 16:49:32 +03:00
|
|
|
-> let ops = AdditiveStructs ma mb in
|
|
|
|
Commute @{ops}
|
|
|
|
-> (Commutative ma, (x,y : a) -> x .+. y = x :+: y)
|
2021-03-03 17:36:33 +03:00
|
|
|
EckmannHilton ma mb prf =
|
2021-03-03 16:49:32 +03:00
|
|
|
let %hint first : Additive a
|
|
|
|
first = AdditiveStruct ma
|
|
|
|
%hint second : Additive2 a
|
|
|
|
second = AdditiveStruct2 mb in
|
|
|
|
let SameUnits : (the a O === O2)
|
|
|
|
SameUnits = Calc $
|
|
|
|
|~ O
|
|
|
|
~~ O .+. O ...(sym $ lftUnit ma O)
|
2021-03-03 17:36:33 +03:00
|
|
|
~~ (O :+: O2) .+. (O2 :+: O) ...(sym $ cong2 (.+.)
|
2021-03-03 16:49:32 +03:00
|
|
|
(rgtUnit mb O)
|
|
|
|
(lftUnit mb O))
|
|
|
|
~~ (O .+. O2) :+: (O2 .+. O) ...(prf O O2 O2 O)
|
|
|
|
~~ O2 :+: O2 ...(cong2 (:+:)
|
|
|
|
(lftUnit ma O2)
|
|
|
|
(rgtUnit ma O2))
|
|
|
|
~~ O2 ...(lftUnit mb O2)
|
2021-03-03 17:36:33 +03:00
|
|
|
|
2021-03-03 16:49:32 +03:00
|
|
|
SameMults : (x,y : a) -> (x .+. y) === (x :+: y)
|
|
|
|
SameMults x y = Calc $
|
|
|
|
|~ x .+. y
|
|
|
|
~~ (x :+: O2) .+. (O2 :+: y) ...(sym $ cong2 (.+.)
|
|
|
|
(rgtUnit mb x)
|
|
|
|
(lftUnit mb y))
|
|
|
|
~~ (x .+. O2) :+: (O2 .+. y) ...(prf x O2 O2 y)
|
2021-03-03 17:36:33 +03:00
|
|
|
~~ (x .+. O ) :+: (O .+. y) ...(cong (\u =>
|
2021-03-03 16:49:32 +03:00
|
|
|
(x .+. u) :+: (u .+. y))
|
|
|
|
(sym SameUnits))
|
|
|
|
~~ x :+: y ...(cong2 (:+:)
|
|
|
|
(rgtUnit ma x)
|
|
|
|
(lftUnit ma y))
|
2021-03-03 17:36:33 +03:00
|
|
|
|
2021-03-03 16:49:32 +03:00
|
|
|
Commutativity : Commutative ma
|
|
|
|
Commutativity x y = Calc $
|
|
|
|
|~ x .+. y
|
|
|
|
~~ (O2 :+: x) .+. (y :+: O2) ...(sym $ cong2 (.+.)
|
|
|
|
(lftUnit mb x)
|
|
|
|
(rgtUnit mb y))
|
|
|
|
~~ (O2 .+. y) :+: (x .+. O2) ...(prf O2 x y O2)
|
2021-03-03 17:36:33 +03:00
|
|
|
~~ (O .+. y) :+: (x .+. O ) ...(cong (\u =>
|
2021-03-03 16:49:32 +03:00
|
|
|
(u .+. y) :+: (x .+. u))
|
|
|
|
(sym SameUnits))
|
|
|
|
~~ y :+: x ...(cong2 (:+:)
|
|
|
|
(lftUnit ma y)
|
|
|
|
(rgtUnit ma x))
|
|
|
|
~~ y .+. x ...(sym $ SameMults y x)
|
|
|
|
in (Commutativity, SameMults)
|