X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Flist%2Flist.ma;h=077e1af6c443682939461692d20bdfc20f60be22;hb=e880d6eab5e1700f4a625ddcd7d0fa8f0cce2dcc;hp=2e0df971e20be45531554c7839d739cc7605e1c4;hpb=8c4659819a1c1f2e450d9a588ecca37d95ae48e9;p=helm.git diff --git a/helm/software/matita/library/list/list.ma b/helm/software/matita/library/list/list.ma index 2e0df971e..077e1af6c 100644 --- a/helm/software/matita/library/list/list.ma +++ b/helm/software/matita/library/list/list.ma @@ -34,9 +34,8 @@ notation "hvbox(l1 break @ l2)" right associative with precedence 47 for @{'append $l1 $l2 }. -interpretation "nil" 'nil = (cic:/matita/list/list/list.ind#xpointer(1/1/1) _). -interpretation "cons" 'cons hd tl = - (cic:/matita/list/list/list.ind#xpointer(1/1/2) _ hd tl). +interpretation "nil" 'nil = (nil ?). +interpretation "cons" 'cons hd tl = (cons ? hd tl). (* theorem test_notation: [O; S O; S (S O)] = O :: S O :: S (S O) :: []. *) @@ -63,7 +62,7 @@ definition tail := \lambda A:Type. \lambda l: list A. [ nil => [] | (cons hd tl) => tl]. -interpretation "append" 'append l1 l2 = (cic:/matita/list/list/append.con _ l1 l2). +interpretation "append" 'append l1 l2 = (append ? l1 l2). theorem append_nil: \forall A:Type.\forall l:list A.l @ [] = l. intros; @@ -110,8 +109,8 @@ with permut1 : list A -> list A -> Prop \def | step : \forall l1,l2:list A. \forall x,y:A. permut1 ? (l1 @ (x :: y :: l2)) (l1 @ (y :: x :: l2)). -include "nat/nat.ma". - +(* + definition x1 \def S O. definition x2 \def S x1. definition x3 \def S x2. @@ -122,15 +121,9 @@ theorem tmp : permutation nat (x1 :: x2 :: x3 :: []) (x1 :: x3 :: x2 :: []). apply (step ? (x1::[]) [] x2 x3). qed. - -(* theorem nil_append_nil_both: \forall A:Type.\forall l1,l2:list A. l1 @ l2 = [] \to l1 = [] \land l2 = []. -*) - -(* -include "nat/nat.ma". theorem test_notation: [O; S O; S (S O)] = O :: S O :: S (S O) :: []. reflexivity. @@ -140,23 +133,32 @@ theorem test_append: [O;O;O;O;O;O] = [O;O;O] @ [O;O] @ [O]. simplify. reflexivity. qed. + *) -let rec nth (A:Type) l d n on n ≝ - match n with - [ O ⇒ - match l with - [ nil ⇒ d - | cons (x : A) _ ⇒ x - ] - | S n' ⇒ nth A (tail ? l) d n' - ]. +definition nth ≝ + λA:Type. + let rec nth l d n on n ≝ + match n with + [ O ⇒ + match l with + [ nil ⇒ d + | cons (x : A) _ ⇒ x + ] + | S n' ⇒ nth (tail ? l) d n'] + in nth. -let rec map (A,B:Type) (f: A → B) (l : list A) on l : list B ≝ - match l with [ nil ⇒ nil ? | cons x tl ⇒ f x :: (map A B f tl)]. +definition map ≝ + λA,B:Type.λf:A→B. + let rec map (l : list A) on l : list B ≝ + match l with [ nil ⇒ nil ? | cons x tl ⇒ f x :: (map tl)] + in map. -let rec foldr (A,B:Type) (f : A → B → B) (b : B) (l : list A) on l : B := - match l with [ nil ⇒ b | (cons a l) ⇒ f a (foldr ? ? f b l)]. +definition foldr ≝ + λA,B:Type.λf:A→B→B.λb:B. + let rec foldr (l : list A) on l : B := + match l with [ nil ⇒ b | (cons a l) ⇒ f a (foldr l)] + in foldr. definition length ≝ λT:Type.λl:list T.foldr T nat (λx,c.S c) O l. @@ -193,3 +195,9 @@ intros;elim l [simplify;apply le_S_S;assumption |simplify;apply le_S;assumption]] qed. + +lemma length_append : ∀A,l,m.length A (l@m) = length A l + length A m. +intros;elim l +[reflexivity +|simplify;rewrite < H;reflexivity] +qed.