X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Flib%2Fbasics%2Flists%2Flist.ma;h=a34ac3c9cd1c7e9026adf18600f312dd38b75740;hb=99c076c7301579c5028a9d6d0ff680ed9e05574f;hp=a330f6224ef1fdb6c94e356b71955f9bad0129e5;hpb=5d54a6d3a0f22bb8784387c491de7bb66e67b625;p=helm.git diff --git a/matita/matita/lib/basics/lists/list.ma b/matita/matita/lib/basics/lists/list.ma index a330f6224..a34ac3c9c 100644 --- a/matita/matita/lib/basics/lists/list.ma +++ b/matita/matita/lib/basics/lists/list.ma @@ -20,7 +20,7 @@ notation "hvbox(hd break :: tl)" right associative with precedence 47 for @{'cons $hd $tl}. -notation "[ list0 x sep ; ]" +notation "[ list0 term 19 x sep ; ]" non associative with precedence 90 for ${fold right @'nil rec acc @{'cons $x $acc}}. @@ -175,8 +175,7 @@ let rec length (A:Type[0]) (l:list A) on l ≝ [ nil ⇒ 0 | cons a tl ⇒ S (length A tl)]. -notation "|M|" non associative with precedence 65 for @{'norm $M}. -interpretation "norm" 'norm l = (length ? l). +interpretation "list length" 'card l = (length ? l). lemma length_tail: ∀A,l. length ? (tail A l) = pred (length ? l). #A #l elim l // @@ -269,6 +268,32 @@ let rec mem A (a:A) (l:list A) on l ≝ [ nil ⇒ False | cons hd tl ⇒ a=hd ∨ mem A a tl ]. + +lemma mem_append: ∀A,a,l1,l2.mem A a (l1@l2) → + mem ? a l1 ∨ mem ? a l2. +#A #a #l1 elim l1 + [#l2 #mema %2 @mema + |#b #tl #Hind #l2 * + [#eqab %1 %1 @eqab + |#Hmema cases (Hind ? Hmema) -Hmema #Hmema [%1 %2 //|%2 //] + ] + ] +qed. + +lemma mem_append_l1: ∀A,a,l1,l2.mem A a l1 → mem A a (l1@l2). +#A #a #l1 #l2 elim l1 + [whd in ⊢ (%→?); @False_ind + |#b #tl #Hind * [#eqab %1 @eqab |#Hmema %2 @Hind //] + ] +qed. + +lemma mem_append_l2: ∀A,a,l1,l2.mem A a l2 → mem A a (l1@l2). +#A #a #l1 #l2 elim l1 [//|#b #tl #Hind #Hmema %2 @Hind //] +qed. + +lemma mem_single: ∀A,a,b. mem A a [b] → a=b. +#A #a #b * // @False_ind +qed. lemma mem_map: ∀A,B.∀f:A→B.∀l,b. mem ? b (map … f l) → ∃a. mem ? a l ∧ f a = b. @@ -420,6 +445,12 @@ lemma All_nth : ∀A,P,n,l. ] ] qed. +let rec Allr (A:Type[0]) (R:relation A) (l:list A) on l : Prop ≝ +match l with +[ nil ⇒ True +| cons a1 l ⇒ match l with [ nil ⇒ True | cons a2 _ ⇒ R a1 a2 ∧ Allr A R l ] +]. + (**************************** Exists *******************************) let rec Exists (A:Type[0]) (P:A → Prop) (l:list A) on l : Prop ≝