]> matita.cs.unibo.it Git - helm.git/commitdiff
...snapshot
authorEnrico Tassi <enrico.tassi@inria.fr>
Thu, 3 Jul 2008 14:20:35 +0000 (14:20 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Thu, 3 Jul 2008 14:20:35 +0000 (14:20 +0000)
helm/software/matita/contribs/dama/dama/models/list_support.ma
helm/software/matita/contribs/dama/dama/models/q_bars.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.    
index c1a017dcccf6ad5cd98f77a732a446b1f6a048a7..d5a7806e799bdc99aaedfa33b0048937920487a2 100644 (file)
@@ -47,6 +47,10 @@ record q_f : Type ≝ {
 lemma nth_nil: ∀T,i.∀def:T. nth [] def i = def.
 intros; elim i; simplify; [reflexivity;] assumption; qed.
 
+lemma len_concat: ∀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).
 
@@ -90,6 +94,77 @@ cases (cmp_nat i (len l));
     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.    
+
+
+alias symbol "pi1" = "pair pi1".
+alias symbol "lt" (instance 6) = "Q less than".
+alias symbol "lt" (instance 2) = "Q less than".
+alias symbol "and" = "logical and".
+lemma sorted_pivot:
+  ∀l1,l2,p. sorted (l1@p::l2) → 
+    (∀i. i < len l1 → nth_base l1 i < \fst p) ∧
+    (∀i. i < len l2 → \fst p < nth_base l2 i).
+intro l; elim l; 
+[1: split; [intros; cases (not_le_Sn_O ? H1);] intros;
+    apply sorted_head_smaller; assumption;
+|2: cases (H ?? (sorted_tail a (l1@p::l2) H1));
+    lapply depth = 0 (sorted_head_smaller (l1@p::l2) a H1) as Hs;
+    split; simplify; intros;
+    [1: cases i in H4; simplify; intros; 
+        [1: lapply depth = 0 (Hs (len l1)) as HS; 
+            unfold nth_base in HS; rewrite > nth_len in HS; apply HS;
+            rewrite > len_concat; simplify; apply lt_n_plus_n_Sm;
+        |2: apply (H2 n); apply le_S_S_to_le; apply H4]
+    |2: apply H3; assumption]]
+qed.
+
 definition eject_NxQ ≝
   λP.λp:∃x:nat × ℚ.P x.match p with [ex_introT p _ ⇒ p].
 coercion eject_NxQ.
@@ -105,12 +180,14 @@ definition value :  ∀f:q_f.∀i:ratio.∃p:ℚ.∃j.value_spec f (Qpos i) 〈j
 intros;
 alias symbol "pi2" = "pair pi2".
 alias symbol "pi1" = "pair pi1".
-alias symbol "lt" (instance 6) = "Q less than".
+alias symbol "lt" (instance 7) = "Q less than".
 alias symbol "leq" = "Q less or equal than".
 letin value_spec_aux ≝ (
-  λf,i,q.
-   \snd q = nth_height f (\fst q) ∧  
-    (nth_base f (\fst q) < i ∧ ∀n.(\fst q) < n → n < len f → i ≤ nth_base f n));
+  λf,i,q. And4
+    (\fst q < len f)
+    (\snd q = nth_height f (\fst q))  
+    (nth_base f (\fst q) < i) 
+    (∀n.(\fst q) < n → n < len f → i ≤ nth_base f n));
 alias symbol "lt" (instance 5) = "Q less than".
 letin value ≝ (
   let rec value (acc: nat × ℚ) (l : list bar) on l : nat × ℚ ≝
@@ -122,25 +199,54 @@ letin value ≝ (
         | q_gt _ ⇒ acc]]
   in value :
   ∀acc,l.∃p:nat × ℚ.
-    ∀story. story @ l = bars f → 
+    ∀story. story @ l = bars f → S (\fst acc) = len story →
     value_spec_aux story (Qpos i) acc → 
