comparison agda/deltaM.agda @ 109:5bd5f4a7ce8d

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