include "basics/lists/list.ma". lemma All_append : ∀A,P,l,r. All A P l → All A P r → All A P (l@r). #A #P #l elim l [ // | #hd #tl #IH #r * #H1 #H2 #H3 % // @IH // ] qed. lemma All_append_l : ∀A,P,l,r. All A P (l@r) → All A P l. #A #P #l elim l [ // | #hd #tl #IH #r * #H1 #H2 % /2/ ] qed. lemma All_append_r : ∀A,P,l,r. All A P (l@r) → All A P r. #A #P #l elim l [ // | #h #t #IH #r * /2/ ] qed. (* An alternative form of All that can be easier to use sometimes. *) lemma All_alt : ∀A,P,l. (∀a,pre,post. l = pre@a::post → P a) → All A P l. #A #P #l #H lapply (refl ? l) change with ([ ] @ l) in ⊢ (???% → ?); generalize in ⊢ (???(??%?) → ?); elim l in ⊢ (? → ???(???%) → %); [ #pre #E % | #a #tl #IH #pre #E % [ @(H a pre tl E) | @(IH (pre@[a])) >associative_append @E ] ] qed. let rec All2 (A,B:Type[0]) (P:A → B → Prop) (la:list A) (lb:list B) on la : Prop ≝ match la with [ nil ⇒ match lb with [ nil ⇒ True | _ ⇒ False ] | cons ha ta ⇒ match lb with [ nil ⇒ False | cons hb tb ⇒ P ha hb ∧ All2 A B P ta tb ] ]. lemma All2_length : ∀A,B,P,la,lb. All2 A B P la lb → |la| = |lb|. #A #B #P #la elim la [ * [ // | #x #y * ] | #ha #ta #IH * [ * | #hb #tb * #H1 #H2 whd in ⊢ (??%%); >(IH … H2) @refl ] qed. lemma All2_mp : ∀A,B,P,Q,la,lb. (∀a,b. P a b → Q a b) → All2 A B P la lb → All2 A B Q la lb. #A #B #P #Q #la elim la [ * [ // | #h #t #_ * ] | #ha #ta #IH * [ // | #hb #tb #H * #H1 #H2 % [ @H @H1 | @(IH … H2) @H ] ] ] qed. let rec map_All (A,B:Type[0]) (P:A → Prop) (f:∀a. P a → B) (l:list A) (H:All A P l) on l : list B ≝ match l return λl. All A P l → ? with [ nil ⇒ λ_. nil B | cons hd tl ⇒ λH. cons B (f hd (proj1 … H)) (map_All A B P f tl (proj2 … H)) ] H. include "utilities/monad.ma". definition Append : ∀A.Aop ? (nil A) ≝ λA.mk_Aop ? ? (append ?) ? ? ?. // qed. definition List ≝ MakeMonadProps list (λX,x.[x]) (λX,Y,l,f.\fold [append ?, [ ]]_{x ∈ l} (f x)) ???. normalize [ / by / | #X#m elim m normalize // | #X#Y#Z #m #f#g elim m normalize [//] #x#l' #Hi <(fold_sum ?? (f x) ? [ ] (Append ?)) >Hi // ] qed. unification hint 0 ≔ X ; N ≟ max_def List, M ≟ m_def N (*---------------------------*)⊢ list X ≡ monad M X.