-    value_spec_aux l (Qpos i) p);
+    value_spec_aux (story @ l) (Qpos i) p);
 [4: clearbody value;  unfold value_spec;
     generalize in match (bars_begin_OQ f);
     generalize in match (bars_sorted f);
     cases (bars_not_nil f) in value; intros (value S); generalize in match (sorted_tail_bigger ?? S);
     clear S; cases (value 〈O,\snd x〉 l) (p Hp); intros; 
-    exists[apply (\snd p)];exists [apply (S (\fst p))] simplify; 
-    cases (Hp [x] (refl_eq ??) ?) (Hg HV); 
-    [unfold; split[reflexivity]simplify;split;
-      [rewrite > H;apply q_pos_OQ;
-      |intros; cases n in H2 H3; [intro X; cases (not_le_Sn_O ? X)]
-       intros; cases (not_le_Sn_O ? (le_S_S_to_le ?? H3))]]
-    split;[rewrite > Hg; reflexivity]split; cases HV; [assumption;]
+    exists[apply (\snd p)];exists [apply (\fst p)] simplify; 
+    cases (Hp [x] (refl_eq ??) (refl_eq ??) ?) (Hg HV); 
+    [unfold; split; [apply le_n|reflexivity|rewrite > H; apply q_pos_OQ;]
+     intros; cases n in H2 H3; [intro X; cases (not_le_Sn_O ? X)]
+     intros; cases (not_le_Sn_O ? (le_S_S_to_le (S n1) O H3))]
+    split;[rewrite > HV; reflexivity] split; [assumption;]
     intros; cases n in H4 H5; intros [cases (not_le_Sn_O ? H4)]
-    apply (H3 n1);apply le_S_S_to_le; assumption;
-|1: unfold value_spec_aux; clear value value_spec_aux H2;intros; split[2:split]
-    [1: apply (q_lt_le_trans ??? (H4 (\fst p))); clear H4 H5; 
+    apply (H3 (S n1)); assumption;
+|1: unfold value_spec_aux; clear value value_spec_aux H2; intros; 
+    cases H4; clear H4; split;
+    [1: apply (trans_lt ??? H5); rewrite > len_concat; simplify; apply lt_n_plus_n_Sm;
+    |2: unfold nth_height; rewrite > nth_concat_lt_len;[2:assumption]assumption; 
+    |3: unfold nth_base; rewrite > nth_concat_lt_len;[2:assumption]
+        apply (q_le_lt_trans ???? H7); apply q_le_n; 
+    |4: intros; (*clear H6 H5 H4 H l;*) lapply (bars_sorted f) as HS;
+        apply (all_bigger_can_concat_bigger story l1 (S (\fst p)));[6:apply q_lt_to_le]try assumption;
+        [1: rewrite < H2 in HS; cases (sorted_pivot ??? HS); assumption
+        |2: rewrite < H2 in HS; cases (sorted_pivot ??? HS); 
+            intros; apply q_lt_to_le; apply H11; assumption;
+        |3: intros; apply H8; assumption;]]
+|3: intro; rewrite > append_nil; intros; assumption;
+|2: intros; cases (value 〈S (\fst p),\snd b〉 l1); unfold; simplify;
+    cases (H6 (story@[b]) ???); 
+    [1: rewrite > associative_append; apply H3;
+    |2: simplify; rewrite > H4; rewrite > len_concat; rewrite > sym_plus; reflexivity;
+    |4: rewrite < (associative_append ? story [b] l1); split; assumption;
+    |3: cases H5; clear H5; split; simplify in match (\snd ?); simplify in match (\fst ?); 
+        [1: rewrite > len_concat; simplify; rewrite < plus_n_SO; apply le_S_S; assumption;
+        |2: 
+        |3: 
+        |4: ]]]
+
+
+
+
+
+
+
+
+
 
 [5: clearbody value; 
     cases (q_cmp i (start f));