comparison agda/deltaM.agda @ 95:cf372fbcebd8

Fix implicit values in deltaM-mu
author Yasutaka Higa <e115763@ie.u-ryukyu.ac.jp>
date Mon, 19 Jan 2015 17:47:55 +0900
parents bcd4fe52a504
children d8cd880f1d78
comparison
equal deleted inserted replaced
94:bcd4fe52a504 95:cf372fbcebd8
80 -> (DeltaM M {functorM} {monadM} A) -> (A -> DeltaM M {functorM} {monadM} B) -> DeltaM M {functorM} {monadM} B 80 -> (DeltaM M {functorM} {monadM} A) -> (A -> DeltaM M {functorM} {monadM} B) -> DeltaM M {functorM} {monadM} B
81 deltaM-bind {l} {A} {B} {M} {functorM} {monadM} (deltaM (mono x)) f = deltaM (mono (bind {l} {A} monadM x (headDeltaM ∙ f))) 81 deltaM-bind {l} {A} {B} {M} {functorM} {monadM} (deltaM (mono x)) f = deltaM (mono (bind {l} {A} monadM x (headDeltaM ∙ f)))
82 deltaM-bind {l} {A} {B} {M} {functorM} {monadM} (deltaM (delta x d)) f = appendDeltaM (deltaM (mono (bind {l} {A} monadM x (headDeltaM ∙ f)))) 82 deltaM-bind {l} {A} {B} {M} {functorM} {monadM} (deltaM (delta x d)) f = appendDeltaM (deltaM (mono (bind {l} {A} monadM x (headDeltaM ∙ f))))
83 (deltaM-bind (deltaM d) (tailDeltaM ∙ f)) 83 (deltaM-bind (deltaM d) (tailDeltaM ∙ f))
84 84
85 deltaM-mu : {l : Level} {A B : Set l} {M : {l' : Level} -> Set l' -> Set l'} 85 deltaM-mu : {l : Level} {A : Set l} {M : {l' : Level} -> Set l' -> Set l'}
86 {functorM : {l' : Level} -> Functor {l'} M} 86 {functorM : {l' : Level} -> Functor {l'} M}
87 {monadM : {l' : Level} {A : Set l'} -> Monad {l'} {A} M functorM} 87 {monadM : {l' : Level} {A : Set l'} -> Monad {l'} {A} M functorM}
88 -> (DeltaM M {functorM} {monadM} (DeltaM M {functorM} {monadM} A)) -> DeltaM M {functorM} {monadM} A 88 -> (DeltaM M {functorM} {monadM} (DeltaM M {functorM} {monadM} A)) -> DeltaM M {functorM} {monadM} A
89 deltaM-mu d = deltaM-bind d id 89 deltaM-mu d = deltaM-bind d id