]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/models/list_support.ma
...snapshot
[helm.git] / helm / software / matita / contribs / dama / dama / models / list_support.ma
index d9188b0a6e75f8ed2f5980a8e2dacef032b77f1e..3aec1a0242a1b3c86c7f4743c29c8f5277c21fb9 100644 (file)
@@ -16,8 +16,8 @@ include "list/list.ma".
 
 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" 
+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}.
 
 definition make_list ≝
@@ -26,18 +26,121 @@ 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 _).   
-notation "'mk_list'" with precedence 90 for @{'mk_list}.
-notation < "'mk_list' \nbsp term 90 f \nbsp term 90 n" 
+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}.
+notation "\len" with precedence 90 for @{'len}.
 interpretation "len" 'len = (length _).
-notation < "'len' \nbsp term 90 l" with precedence 69 for @{'len_appl $l}.
+notation < "\len \nbsp term 90 l" with precedence 69 for @{'len_appl $l}.
 interpretation "len appl" 'len_appl l = (length _ l).
 
-lemma len_mk_list : ∀T:Type.∀f:nat→T.∀n.len (mk_list f n) = n.
+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.
 
+inductive sorted (T:Type) (lt : T → T → Prop): list T → Prop ≝ 
+| sorted_nil : sorted T lt []
+| sorted_one : ∀x. sorted T lt [x]
+| sorted_cons : ∀x,y,tl. lt x y → sorted T lt (y::tl) → sorted T lt (x::y::tl). 
+
+lemma nth_nil: ∀T,i.∀def:T. \nth [] def i = def.
+intros; elim i; simplify; [reflexivity;] assumption; qed.
+
+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 := 
+| show_head: ∀x,l. non_empty_list A (x::l).
+
+lemma len_gt_non_empty : 
+  ∀T.∀l:list T.O < \len l → non_empty_list T l.
+intros; cases l in H; [intros; cases (not_le_Sn_O ? H);] intros; constructor 1;
+qed. 
+
+lemma sorted_tail: ∀T,r,x,l.sorted T r (x::l) → sorted T r l.
+intros; inversion H; intros; [destruct H1;|destruct H1;constructor 1;]
+destruct H4; assumption;
+qed.
+
+lemma sorted_skip: 
+ ∀T,r,x,y,l. 
+  transitive T r → sorted T r (x::y::l) → sorted T r (x::l).
+intros; inversion H1; intros; [1,2: destruct H2]
+destruct H5; inversion H3; intros; [destruct H5]
+[1: destruct H5; constructor 2;
+|2: destruct H8; constructor 3; [apply (H ??? H2 H5);]
+    apply (sorted_tail ???? H3);]
+qed.
+
+lemma sorted_tail_bigger : ∀T,r,x,l,d.sorted T r (x::l) → ∀i. i < \len l → r x (\nth l d i).
+intros 4; elim l; [ cases (not_le_Sn_O i H1);]
+cases i in H2;
+[2: intros; apply (H d ? n);[apply (sorted_skip ????? H1)|apply le_S_S_to_le; apply H2]
+|1: intros; inversion H1; intros; [1,2: destruct H3]
+    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.
+
+lemma lt_n_plus_n_Sm : ∀n,m:nat.n < n + S m.
+intros; rewrite > sym_plus; apply (le_S_S n (m+n)); apply (le_plus_n m n); qed.
+
+lemma nth_concat_lt_len:
+  ∀T:Type.∀l1,l2:list T.∀def.∀i.i < len l1 → nth (l1@l2) def i = nth l1 def i.
+intros 4; elim l1; [cases (not_le_Sn_O ? H)] cases i in H H1; simplify; intros;
+[reflexivity| rewrite < H;[reflexivity] apply le_S_S_to_le; apply H1]
+qed. 
+
+lemma nth_concat_ge_len:
+  ∀T:Type.∀l1,l2:list T.∀def.∀i.
+    len l1 ≤ i → nth (l1@l2) def i = nth l2 def (i - len l1).
+intros 4; elim l1; [ rewrite < minus_n_O; reflexivity]
+cases i in H1; simplify; intros; [cases (not_le_Sn_O ? H1)]
+apply H; apply le_S_S_to_le; apply H1;
+qed. 
+
+lemma nth_len:
+  ∀T:Type.∀l1,l2:list T.∀def,x.
+    nth (l1@x::l2) def (len l1) = x.
+intros 2; elim l1;[reflexivity] simplify; apply H; qed.
+
+lemma all_bigger_can_concat_bigger:
+   ∀l1,l2,start,b,x,n.
+    (∀i.i< len l1 → nth_base l1 i < \fst b) →
+    (∀i.i< len l2 → \fst b ≤ nth_base l2 i) →
+    (∀i.i< len l1 → start ≤ i → x ≤ nth_base l1 i) →
+    start ≤ n → n < len (l1@b::l2) → x ≤ \fst b → x ≤ nth_base (l1@b::l2) n.
+intros; cases (cmp_nat n (len l1));
+[1: unfold nth_base;  rewrite > (nth_concat_lt_len ????? H6);
+    apply (H2 n); assumption;
+|2: rewrite > H6; unfold nth_base; rewrite > nth_len; assumption;
+|3: unfold nth_base; rewrite > nth_concat_ge_len; [2: apply lt_to_le; assumption]
+    rewrite > len_concat in H4; simplify in H4; rewrite < plus_n_Sm in H4;
+    lapply linear le_S_S_to_le to H4 as K; rewrite > sym_plus in K;
+    lapply linear le_plus_to_minus to K as X; 
+    generalize in match X; generalize in match (n - len l1); intro W; cases W; clear W X;
+    [intros; assumption] intros;
+    apply (q_le_trans ??? H5); apply (H1 n1); assumption;]
+qed. 
+
+lemma sorted_head_smaller:
+  ∀l,p. sorted (p::l) → ∀i.i < len l → \fst p < nth_base l i.
+intro l; elim l; intros; [cases (not_le_Sn_O ? H1)] cases i in H2; simplify; intros;
+[1: inversion H1; [1,2: simplify; intros; destruct H3] intros; destruct H6; assumption; 
+|2: apply (H p ? n ?); [apply (sorted_skip ??? H1)] apply le_S_S_to_le; apply H2]
+qed.