1 | include "basics/types.ma". |
---|
2 | include "basics/relations.ma". |
---|
3 | include "utilities/setoids.ma". |
---|
4 | |
---|
5 | definition pred_transformer ≝ λA,B : Type[0].(A → Prop) → B → Prop. |
---|
6 | |
---|
7 | definition modus_ponens ≝ λA,B.λPT : pred_transformer A B. |
---|
8 | ∀P,Q.(∀x.P x → Q x) → ∀y.PT P y → PT Q y. |
---|
9 | |
---|
10 | lemma mp_transitive : |
---|
11 | ∀A,B,C,PT1,PT2.modus_ponens A B PT1 → modus_ponens B C PT2 → |
---|
12 | modus_ponens A C (PT2 ∘ PT1). /4/ qed. |
---|
13 | |
---|
14 | definition rel_transformer ≝ λA,B,C,D : Type[0]. |
---|
15 | (A → B → Prop) → C → D → Prop. |
---|
16 | |
---|
17 | definition rel_modus_ponens ≝ λA,B,C,D.λRT : rel_transformer A B C D. |
---|
18 | ∀P,Q.(∀x,y.P x y → Q x y) → ∀z,w.RT P z w → RT Q z w. |
---|
19 | |
---|
20 | lemma rel_mp_transitive : |
---|
21 | ∀A,B,C,D,E,F,RT1,RT2.rel_modus_ponens A B C D RT1 → rel_modus_ponens C D E F RT2 → |
---|
22 | rel_modus_ponens … (RT2 ∘ RT1). /4/ qed. |
---|
23 | |
---|
24 | record Monad : Type[1] ≝ { |
---|
25 | monad :1> Type[0] → Type[0] ; |
---|
26 | m_return : ∀X. X → monad X ; |
---|
27 | m_bind : ∀X,Y. monad X → (X → monad Y) → monad Y |
---|
28 | }. |
---|
29 | |
---|
30 | notation "m »= f" with precedence 49 |
---|
31 | for @{'m_bind $m $f }. |
---|
32 | |
---|
33 | notation > "!_ e; e'" |
---|
34 | with precedence 48 for @{'m_bind ${e} (λ_. ${e'})}. |
---|
35 | notation > "! ident v ← e; e'" |
---|
36 | with precedence 48 for @{'m_bind ${e} (λ${ident v}. ${e'})}. |
---|
37 | notation > "! ident v : ty ← e; e'" |
---|
38 | with precedence 48 for @{'m_bind ${e} (λ${ident v} : ${ty}. ${e'})}. |
---|
39 | notation < "vbox(! \nbsp ident v ← e ; break e')" |
---|
40 | with precedence 48 for @{'m_bind ${e} (λ${ident v}.${e'})}. |
---|
41 | notation > "! ident v ← e : ty ; e'" |
---|
42 | with precedence 48 for @{'m_bind (${e} : ${ty}) (λ${ident v}.${e'})}. |
---|
43 | notation < "vbox(! \nbsp ident v : ty \nbsp ← \nbsp e ; break e')" |
---|
44 | with precedence 48 for @{'m_bind ${e} (λ${ident v} : ${ty}. ${e'})}. |
---|
45 | notation > "! ident v : ty ← e : ty' ; e'" |
---|
46 | with precedence 48 for @{'m_bind (${e} : ${ty'}) (λ${ident v} : ${ty}. ${e'})}. |
---|
47 | notation > "! 〈ident v1, ident v2〉 ← e ; e'" |
---|
48 | with precedence 48 for @{'m_bind2 ${e} (λ${ident v1}. λ${ident v2}. ${e'})}. |
---|
49 | notation > "! 〈ident v1, ident v2〉 ← e : ty ; e'" |
---|
50 | with precedence 48 for @{'m_bind2 (${e} : $ty) (λ${ident v1}. λ${ident v2}. ${e'})}. |
---|
51 | notation > "! 〈ident v1 : ty1, ident v2 : ty2〉 ← e ; e'" |
---|
52 | with precedence 48 for @{'m_bind2 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. ${e'})}. |
---|
53 | notation < "vbox(! \nbsp 〈ident v1, ident v2〉 ← e ; break e')" |
---|
54 | with precedence 48 for @{'m_bind2 ${e} (λ${ident v1}. λ${ident v2}. ${e'})}. |
---|
55 | notation < "vbox(! \nbsp 〈ident v1 : ty1, ident v2 : ty2〉 \nbsp ← \nbsp e ; break e')" |
---|
56 | with precedence 48 for @{'m_bind2 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. ${e'})}. |
---|
57 | notation > "! 〈ident v1, ident v2, ident v3〉 ← e ; e'" |
---|
58 | with precedence 48 for @{'m_bind3 ${e} (λ${ident v1}. λ${ident v2}. λ${ident v3}. ${e'})}. |
---|
59 | notation > "! 〈ident v1, ident v2, ident v3〉 ← e : ty ; e'" |
---|
60 | with precedence 48 for @{'m_bind3 (${e} : ${ty}) (λ${ident v1}. λ${ident v2}. λ${ident v3}. ${e'})}. |
---|
61 | notation > "! 〈ident v1 : ty1, ident v2 : ty2, ident v3 : ty3〉 ← e ; e'" |
---|
62 | with precedence 48 for @{'m_bind3 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. λ${ident v3} : ${ty3}. ${e'})}. |
---|
63 | notation < "vbox(! \nbsp 〈ident v1, ident v2, ident v3〉 \nbsp ← \nbsp e ; break e')" |
---|
64 | with precedence 48 for @{'m_bind3 ${e} (λ${ident v1}. λ${ident v2}. λ${ident v3}. ${e'})}. |
---|
65 | notation < "vbox(! \nbsp 〈ident v1 : ty1, ident v2 : ty2, ident v3 : ty3〉 \nbsp ← e \nbsp ; break e')" |
---|
66 | with precedence 48 for @{'m_bind3 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. λ${ident v3} : ${ty3}. ${e'})}. |
---|
67 | |
---|
68 | (* dependent pair versions *) |
---|
69 | notation > "! «ident v1, ident v2» ← e ; e'" |
---|
70 | with precedence 48 for |
---|
71 | @{'m_bind ${e} (λ${fresh p_sig}. |
---|
72 | match ${fresh p_sig} with [mk_Sig ${ident v1} ${ident v2} ⇒ ${e'}])}. |
---|
73 | notation < "vbox(! \nbsp «ident v1, ident v2» \nbsp ← \nbsp e ; break e')" |
---|
74 | with precedence 48 for |
---|
75 | @{'m_bind ${e} (λ${fresh p_sig}. |
---|
76 | match ${fresh p_sig} with [mk_Sig ${ident v1} ${ident v2} ⇒ ${e'}])}. |
---|
77 | |
---|
78 | notation > "! «ident v1, ident v2, ident H» ← e ; e'" |
---|
79 | with precedence 48 for |
---|
80 | @{'m_sigbind2 ${e} (λ${ident v1},${ident v2},${ident H}. ${e'})}. |
---|
81 | notation < "vbox(! \nbsp «ident v1, ident v2, ident H» \nbsp ← \nbsp e ; e')" |
---|
82 | with precedence 48 for |
---|
83 | @{'m_sigbind2 ${e} (λ${ident v1} : ${ty1}.λ${ident v2} : ${ty2}. |
---|
84 | λ${ident H} : ${ty3}. ${e'})}. |
---|
85 | |
---|
86 | |
---|
87 | (*alternative do notation for backward compatibility *) |
---|
88 | notation > "'do'_ e; e'" |
---|
89 | with precedence 48 for @{'m_bind ${e} (λ_. ${e'})}. |
---|
90 | notation > "'do' ident v ← e; e'" |
---|
91 | with precedence 48 for @{'m_bind ${e} (λ${ident v}. ${e'})}. |
---|
92 | notation > "'do' ident v : ty ← e; e'" |
---|
93 | with precedence 48 for @{'m_bind ${e} (λ${ident v} : ${ty}. ${e'})}. |
---|
94 | notation > "'do' ident v ← e : ty ; e'" |
---|
95 | with precedence 48 for @{'m_bind (${e} : ${ty}) (λ${ident v}.${e'})}. |
---|
96 | notation > "'do' ident v : ty ← e : ty' ; e'" |
---|
97 | with precedence 48 for @{'m_bind (${e} : ${ty'}) (λ${ident v} : ${ty}. ${e'})}. |
---|
98 | notation > "'do' 〈ident v1, ident v2〉 ← e ; e'" |
---|
99 | with precedence 48 for @{'m_bind2 ${e} (λ${ident v1}. λ${ident v2}. ${e'})}. |
---|
100 | notation > "'do' 〈ident v1 : ty1, ident v2 : ty2〉 ← e ; e'" |
---|
101 | with precedence 48 for @{'m_bind2 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. ${e'})}. |
---|
102 | notation > "'do' 〈ident v1, ident v2, ident v3〉 ← e ; e'" |
---|
103 | with precedence 48 for @{'m_bind3 ${e} (λ${ident v1}. λ${ident v2}. λ${ident v3}. ${e'})}. |
---|
104 | notation > "'do' 〈ident v1 : ty1, ident v2 : ty2, ident v3 : ty3〉 ← e ; e'" |
---|
105 | with precedence 48 for @{'m_bind3 ${e} (λ${ident v1} : ${ty1}. λ${ident v2} : ${ty2}. λ${ident v3} : ${ty3}. ${e'})}. |
---|
106 | |
---|
107 | (* dependent pair versions *) |
---|
108 | notation > "'do' «ident v1, ident v2» ← e ; e'" |
---|
109 | with precedence 48 for |
---|
110 | @{'m_bind ${e} (λ${fresh p_sig}. |
---|
111 | match ${fresh p_sig} with [mk_Sig ${ident v1} ${ident v2} ⇒ ${e'}])}. |
---|
112 | |
---|
113 | notation > "'do' «ident v1, ident v2, ident H» ← e ; e'" |
---|
114 | with precedence 48 for |
---|
115 | @{'m_sigbind2 ${e} (λ${ident v1},${ident v2},${ident H}. ${e'})}. |
---|
116 | |
---|
117 | notation > "'do' ❬ident v1, ident v2❭ ← e ; e'" |
---|
118 | with precedence 48 for |
---|
119 | @{'m_bind ${e} (λ${fresh p_sig}. |
---|
120 | match ${fresh p_sig} with [mk_DPair ${ident v1} ${ident v2} ⇒ ${e'}])}. |
---|
121 | |
---|
122 | notation > "'return' t" with precedence 49 for @{'m_return $t}. |
---|
123 | notation < "'return' \nbsp t" with precedence 49 for @{'m_return $t}. |
---|
124 | |
---|
125 | interpretation "monad bind" 'm_bind m f = (m_bind ? ? ? m f). |
---|
126 | interpretation "monad return" 'm_return x = (m_return ? ? x). |
---|
127 | |
---|
128 | |
---|
129 | record MonadProps : Type[1] ≝ |
---|
130 | { max_def :> Monad |
---|
131 | ; m_return_bind : ∀X,Y.∀x : X.∀f : X → max_def Y. ! y ← return x ; f y = f x |
---|
132 | ; m_bind_return : ∀X.∀m : max_def X.! x ← m ; return x = m |
---|
133 | ; m_bind_bind : ∀X,Y,Z. ∀m : max_def X.∀f : X → max_def Y. |
---|
134 | ∀g : Y → max_def Z.! y ← (! x ← m ; f x) ; g y = ! x ← m; ! y ← f x ; g y |
---|
135 | ; m_bind_ext_eq : ∀X,Y.∀m : max_def X.∀f,g : X → max_def Y. |
---|
136 | (∀x.f x = g x) → ! x ← m ; f x = ! x ← m ; g x |
---|
137 | }. |
---|
138 | |
---|
139 | record SetoidMonadProps : Type[1] ≝ |
---|
140 | { smax_def :> Monad |
---|
141 | ; sm_eq : ∀X.relation (smax_def X) |
---|
142 | ; sm_eq_refl : ∀X.reflexive ? (sm_eq X) |
---|
143 | ; sm_eq_trans : ∀X.transitive ? (sm_eq X) |
---|
144 | ; sm_eq_sym : ∀X.symmetric ? (sm_eq X) |
---|
145 | ; sm_return_proper : ∀X,x.sm_eq X (return x) (return x) |
---|
146 | ; 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) |
---|
147 | ; sm_return_bind : ∀X,Y.∀x : X.∀f : X → smax_def Y. |
---|
148 | sm_eq Y (! y ← return x ; f y) (f x) |
---|
149 | ; sm_bind_return : ∀X.∀m : smax_def X.sm_eq X (! x ← m ; return x) m |
---|
150 | ; sm_bind_bind : ∀X,Y,Z. ∀m : smax_def X.∀f : X → smax_def Y. |
---|
151 | ∀g : Y → smax_def Z.sm_eq Z (! y ← (! x ← m ; f x) ; g y) (! x ← m; ! y ← f x ; g y) |
---|
152 | }. |
---|
153 | |
---|
154 | definition setoid_of_monad : ∀M : SetoidMonadProps.∀X : Type[0]. |
---|
155 | Setoid ≝ |
---|
156 | λ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). |
---|
157 | |
---|
158 | include "hints_declaration.ma". |
---|
159 | alias symbol "hint_decl" (instance 1) = "hint_decl_Type1". |
---|
160 | unification hint 0 ≔ M, X; |
---|
161 | M' ≟ smax_def M, S ≟ setoid_of_monad M X |
---|
162 | (*-----------------------------*)⊢ |
---|
163 | monad M' X ≡ std_supp S. |
---|
164 | |
---|
165 | include "basics/lists/list.ma". |
---|
166 | |
---|
167 | definition m_map ≝ λM : Monad.λX,Y.λf : X → Y.λm : M X. |
---|
168 | ! x ← m ; return f x. |
---|
169 | |
---|
170 | definition m_map2 ≝ λM : Monad.λX,Y,Z.λf : X → Y → Z.λm : M X.λn : M Y. |
---|
171 | ! x ← m ; ! y ← n ; return f x y. |
---|
172 | |
---|
173 | definition m_bind2 ≝ λM : Monad.λX,Y,Z.λm : M (X × Y).λf : X → Y → M Z. |
---|
174 | ! p ← m ; f (\fst p) (\snd p). |
---|
175 | |
---|
176 | interpretation "monad bind2" 'm_bind2 m f = (m_bind2 ? ? ? ? m f). |
---|
177 | |
---|
178 | definition m_bind3 : |
---|
179 | ∀M : Monad.∀X,Y,Z,W.M (X×Y×Z) → (X → Y → Z → M W) → M W ≝ |
---|
180 | λM,X,Y,Z,W,m,f. |
---|
181 | ! p ← m ; f (\fst (\fst p)) (\snd (\fst p)) (\snd p). |
---|
182 | |
---|
183 | interpretation "monad bind3" 'm_bind3 m f = (m_bind3 ? ? ? ? ? m f). |
---|
184 | |
---|
185 | definition m_join : ∀M : Monad.∀X.M (M X) → M X ≝ |
---|
186 | λM,X,m.! x ← m ; x. |
---|
187 | |
---|
188 | definition m_sigbind2 : |
---|
189 | ∀M : Monad.∀A,B,C.∀P:A×B → Prop. M (Σx:A×B.P x) → |
---|
190 | (∀a,b. P 〈a,b〉 → M C) → M C ≝ |
---|
191 | λM,A,B,C,P,e,f. |
---|
192 | ! e_sig ← e ; |
---|
193 | match e_sig with |
---|
194 | [ mk_Sig p p_prf ⇒ |
---|
195 | match p return λx.x = p → ? with |
---|
196 | [ mk_Prod a b ⇒ |
---|
197 | λeq_p. f a b (eq_ind_r ?? (λx.λ_.P x) p_prf ? eq_p) |
---|
198 | ](refl …) |
---|
199 | ]. |
---|
200 | |
---|
201 | interpretation "monad sigbind2" 'm_sigbind2 m f = (m_sigbind2 ????? m f). |
---|
202 | |
---|
203 | definition m_list_map : |
---|
204 | ∀M : Monad.∀X,Y.(X → M Y) → list X → M (list Y) ≝ |
---|
205 | λM,X,Y,f,l.foldr … (λel,macc.! r ← f el; !acc ← macc; return (r :: acc)) (return [ ]) l. |
---|
206 | |
---|
207 | definition m_list_map_sigma : |
---|
208 | ∀M : Monad.∀X,Y,P.(X → M (Σy : Y.P y)) → list X → M (Σl : list Y.All ? P l) ≝ |
---|
209 | λM,X,Y,P,f,l.foldr … (λel,macc. |
---|
210 | ! «r, r_prf» ← f el ; |
---|
211 | ! «acc, acc_prf» ← macc ; |
---|
212 | return (mk_Sig ?? (r :: acc) ?)) |
---|
213 | (return (mk_Sig … [ ] ?)) l. % assumption qed. |
---|
214 | |
---|
215 | definition m_bin_op : |
---|
216 | ∀M : Monad.∀X,Y,Z.(X → Y → Z) → M X → M Y → M Z ≝ |
---|
217 | λM,X,Y,Z,op,m,n. ! x ← m ; ! y ← n ; return op x y. |
---|
218 | |
---|
219 | unification hint 0 ≔ M, X, Y, Z, m, f ; |
---|
220 | P ≟ Prod X Y, F ≟ λp.f (\fst p) (\snd p) ⊢ |
---|
221 | m_bind M P Z m F ≡ m_bind2 M X Y Z m f. |
---|
222 | |
---|
223 | unification hint 0 ≔ M, X, Y, Z, W, m, f ; |
---|
224 | P' ≟ Prod X Y, P ≟ Prod P' Z, F ≟ λp.f (\fst (\fst p)) (\snd (\fst p)) (\snd p) ⊢ |
---|
225 | m_bind M P W m F ≡ m_bind3 M X Y Z W m f. |
---|
226 | |
---|
227 | definition MakeMonadProps : ?→?→?→?→?→?→?→MonadProps ≝ λmonad,m_return,m_bind. |
---|
228 | mk_MonadProps (mk_Monad monad m_return m_bind). |
---|
229 | |
---|
230 | definition MakeSetoidMonadProps : ?→?→?→?→?→?→?→?→?→?→?→?→SetoidMonadProps ≝ |
---|
231 | λmonad,m_return,m_bind. |
---|
232 | mk_SetoidMonadProps (mk_Monad monad m_return m_bind). |
---|
233 | |
---|
234 | include "basics/russell.ma". |
---|
235 | |
---|
236 | record MonadPred (M : Monad) : Type[1] ≝ |
---|
237 | { m_pred :2> ∀X.(X → Prop) → (M X → Prop) |
---|
238 | ; mp_return : ∀X,P,x.P x → m_pred X P (return x) |
---|
239 | ; mp_bind : ∀X,Y,Pin,Pout,m.m_pred X Pin m → |
---|
240 | ∀f.(∀x.Pin x → m_pred Y Pout (f x)) → |
---|
241 | m_pred ? Pout (m_bind … m f) |
---|
242 | ; m_pred_mp : ∀X.modus_ponens ?? (m_pred X) |
---|
243 | }. |
---|
244 | |
---|
245 | record InjMonadPred (M : Monad) : Type[1] ≝ |
---|
246 | { im_pred :> MonadPred M |
---|
247 | ; mp_inject : ∀X.∀P : X → Prop.(Σm.im_pred P m) → M (Σx.P x) |
---|
248 | ; mp_inject_eq : ∀X,P,m.pi1 ?? m = ! x ← mp_inject X P m ; return pi1 … x |
---|
249 | }. |
---|
250 | |
---|
251 | coercion coerc_mp_inject : ∀M.∀MP:InjMonadPred M. |
---|
252 | ∀X.∀P : X → Prop.∀m : Σm.MP P m.M (Σx.P x) ≝ |
---|
253 | mp_inject on _m : Sig (monad ??) (λm.im_pred ??? m) to monad ? (Sig ? (λx.? x)). |
---|
254 | |
---|
255 | lemma mp_inject_bind : ∀M : MonadProps.∀MP,X,P,Y.∀m.∀f : X → M Y. |
---|
256 | ! x ← mp_inject M MP X P m ; f (pi1 … x) = ! x ← pi1 … m ; f x. |
---|
257 | #M#MP#X#P#Y#m#f >mp_inject_eq >m_bind_bind @m_bind_ext_eq #x >m_return_bind % qed. |
---|
258 | |
---|
259 | record MonadRel (M1 : Monad) (M2 : Monad) : Type[1] ≝ |
---|
260 | { m_rel :3> ∀X,Y.(X → Y → Prop) → (M1 X → M2 Y → Prop) |
---|
261 | ; mr_return : ∀X,Y,rel,x,y.rel x y → m_rel X Y rel (return x) (return y) |
---|
262 | ; 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)) → |
---|
263 | m_rel ?? relout (m_bind … m f) (m_bind … n g) |
---|
264 | ; m_rel_mp : ∀X,Y.rel_modus_ponens ???? (m_rel X Y) |
---|
265 | }. |
---|