]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/models/list_support.ma
nat model ported to the dualized version, but not itself dualized dedekind-sig-compl...
[helm.git] / helm / software / matita / contribs / dama / dama / models / list_support.ma
index 8dad0c436b9b224bcb73e7527e86bba5fde54e1b..eb70322a9a30a0307aceb63500bf1d3085060392 100644 (file)
@@ -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<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]
+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.
@@ -68,7 +73,7 @@ lemma len_append: ∀T:Type.∀l1,l2:list T. \len (l1@l2) = \len l1 + \len l2.
 intros; elim l1; [reflexivity] simplify; rewrite < H; reflexivity;
 qed.
 
-inductive non_empty_list (A:Type) : list A → Type := 
+coinductive non_empty_list (A:Type) : list A → Type := 
 | show_head: ∀x,l. non_empty_list A (x::l).
 
 lemma len_gt_non_empty : 
@@ -97,21 +102,6 @@ cases i in H2;
     destruct H6; simplify; assumption;]
 qed. 
 
-(*
-lemma all_bases_positive : ∀f:q_f.∀i. OQ < nth_base (bars f) (S i).
-intro f; generalize in match (bars_begin_OQ f); generalize in match (bars_sorted f);
-cases (bars_not_nil f); intros;
-cases (cmp_nat i (len l));
-[1: lapply (sorted_tail_bigger ?? H ? H2) as K; simplify in H1;
-    rewrite > H1 in K; apply K;
-|2: rewrite > H2; simplify; elim l; simplify; [apply (q_pos_OQ one)] 
-    assumption;
-|3: simplify; elim l in i H2;[simplify; rewrite > nth_nil; apply (q_pos_OQ one)]
-    cases n in H3; intros; [cases (not_le_Sn_O ? H3)] apply (H2 n1);
-    apply (le_S_S_to_le ?? H3);]
-qed.
-*)
-
 (* move in nat/ *)
 lemma lt_n_plus_n_Sm : ∀n,m:nat.n < n + S m.
 intros; rewrite > sym_plus; apply (le_S_S n (m+n)); alias id "le_plus_n" = "cic:/matita/nat/le_arith/le_plus_n.con".
@@ -162,7 +152,7 @@ intros 2 (r l); elim l;
     |2: apply H3; assumption]]
 qed.
 
-inductive cases_bool (p:bool) : bool → CProp ≝
+coinductive cases_bool (p:bool) : bool → CProp ≝
 | case_true : p = true → cases_bool p true
 | cases_false : p = false → cases_bool p false.
 
@@ -170,7 +160,23 @@ lemma case_b : ∀A:Type.∀f:A → bool. ∀x.cases_bool (f x) (f x).
 intros; cases (f x);[left;|right] reflexivity;
 qed. 
 
-include "cprop_connectives.ma".
+coinductive break_spec (T : Type) (n : nat) (l : list T) : list T → CProp ≝
+| break_to: ∀l1,x,l2. \len l1 = n → l = l1 @ [x] @ l2 → break_spec T n l l.     
+
+lemma list_break: ∀T,n,l. n < \len l → break_spec T n l l.
+intros 2; elim n;
+[1: elim l in H; [cases (not_le_Sn_O ? H)]
+    apply (break_to ?? ? [] a l1); reflexivity;
+|2: cases (H l); [2: apply lt_S_to_lt; assumption;] cases l2 in H3; intros;
+    [1: rewrite < H2 in H1; rewrite > H3 in H1; rewrite > append_nil in H1;
+        rewrite > len_append in H1; rewrite > plus_n_SO in H1;
+        cases (not_le_Sn_n ? H1);
+    |2: apply (break_to ?? ? (l1@[x]) t l3); 
+        [2: simplify; rewrite > associative_append; assumption;
+        |1: rewrite < H2; rewrite > len_append; rewrite > plus_n_SO; reflexivity]]]
+qed.
+
+include "logic/cprop_connectives.ma".
 
 definition eject_N ≝
   λP.λp:∃x:nat.P x.match p with [ex_introT p _ ⇒ p].
@@ -178,7 +184,7 @@ coercion eject_N.
 definition inject_N ≝ λP.λp:nat.λh:P p. ex_introT ? P p h.
 coercion inject_N with 0 1 nocomposites.
 
-inductive find_spec (T:Type) (P:T→bool) (l:list T) (d:T) (res : nat) : nat → CProp ≝
+coinductive find_spec (T:Type) (P:T→bool) (l:list T) (d:T) (res : nat) : nat → CProp ≝
 | found: 
     ∀i. i < \len l → P (\nth l d i) = true → res = i →  
     (∀j. j < i → P (\nth l d j) = false) → find_spec T P l d res i  
@@ -241,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).
+