X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fmodels%2Flist_support.ma;h=eb70322a9a30a0307aceb63500bf1d3085060392;hb=7deb4b1f322850b8ff03d5626f7828736d074ec8;hp=4a335e01edeb8b941dc749250ba16e3880763861;hpb=8d367045e504f594c280d2c87f906695ef9671ee;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/models/list_support.ma b/helm/software/matita/contribs/dama/dama/models/list_support.ma index 4a335e01e..eb70322a9 100644 --- a/helm/software/matita/contribs/dama/dama/models/list_support.ma +++ b/helm/software/matita/contribs/dama/dama/models/list_support.ma @@ -38,6 +38,11 @@ interpretation "len" 'len = (length _). notation < "\len \nbsp term 90 l" with precedence 69 for @{'len_appl $l}. interpretation "len appl" 'len_appl l = (length _ l). +lemma mk_list_ext: ∀T:Type.∀f1,f2:nat→T.∀n. (∀x.x H1; [2: apply le_n] +apply eq_f; apply H; intros; apply H1; apply (trans_le ??? H2); apply le_S; apply le_n; +qed. + lemma len_mk_list : ∀T:Type.∀f:nat→T.∀n.\len (\mk_list f n) = n. intros; elim n; [reflexivity] simplify; rewrite > H; reflexivity; qed. @@ -171,7 +176,7 @@ intros 2; elim n; |1: rewrite < H2; rewrite > len_append; rewrite > plus_n_SO; reflexivity]]] qed. -include "cprop_connectives.ma". +include "logic/cprop_connectives.ma". definition eject_N ≝ λP.λp:∃x:nat.P x.match p with [ex_introT p _ ⇒ p]. @@ -242,3 +247,36 @@ intros; cases (find_lemma ? f l t); apply w; qed. lemma cases_find: ∀T,P,l,d. find_spec T P l d (find T P l d) (find T P l d). intros; unfold find; cases (find_lemma T P l d); simplify; assumption; qed. + +lemma list_elim_with_len: + ∀T:Type.∀P: nat → list T → CProp. + P O [] → (∀l,a,n.P (\len l) l → P (S n) (a::l)) → + ∀l.P (\len l) l. +intros;elim l; [assumption] simplify; apply H1; apply H2; +qed. + +lemma sorted_near: + ∀r,l. sorted r l → ∀i,d. S i < \len l → r (\nth l d i) (\nth l d (S i)). + intros 3; elim H; + [1: cases (not_le_Sn_O ? H1); + |2: simplify in H1; cases (not_le_Sn_O ? (le_S_S_to_le ?? H1)); + |3: simplify; cases i in H4; intros; [apply H1] + apply H3; apply le_S_S_to_le; apply H4] +qed. + +definition last ≝ + λT:Type.λd.λl:list T. \nth l d (pred (\len l)). + +notation > "\last" non associative with precedence 90 for @{'last}. +notation < "\last d l" non associative with precedence 70 for @{'last2 $d $l}. +interpretation "list last" 'last = (last _). +interpretation "list last applied" 'last2 d l = (last _ d l). + +definition head ≝ + λT:Type.λd.λl:list T.\nth l d O. + +notation > "\hd" non associative with precedence 90 for @{'hd}. +notation < "\hd d l" non associative with precedence 70 for @{'hd2 $d $l}. +interpretation "list head" 'hd = (head _). +interpretation "list head applied" 'hd2 d l = (head _ d l). +