]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/models/list_support.ma
nasty change in the lexer/parser:
[helm.git] / helm / software / matita / contribs / dama / dama / models / list_support.ma
index a9d26a9ce890c6636932c5015dbc5ff39a204c60..b9b60f5b9073fa1637ae943a870ee4132340d2f8 100644 (file)
@@ -15,8 +15,8 @@
 include "nat/minus.ma".
 include "list/list.ma".
 
-interpretation "list nth" 'nth = (nth _).
-interpretation "list nth" 'nth_appl l d i = (nth _ l d i).
+interpretation "list nth" 'nth = (nth ?).
+interpretation "list nth" 'nth_appl l d i = (nth ? l d i).
 notation "\nth" with precedence 90 for @{'nth}.
 notation < "\nth \nbsp term 90 l \nbsp term 90 d \nbsp term 90 i" 
 with precedence 69 for @{'nth_appl $l $d $i}.
@@ -27,16 +27,16 @@ definition make_list ≝
       match n with [ O ⇒ nil ? | S m ⇒ def m :: make_list m]
     in make_list.
 
-interpretation "\mk_list appl" 'mk_list_appl f n = (make_list _ f n).
-interpretation "\mk_list" 'mk_list = (make_list _).   
+interpretation "\mk_list appl" 'mk_list_appl f n = (make_list ? f n).
+interpretation "\mk_list" 'mk_list = (make_list ?).   
 notation "\mk_list" with precedence 90 for @{'mk_list}.
 notation < "\mk_list \nbsp term 90 f \nbsp term 90 n" 
 with precedence 69 for @{'mk_list_appl $f $n}.
 
 notation "\len" with precedence 90 for @{'len}.
-interpretation "len" 'len = (length _).
+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).
+interpretation "len appl" 'len_appl l = (length ? l).
 
 lemma mk_list_ext: ∀T:Type.∀f1,f2:nat→T.∀n. (∀x.x<n → f1 x = f2 x) → \mk_list f1 n = \mk_list f2 n.
 intros 4;elim n; [reflexivity] simplify; rewrite > H1; [2: apply le_n]
@@ -264,3 +264,19 @@ lemma sorted_near:
      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).
+