comparison agda/deltaM.agda @ 112:0a3b6cb91a05

Prove left-unity-law for DeltaM
author Yasutaka Higa <e115763@ie.u-ryukyu.ac.jp>
date Fri, 30 Jan 2015 21:57:31 +0900
parents 9fe3d0bd1149
children e6bcc7467335
comparison
equal deleted inserted replaced
111:9fe3d0bd1149 112:0a3b6cb91a05
9 module deltaM where 9 module deltaM where
10 10
11 -- DeltaM definitions 11 -- DeltaM definitions
12 12
13 data DeltaM {l : Level} 13 data DeltaM {l : Level}
14 (M : {l' : Level} -> Set l' -> Set l') 14 (M : Set l -> Set l)
15 {functorM : {l' : Level} -> Functor {l'} M} 15 {functorM : Functor M}
16 {monadM : {l' : Level} {A : Set l'} -> Monad {l'} M functorM} 16 {monadM : Monad M functorM}
17 (A : Set l) 17 (A : Set l)
18 : (Nat -> Set l) where 18 : (Nat -> Set l) where
19 deltaM : {n : Nat} -> Delta (M A) (S n) -> DeltaM M {functorM} {monadM} A (S n) 19 deltaM : {n : Nat} -> Delta (M A) (S n) -> DeltaM M {functorM} {monadM} A (S n)
20 20
21 21
22 -- DeltaM utils 22 -- DeltaM utils
23 23
24 headDeltaM : {l : Level} {A : Set l} {n : Nat} 24 headDeltaM : {l : Level} {A : Set l} {n : Nat}
25 {M : {l' : Level} -> Set l' -> Set l'} 25 {M : Set l -> Set l} {functorM : Functor M} {monadM : Monad M functorM}
26 {functorM : {l' : Level} -> Functor {l'} M}
27 {monadM : {l' : Level} -> Monad {l'} M functorM}
28 -> DeltaM M {functorM} {monadM} A (S n) -> M A 26 -> DeltaM M {functorM} {monadM} A (S n) -> M A
29 headDeltaM (deltaM d) = headDelta d 27 headDeltaM (deltaM d) = headDelta d
30 28
31 29
32 tailDeltaM : {l : Level} {A : Set l} {n : Nat} 30 tailDeltaM : {l : Level} {A : Set l} {n : Nat}
33 {M : {l' : Level} -> Set l' -> Set l'} 31 {M : Set l -> Set l} {functorM : Functor M} {monadM : Monad M functorM}
34 {functorM : {l' : Level} -> Functor {l'} M}
35 {monadM : {l' : Level} -> Monad {l'} M functorM}
36 -> DeltaM {l} M {functorM} {monadM} A (S (S n)) -> DeltaM M {functorM} {monadM} A (S n) 32 -> DeltaM {l} M {functorM} {monadM} A (S (S n)) -> DeltaM M {functorM} {monadM} A (S n)
37 tailDeltaM {_} {n} (deltaM d) = deltaM (tailDelta d) 33 tailDeltaM {_} {n} (deltaM d) = deltaM (tailDelta d)
38 34
39 35
40 appendDeltaM : {l : Level} {A : Set l} {n m : Nat} 36 appendDeltaM : {l : Level} {A : Set l} {n m : Nat}
41 {M : {l' : Level} -> Set l' -> Set l'} 37 {M : Set l -> Set l} {functorM : Functor M} {monadM : Monad M functorM} ->
42 {functorM : {l' : Level} -> Functor {l'} M}
43 {monadM : {l' : Level} -> Monad {l'} M functorM} ->
44 DeltaM M {functorM} {monadM} A (S n) -> 38 DeltaM M {functorM} {monadM} A (S n) ->
45 DeltaM M {functorM} {monadM} A (S m) -> 39 DeltaM M {functorM} {monadM} A (S m) ->
46 DeltaM M {functorM} {monadM} A ((S n) + (S m)) 40 DeltaM M {functorM} {monadM} A ((S n) + (S m))
47 appendDeltaM (deltaM d) (deltaM dd) = deltaM (deltaAppend d dd) 41 appendDeltaM (deltaM d) (deltaM dd) = deltaM (deltaAppend d dd)
48 42
49 43
50 44
51 45
52 -- functor definitions 46 -- functor definitions
53 open Functor 47 open Functor
54 deltaM-fmap : {l : Level} {A B : Set l} {n : Nat} 48 deltaM-fmap : {l : Level} {A B : Set l} {n : Nat}
55 {M : {l' : Level} -> Set l' -> Set l'} 49 {M : Set l -> Set l} {functorM : Functor M} {monadM : Monad M functorM}
56 {functorM : {l' : Level} -> Functor {l'} M}
57 {monadM : {l' : Level} -> Monad {l'} M functorM}
58 -> (A -> B) -> DeltaM M {functorM} {monadM} A (S n) -> DeltaM M {functorM} {monadM} B (S n) 50 -> (A -> B) -> DeltaM M {functorM} {monadM} A (S n) -> DeltaM M {functorM} {monadM} B (S n)
59 deltaM-fmap {l} {A} {B} {n} {M} {functorM} f (deltaM d) = deltaM (fmap delta-is-functor (fmap functorM f) d) 51 deltaM-fmap {l} {A} {B} {n} {M} {functorM} f (deltaM d) = deltaM (fmap delta-is-functor (fmap functorM f) d)
60 52
61 53
62 54
63 55
64 -- monad definitions 56 -- monad definitions
65 open Monad 57 open Monad
66 58
67 deltaM-eta : {l : Level} {A : Set l} {n : Nat} 59 deltaM-eta : {l : Level} {A : Set l} {n : Nat}
68 {M : {l' : Level} -> Set l' -> Set l'} 60 {M : Set l -> Set l} {functorM : Functor M} {monadM : Monad M functorM} ->
69 {functorM : {l' : Level} -> Functor {l'} M}
70 {monadM : {l' : Level} -> Monad {l'} M functorM} ->
71 A -> (DeltaM M {functorM} {monadM} A (S n)) 61 A -> (DeltaM M {functorM} {monadM} A (S n))
72 deltaM-eta {n = n} {monadM = mm} x = deltaM (delta-eta {n = n} (eta mm x)) 62 deltaM-eta {n = n} {monadM = mm} x = deltaM (delta-eta {n = n} (eta mm x))
73 63
74 deltaM-mu : {l : Level} {A : Set l} {n : Nat} 64 deltaM-mu : {l : Level} {A : Set l} {n : Nat}
75 {M : {l' : Level} -> Set l' -> Set l'} 65 {M : Set l -> Set l} {functorM : Functor M} {monadM : Monad M functorM} ->
76 {functorM : {l' : Level} -> Functor {l'} M}
77 {monadM : {l' : Level} -> Monad {l'} M functorM} ->
78 DeltaM M {functorM} {monadM} (DeltaM M {functorM} {monadM} A (S n)) (S n) -> 66 DeltaM M {functorM} {monadM} (DeltaM M {functorM} {monadM} A (S n)) (S n) ->
79 DeltaM M {functorM} {monadM} A (S n) 67 DeltaM M {functorM} {monadM} A (S n)
80 deltaM-mu {n = O} {functorM = fm} {monadM = mm} (deltaM (mono x)) = deltaM (mono (mu mm (fmap fm headDeltaM x))) 68 deltaM-mu {n = O} {functorM = fm} {monadM = mm} (deltaM (mono x)) = deltaM (mono (mu mm (fmap fm headDeltaM x)))
81 deltaM-mu {n = S n} {functorM = fm} {monadM = mm} (deltaM (delta x d)) = appendDeltaM (deltaM (mono (mu mm (fmap fm headDeltaM x)))) 69 deltaM-mu {n = S n} {functorM = fm} {monadM = mm} (deltaM (delta x d)) = appendDeltaM (deltaM (mono (mu mm (fmap fm headDeltaM x))))
82 (deltaM-mu (deltaM-fmap tailDeltaM (deltaM d))) 70 (deltaM-mu (deltaM-fmap tailDeltaM (deltaM d)))
83 71
84 72
85 deltaM-bind : {l : Level} {A B : Set l} 73 deltaM-bind : {l : Level} {A B : Set l}
86 {n : Nat} 74 {n : Nat}
87 {M : {l' : Level} -> Set l' -> Set l'} 75 {M : Set l -> Set l}
88 {functorM : {l' : Level} -> Functor {l'} M} 76 {functorM : Functor M}
89 {monadM : {l' : Level} -> Monad {l'} M functorM} -> 77 {monadM : Monad M functorM} ->
90 (DeltaM M {functorM} {monadM} A (S n)) -> 78 (DeltaM M {functorM} {monadM} A (S n)) ->
91 (A -> DeltaM M {functorM} {monadM} B (S n)) 79 (A -> DeltaM M {functorM} {monadM} B (S n))
92 -> DeltaM M {functorM} {monadM} B (S n) 80 -> DeltaM M {functorM} {monadM} B (S n)
93 deltaM-bind {n = O} {monadM = mm} (deltaM (mono x)) f = deltaM (mono (bind mm x (headDeltaM ∙ f))) 81 deltaM-bind {n = O} {monadM = mm} (deltaM (mono x)) f = deltaM (mono (bind mm x (headDeltaM ∙ f)))
94 deltaM-bind {n = S n} {monadM = mm} (deltaM (delta x d)) f = appendDeltaM (deltaM (mono (bind mm x (headDeltaM ∙ f)))) 82 deltaM-bind {n = S n} {monadM = mm} (deltaM (delta x d)) f = appendDeltaM (deltaM (mono (bind mm x (headDeltaM ∙ f))))