diff options
author | herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2006-10-26 19:20:42 +0000 |
---|---|---|
committer | herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2006-10-26 19:20:42 +0000 |
commit | b0c6538b5746a7507fc34eac2cf718da13dd70a6 (patch) | |
tree | 05be6d072babe70a86f43815bd3f728cbd9cde2a /theories/Lists | |
parent | 940f383ce173c178b9dd1a1d2890c3ae8198e940 (diff) |
Déplacement des propriétés générales de BinList dans List et des tactiques de
BinList dans un nouveau ListTactics.
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@9290 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'theories/Lists')
-rw-r--r-- | theories/Lists/List.v | 22 | ||||
-rw-r--r-- | theories/Lists/ListTactics.v | 69 |
2 files changed, 86 insertions, 5 deletions
diff --git a/theories/Lists/List.v b/theories/Lists/List.v index cd6ce35cf..2232d084e 100644 --- a/theories/Lists/List.v +++ b/theories/Lists/List.v @@ -39,6 +39,12 @@ Section Lists. | x :: _ => value x end. + Definition hd (default:A) (l:list) := + match l with + | nil => default + | x :: _ => x + end. + Definition tail (l:list) : list := match l with | nil => nil @@ -670,21 +676,27 @@ Section ListOps. (** An alternative tail-recursive definition for reverse *) - Fixpoint rev_acc (l l': list A) {struct l} : list A := + Fixpoint rev_append (l l': list A) {struct l} : list A := match l with | nil => l' - | a::l => rev_acc l (a::l') + | a::l => rev_append l (a::l') end. - Lemma rev_acc_rev : forall l l', rev_acc l l' = rev l ++ l'. + Definition rev' l : list A := rev_append l nil. + + Notation rev_acc := rev_append (only parsing). + + Lemma rev_append_rev : forall l l', rev_acc l l' = rev l ++ l'. Proof. induction l; simpl; auto; intros. rewrite <- ass_app; firstorder. Qed. - Lemma rev_alt : forall l, rev l = rev_acc l nil. + Notation rev_acc_rev := rev_append_rev (only parsing). + + Lemma rev_alt : forall l, rev l = rev_append l nil. Proof. - intros; rewrite rev_acc_rev. + intros; rewrite rev_append_rev. apply app_nil_end. Qed. diff --git a/theories/Lists/ListTactics.v b/theories/Lists/ListTactics.v new file mode 100644 index 000000000..961101d65 --- /dev/null +++ b/theories/Lists/ListTactics.v @@ -0,0 +1,69 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(*i $Id$ i*) + +Require Import BinPos. +Require Import List. + +Ltac list_fold_right fcons fnil l := + match l with + | (cons ?x ?tl) => fcons x ltac:(list_fold_right fcons fnil tl) + | nil => fnil + end. + +Ltac list_fold_left fcons fnil l := + match l with + | (cons ?x ?tl) => list_fold_left fcons ltac:(fcons x fnil) tl + | nil => fnil + end. + +Ltac list_iter f l := + match l with + | (cons ?x ?tl) => f x; list_iter f tl + | nil => idtac + end. + +Ltac list_iter_gen seq f l := + match l with + | (cons ?x ?tl) => + let t1 _ := f x in + let t2 _ := list_iter_gen seq f tl in + seq t1 t2 + | nil => idtac + end. + +Ltac AddFvTail a l := + match l with + | nil => constr:(cons a l) + | (cons a _) => l + | (cons ?x ?l) => let l' := AddFvTail a l in constr:(cons x l') + end. + +Ltac Find_at a l := + let rec find n l := + match l with + | nil => fail 100 "anomaly: Find_at" + | (cons a _) => eval compute in n + | (cons _ ?l) => find (Psucc n) l + end + in find 1%positive l. + +Ltac check_is_list t := + match t with + | cons _ ?l => check_is_list l + | nil => idtac + | _ => fail 100 "anomaly: failed to build a canonical list" + end. + +Ltac check_fv l := + check_is_list l; + match type of l with + | list _ => idtac + | _ => fail 100 "anomaly: built an ill-typed list" + end. |