include "basics/types.ma". include "basics/relations.ma". include "utilities/setoids.ma". definition pred_transformer ≝ λA,B : Type[0].(A → Prop) → B → Prop. definition modus_ponens ≝ λA,B.λPT : pred_transformer A B. ∀P,Q.(∀x.P x → Q x) → ∀y.PT P y → PT Q y. lemma mp_transitive : ∀A,B,C,PT1,PT2.modus_ponens A B PT1 → modus_ponens B C PT2 → modus_ponens A C (PT2 ∘ PT1). /4/ qed. definition rel_transformer ≝ λA,B,C,D : Type[0]. (A → B → Prop) → C → D → Prop. definition rel_modus_ponens ≝ λA,B,C,D.λRT : rel_transformer A B C D. ∀P,Q.(∀x,y.P x y → Q x y) → ∀z,w.RT P z w → RT Q z w. lemma rel_mp_transitive : ∀A,B,C,D,E,F,RT1,RT2.rel_modus_ponens A B C D RT1 → rel_modus_ponens C D E F RT2 → rel_modus_ponens … (RT2 ∘ RT1). /4/ qed. record Monad : Type[1] ≝ { monad :1> 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). record MonadProps : Type[1] ≝ { max_def :> Monad ; m_return_bind : ∀X,Y.∀x : X.∀f : X → max_def Y. ! y ← return x ; f y = f x ; m_bind_return : ∀X.∀m : max_def X.! x ← m ; return x = m ; m_bind_bind : ∀X,Y,Z. ∀m : max_def X.∀f : X → max_def Y. ∀g : Y → max_def Z.! y ← (! x ← m ; f x) ; g y = ! x ← m; ! y ← f x ; g y ; m_bind_ext_eq : ∀X,Y.∀m : max_def X.∀f,g : X → max_def Y. (∀x.f x = g x) → ! x ← m ; f x = ! x ← m ; g x }. record SetoidMonadProps : Type[1] ≝ { smax_def :> Monad ; sm_eq : ∀X.relation (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,x.sm_eq X (return x) (return x) ; sm_bind_proper : ∀X,Y,x,y,f,g.sm_eq X x y → (∀x.sm_eq Y (f x) (g x)) → sm_eq Y (x »= f) (y »= g) ; sm_return_bind : ∀X,Y.∀x : X.∀f : X → smax_def Y. sm_eq Y (! y ← return x ; f y) (f x) ; sm_bind_return : ∀X.∀m : smax_def X.sm_eq X (! x ← m ; return x) m ; sm_bind_bind : ∀X,Y,Z. ∀m : smax_def X.∀f : X → smax_def Y. ∀g : Y → smax_def Z.sm_eq Z (! y ← (! x ← m ; f x) ; g y) (! x ← m; ! y ← f x ; g y) }. definition setoid_of_monad : ∀M : SetoidMonadProps.∀X : Type[0]. Setoid ≝ λM,X.mk_Setoid (M X) (sm_eq M X) (sm_eq_refl M X) (sm_eq_trans M X) (sm_eq_sym M X). include "hints_declaration.ma". alias symbol "hint_decl" (instance 1) = "hint_decl_Type1". unification hint 0 ≔ M, X; M' ≟ smax_def M, S ≟ setoid_of_monad M X (*-----------------------------*)⊢ monad M' X ≡ std_supp S. include "basics/lists/list.ma". definition m_map ≝ λM : Monad.λX,Y.λf : X → Y.λm : M X. ! x ← m ; return f x. definition m_map2 ≝ λM : Monad.λX,Y,Z.λf : X → Y → Z.λm : M X.λn : M Y. ! x ← m ; ! y ← n ; return f x y. definition m_bind2 ≝ λM : Monad.λX,Y,Z.λm : M (X × Y).λf : X → Y → M Z. ! p ← m ; f (\fst p) (\snd p). interpretation "monad bind2" 'm_bind2 m f = (m_bind2 ? ? ? ? m f). definition m_bind3 : ∀M : Monad.∀X,Y,Z,W.M (X×Y×Z) → (X → Y → Z → M W) → M W ≝ λM,X,Y,Z,W,m,f. ! p ← m ; f (\fst (\fst p)) (\snd (\fst p)) (\snd p). interpretation "monad bind3" 'm_bind3 m f = (m_bind3 ? ? ? ? ? m f). definition m_join : ∀M : Monad.∀X.M (M X) → M X ≝ λM,X,m.! x ← m ; x. definition m_sigbind2 : ∀M : Monad.∀A,B,C.∀P:A×B → Prop. M (Σx:A×B.P x) → (∀a,b. P 〈a,b〉 → M C) → M C ≝ λM,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 …) ]. interpretation "monad sigbind2" 'm_sigbind2 m f = (m_sigbind2 ????? m f). definition m_list_map : ∀M : Monad.∀X,Y.(X → M Y) → list X → M (list Y) ≝ λM,X,Y,f,l.foldr … (λel,macc.! r ← f el; !acc ← macc; return (r :: acc)) (return [ ]) l. definition m_list_map_sigma : ∀M : Monad.∀X,Y,P.(X → M (Σy : Y.P y)) → list X → M (Σl : list Y.All ? P l) ≝ λM,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 m_bin_op : ∀M : Monad.∀X,Y,Z.(X → Y → Z) → M X → M Y → M Z ≝ λM,X,Y,Z,op,m,n. ! x ← m ; ! y ← n ; return op x y. unification hint 0 ≔ M, X, Y, Z, m, f ; P ≟ Prod X Y, F ≟ λp.f (\fst p) (\snd p) ⊢ m_bind M P Z m F ≡ m_bind2 M X Y Z m f. unification hint 0 ≔ M, X, Y, Z, W, m, f ; P' ≟ Prod X Y, P ≟ Prod P' Z, F ≟ λp.f (\fst (\fst p)) (\snd (\fst p)) (\snd p) ⊢ m_bind M P W m F ≡ m_bind3 M X Y Z W m f. definition MakeMonadProps : ?→?→?→?→?→?→?→MonadProps ≝ λmonad,m_return,m_bind. mk_MonadProps (mk_Monad monad m_return m_bind). definition MakeSetoidMonadProps : ?→?→?→?→?→?→?→?→?→?→?→?→SetoidMonadProps ≝ λmonad,m_return,m_bind. mk_SetoidMonadProps (mk_Monad monad m_return m_bind). include "basics/russell.ma". record MonadPred (M : Monad) : Type[1] ≝ { m_pred :2> ∀X.(X → Prop) → (M X → Prop) ; mp_return : ∀X,P,x.P x → m_pred X P (return x) ; mp_bind : ∀X,Y,Pin,Pout,m.m_pred X Pin m → ∀f.(∀x.Pin x → m_pred Y Pout (f x)) → m_pred ? Pout (m_bind … m f) ; m_pred_mp : ∀X.modus_ponens ?? (m_pred X) }. record InjMonadPred (M : Monad) : Type[1] ≝ { im_pred :> MonadPred M ; mp_inject : ∀X.∀P : X → Prop.(Σm.im_pred P m) → M (Σx.P x) ; mp_inject_eq : ∀X,P,m.pi1 ?? m = ! x ← mp_inject X P m ; return pi1 … x }. coercion coerc_mp_inject : ∀M.∀MP:InjMonadPred M. ∀X.∀P : X → Prop.∀m : Σm.MP P m.M (Σx.P x) ≝ mp_inject on _m : Sig (monad ??) (λm.im_pred ??? m) to monad ? (Sig ? (λx.? x)). lemma mp_inject_bind : ∀M : MonadProps.∀MP,X,P,Y.∀m.∀f : X → M Y. ! x ← mp_inject M MP X P m ; f (pi1 … x) = ! x ← pi1 … m ; f x. #M#MP#X#P#Y#m#f >mp_inject_eq >m_bind_bind @m_bind_ext_eq #x >m_return_bind % qed. record MonadRel (M1 : Monad) (M2 : Monad) : Type[1] ≝ { m_rel :3> ∀X,Y.(X → Y → Prop) → (M1 X → M2 Y → Prop) ; mr_return : ∀X,Y,rel,x,y.rel x y → m_rel X Y rel (return x) (return y) ; mr_bind : ∀X,Y,Z,W,relin,relout,m,n.m_rel X Y relin m n → ∀f,g.(∀x,y.relin x y → m_rel Z W relout (f x) (g y)) → m_rel ?? relout (m_bind … m f) (m_bind … n g) ; m_rel_mp : ∀X,Y.rel_modus_ponens ???? (m_rel X Y) }.