include "basics/types.ma". include "basics/relations.ma". include "utilities/setoids.ma". include "utilities/proper.ma". record MonadDefinition : Type[1] ≝ { monad : Type[0] → Type[0] ; m_return : ∀X. X → monad X ; m_bind : ∀X,Y. monad X → (X → monad Y) → monad Y }. notation "m »= f" with precedence 49 for @{'m_bind $m $f) }. notation > "!_ e; e'" with precedence 48 for @{'m_bind ${e} (λ_. ${e'})}. notation > "! ident v ← e; e'" with precedence 48 for @{'m_bind ${e} (λ${ident v}. ${e'})}. notation > "! ident v : ty ← e; e'" with precedence 48 for @{'m_bind ${e} (λ${ident v} : ${ty}. ${e'})}. notation < "vbox(! \nbsp ident v ← e ; break e')" with precedence 48 for @{'m_bind ${e} (λ${ident v}.${e'})}. notation > "! ident v ← e : ty ; e'" with precedence 48 for @{'m_bind (${e} : ${ty}) (λ${ident v}.${e'})}. notation < "vbox(! \nbsp ident v : ty \nbsp ← \nbsp e ; break e')" with precedence 48 for @{'m_bind ${e} (λ${ident v} : ${ty}. ${e'})}. notation > "! ident v : ty ← e : ty' ; e'" with precedence 48 for @{'m_bind (${e} : ${ty'}) (λ${ident v} : ${ty}. ${e'})}. notation > "! 〈ident v1, ident v2〉 ← e ; e'" with precedence 48 for @{'m_bind2 ${e} (λ${ident v1}. λ${ident v2}. ${e'})}. notation > "! 〈ident v1, ident v2〉 ← e : ty ; e'" with precedence 48 for @{'m_bind2 (${e} : $ty) (λ${ident v1}. λ${ident v2}. ${e'})}. notation > "! 〈ident v1 : ty1, ident v2 : ty2〉 ← e ; e'" with precedence 48 for @{'m_bind2 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. ${e'})}. notation < "vbox(! \nbsp 〈ident v1, ident v2〉 ← e ; break e')" with precedence 48 for @{'m_bind2 ${e} (λ${ident v1}. λ${ident v2}. ${e'})}. notation < "vbox(! \nbsp 〈ident v1 : ty1, ident v2 : ty2〉 \nbsp ← \nbsp e ; break e')" with precedence 48 for @{'m_bind2 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. ${e'})}. notation > "! 〈ident v1, ident v2, ident v3〉 ← e ; e'" with precedence 48 for @{'m_bind3 ${e} (λ${ident v1}. λ${ident v2}. λ${ident v3}. ${e'})}. notation > "! 〈ident v1, ident v2, ident v3〉 ← e : ty ; e'" with precedence 48 for @{'m_bind3 (${e} : ${ty}) (λ${ident v1}. λ${ident v2}. λ${ident v3}. ${e'})}. notation > "! 〈ident v1 : ty1, ident v2 : ty2, ident v3 : ty3〉 ← e ; e'" with precedence 48 for @{'m_bind3 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. λ${ident v3} : ${ty3}. ${e'})}. notation < "vbox(! \nbsp 〈ident v1, ident v2, ident v3〉 \nbsp ← \nbsp e ; break e')" with precedence 48 for @{'m_bind3 ${e} (λ${ident v1}. λ${ident v2}. λ${ident v3}. ${e'})}. notation < "vbox(! \nbsp 〈ident v1 : ty1, ident v2 : ty2, ident v3 : ty3〉 \nbsp ← e \nbsp ; break e')" with precedence 48 for @{'m_bind3 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. λ${ident v3} : ${ty3}. ${e'})}. (* dependent pair versions *) notation > "! «ident v1, ident v2» ← e ; e'" with precedence 48 for @{'m_bind ${e} (λ${fresh p_sig}. match ${fresh p_sig} with [mk_Sig ${ident v1} ${ident v2} ⇒ ${e'}])}. notation < "vbox(! \nbsp «ident v1, ident v2» \nbsp ← \nbsp e ; break e')" with precedence 48 for @{'m_bind ${e} (λ${fresh p_sig}. match ${fresh p_sig} with [mk_Sig ${ident v1} ${ident v2} ⇒ ${e'}])}. notation > "! «ident v1, ident v2, ident H» ← e ; e'" with precedence 48 for @{'m_sigbind2 ${e} (λ${ident v1},${ident v2},${ident H}. ${e'})}. notation < "vbox(! \nbsp «ident v1, ident v2, ident H» \nbsp ← \nbsp e ; e')" with precedence 48 for @{'m_sigbind2 ${e} (λ${ident v1} : ${ty1}.λ${ident v2} : ${ty2}. λ${ident H} : ${ty3}. ${e'})}. (*alternative do notation for backward compatibility *) notation > "'do'_ e; e'" with precedence 48 for @{'m_bind ${e} (λ_. ${e'})}. notation > "'do' ident v ← e; e'" with precedence 48 for @{'m_bind ${e} (λ${ident v}. ${e'})}. notation > "'do' ident v : ty ← e; e'" with precedence 48 for @{'m_bind ${e} (λ${ident v} : ${ty}. ${e'})}. notation > "'do' ident v ← e : ty ; e'" with precedence 48 for @{'m_bind (${e} : ${ty}) (λ${ident v}.${e'})}. notation > "'do' ident v : ty ← e : ty' ; e'" with precedence 48 for @{'m_bind (${e} : ${ty'}) (λ${ident v} : ${ty}. ${e'})}. notation > "'do' 〈ident v1, ident v2〉 ← e ; e'" with precedence 48 for @{'m_bind2 ${e} (λ${ident v1}. λ${ident v2}. ${e'})}. notation > "'do' 〈ident v1 : ty1, ident v2 : ty2〉 ← e ; e'" with precedence 48 for @{'m_bind2 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. ${e'})}. notation > "'do' 〈ident v1, ident v2, ident v3〉 ← e ; e'" with precedence 48 for @{'m_bind3 ${e} (λ${ident v1}. λ${ident v2}. λ${ident v3}. ${e'})}. notation > "'do' 〈ident v1 : ty1, ident v2 : ty2, ident v3 : ty3〉 ← e ; e'" with precedence 48 for @{'m_bind3 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. λ${ident v3} : ${ty3}. ${e'})}. (* dependent pair versions *) notation > "'do' «ident v1, ident v2» ← e ; e'" with precedence 48 for @{'m_bind ${e} (λ${fresh p_sig}. match ${fresh p_sig} with [mk_Sig ${ident v1} ${ident v2} ⇒ ${e'}])}. notation > "'do' «ident v1, ident v2, ident H» ← e ; e'" with precedence 48 for @{'m_sigbind2 ${e} (λ${ident v1},${ident v2},${ident H}. ${e'})}. notation > "'return' t" with precedence 49 for @{'m_return $t}. notation < "'return' \nbsp t" with precedence 49 for @{'m_return $t}. interpretation "monad bind" 'm_bind m f = (m_bind ? ? ? m f). interpretation "monad return" 'm_return x = (m_return ? ? x). include "basics/lists/list.ma". (* add structure and properties as needed *) record Monad : Type[1] ≝ { m_def :> MonadDefinition ; m_map : ∀X, Y. (X → Y) → monad m_def X → monad m_def Y ; m_map2 : ∀X, Y, Z. (X → Y → Z) → monad m_def X → monad m_def Y → monad m_def Z; m_bind2 : ∀X,Y,Z.monad m_def (X×Y) → (X → Y → monad m_def Z) → monad m_def Z; m_bind3 : ∀X,Y,Z,W.monad m_def (X×Y×Z) → (X → Y → Z → monad m_def W) → monad m_def W; m_join : ∀X.monad m_def (monad m_def X) → monad m_def X; m_sigbind2 : ∀A,B,C.∀P:A×B → Prop. monad m_def (Σx:A×B.P x) → (∀a,b. P 〈a,b〉 → monad m_def C) → monad m_def C; m_mmap : ∀X,Y.(X → monad m_def Y) → list X → monad m_def (list Y); m_mmap_sigma : ∀X,Y,P.(X → monad m_def (Σy : Y.P y)) → list X → monad m_def (Σl : list Y.All ? P l) }. interpretation "monad bind2" 'm_bind2 m f = (m_bind2 ? ? ? ? m f). interpretation "monad bind3" 'm_bind3 m f = (m_bind3 ? ? ? ? ? m f). interpretation "monad sigbind2" 'm_sigbind2 m f = (m_sigbind2 ????? m f). (* check (λM : Monad.λA,B,C,P.λp.λf : ∀a,b.P 〈a,b〉 → monad M C.! «a,b,H» ← p; f a b H) *) record MonadProps : Type[1] ≝ { max_def :> Monad ; m_return_bind : ∀X,Y.∀x : X.∀f : X → monad max_def Y. ! y ← return x ; f y = f x ; m_bind_return : ∀X.∀m : monad max_def X.! x ← m ; return x = m ; m_bind_bind : ∀X,Y,Z. ∀m : monad max_def X.∀f : X → monad max_def Y. ∀g : Y → monad max_def Z.! y ← (! x ← m ; f x) ; g y = ! x ← m; ! y ← f x ; g y }. record SetoidMonadProps : Type[1] ≝ { smax_def :> Monad ; sm_eq : ∀X.relation (monad smax_def X) ; sm_eq_refl : ∀X.reflexive ? (sm_eq X) ; sm_eq_trans : ∀X.transitive ? (sm_eq X) ; sm_eq_sym : ∀X.symmetric ? (sm_eq X) ; sm_return_proper : ∀X .m_return smax_def X ⊨ eq ? ++> sm_eq ? ; sm_bind_proper : ∀X, Y.m_bind smax_def X Y ⊨ sm_eq ? ++> (eq ? ++> sm_eq ?) ++> sm_eq ? ; sm_return_bind : ∀X,Y.∀x : X.∀f : X → monad smax_def Y. sm_eq Y (! y ← return x ; f y) (f x) ; sm_bind_return : ∀X.∀m : monad smax_def X.sm_eq X (! x ← m ; return x) m ; sm_bind_bind : ∀X,Y,Z. ∀m : monad smax_def X.∀f : X → monad smax_def Y. ∀g : Y → monad smax_def Z.sm_eq Z (! y ← (! x ← m ; f x) ; g y) (! x ← m; ! y ← f x ; g y) }. interpretation "monad setoid equality" 'std_eq x y = (sm_eq ?? x y). definition MakeMonad : ?→?→?→Monad ≝ λmonad,m_return,m_bind. mk_Monad (mk_MonadDefinition monad m_return m_bind) (* map *) (λX,Y,f,m. ! x ← m ; return f x) (* map2 *) (λX,Y,Z,f,m,n. ! x ← m ; ! y ← n ; return (f x y)) (* bind2 *) (λX,Y,Z,m,f. ! p ← m ; let 〈x, y〉 ≝ p in f x y) (λX,Y,Z,W,m,f. ! p ← m ; let 〈x, y, z〉 ≝ p in f x y z) (* join *) (λX,m.! x ← m ; x) (* m_sigbind2 *) (λA,B,C,P,e,f. ! e_sig ← e ; match e_sig with [ mk_Sig p p_prf ⇒ match p return λx.x = p → ? with [ mk_Prod a b ⇒ λeq_p. f a b (eq_ind_r ?? (λx.λ_.P x) p_prf ? eq_p) ](refl …) ]) (λX,Y,f,l.foldr … (λel,macc.! r ← f el; !acc ← macc; return (r :: acc)) (return [ ]) l) (λX,Y,P,f,l.foldr … (λel,macc. ! «r, r_prf» ← f el ; ! «acc, acc_prf» ← macc ; return (mk_Sig ?? (r :: acc) ?)) (return (mk_Sig … [ ] ?)) l). % assumption qed. definition MakeMonadProps : ?→?→?→?→?→?→MonadProps ≝ λmonad,m_return,m_bind,p1,p2,p3. mk_MonadProps (MakeMonad monad m_return m_bind) p1 p2 p3. definition MakeSetoidMonadProps : ?→?→?→?→?→?→?→?→?→?→?→?→SetoidMonadProps ≝ λmonad,m_return,m_bind,sm_eq,p1,p2,p3,p4,p5,p6,p7,p8. mk_SetoidMonadProps (MakeMonad monad m_return m_bind) sm_eq p1 p2 p3 p4 p5 p6 p7 p8.