From cf89508024f0a19023bb1bac343012d54e860d9d Mon Sep 17 00:00:00 2001 From: Andrea Asperti Date: Tue, 16 Feb 2010 09:21:04 +0000 Subject: [PATCH] More theorems --- helm/software/matita/nlibrary/basics/list.ma | 20 +++++++++++++++++++- 1 file changed, 19 insertions(+), 1 deletion(-) diff --git a/helm/software/matita/nlibrary/basics/list.ma b/helm/software/matita/nlibrary/basics/list.ma index 62210608a..b5664d134 100644 --- a/helm/software/matita/nlibrary/basics/list.ma +++ b/helm/software/matita/nlibrary/basics/list.ma @@ -54,7 +54,12 @@ nlet rec append A (l1: list A) l2 on l1 := [ nil ⇒ l2 | cons hd tl ⇒ hd :: append A tl l2 ]. -ndefinition tail := λA:Type.λl: list A. +ndefinition hd ≝ λA:Type.λl: list A.λd:A. + match l with + [ nil ⇒ d + | cons a _ ⇒ a]. + +ndefinition tail ≝ λA:Type.λl: list A. match l with [ nil ⇒ [] | cons hd tl ⇒ tl]. @@ -76,6 +81,19 @@ ntheorem cons_append_commute: ntheorem append_cons:∀A.∀a:A.∀l,l1.l@(a::l1)=(l@[a])@l1. /2/; nqed. +ntheorem nil_append_elim: ∀A.∀l1,l2: list A.∀P: list A → list A → Prop. + l1@l2 = [] → P (nil A) (nil A) → P l1 l2. +#A;#l1; #l2; #P; ncases l1; nnormalize;//; +#a; #l3; #heq; ndestruct; +nqed. + +ntheorem nil_to_nil: ∀A.∀l1,l2:list A. + l1@l2 = [] → l1 = [] ∧ l2 = []. +#A; #l1; #l2; #isnil; napply (nil_append_elim A l1 l2);/2/; +nqed. + +(* ierators *) + nlet 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)]. -- 2.39.2