]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/models/q_bars.ma
...snapshot
[helm.git] / helm / software / matita / contribs / dama / dama / models / q_bars.ma
index 2186890f0cf29b007e36fd04eb82380637633e49..d5a7806e799bdc99aaedfa33b0048937920487a2 100644 (file)
@@ -17,137 +17,237 @@ include "models/q_support.ma".
 include "models/list_support.ma".
 include "cprop_connectives.ma". 
 
-definition bar ≝ ratio × ℚ. (* base (Qpos) , height *)
-record q_f : Type ≝ { start : ℚ; bars: list bar }.
+definition bar ≝ ℚ × ℚ.
 
 notation < "\rationals \sup 2" non associative with precedence 90 for @{'q2}.
 interpretation "Q x Q" 'q2 = (Prod Q Q).
 
-definition empty_bar : bar ≝ 〈one,OQ〉.
+definition empty_bar : bar ≝ 〈Qpos one,OQ〉.
 notation "\rect" with precedence 90 for @{'empty_bar}.
 interpretation "q0" 'empty_bar = empty_bar.
 
 notation < "\ldots\rect\square\EmptySmallSquare\ldots" with precedence 90 for @{'lq2}.
-interpretation "lq2" 'lq2 = (list bar).
-
-let rec sum_bases (l:list bar) (i:nat) on i ≝
-    match i with
-    [ O ⇒ OQ
-    | S m ⇒ 
-         match l with
-         [ nil ⇒ sum_bases [] m + Qpos one
-         | cons x tl ⇒ sum_bases tl m + Qpos (\fst x)]].
-         
-axiom sum_bases_empty_nat_of_q_ge_OQ:
-  ∀q:ℚ.OQ ≤ sum_bases [] (nat_of_q q). 
-axiom sum_bases_empty_nat_of_q_le_q:
-  ∀q:ℚ.sum_bases [] (nat_of_q q) ≤ q.
-axiom sum_bases_empty_nat_of_q_le_q_one:
-  ∀q:ℚ.q < sum_bases [] (nat_of_q q) + Qpos one.
-
-lemma sum_bases_ge_OQ:
-  ∀l,n. OQ ≤ sum_bases l n.
-intro; elim l; simplify; intros;
-[1: elim n; [apply q_eq_to_le;reflexivity] simplify;
-    apply q_le_plus_trans; try assumption; apply q_lt_to_le; apply q_pos_lt_OQ;
-|2: cases n; [apply q_eq_to_le;reflexivity] simplify;    
-    apply q_le_plus_trans; [apply H| apply q_lt_to_le; apply q_pos_lt_OQ;]]
+interpretation "lq2" 'lq2 = (list bar). 
+
+inductive sorted : list bar → Prop ≝ 
+| sorted_nil : sorted []
+| sorted_one : ∀x. sorted [x]
+| sorted_cons : ∀x,y,tl. \fst x < \fst y → sorted (y::tl) → sorted (x::y::tl). 
+
+definition nth_base ≝ λf,n. \fst (nth f ▭ n).
+definition nth_height ≝ λf,n. \snd (nth f ▭ n).
+
+record q_f : Type ≝ {
+ bars: list bar; 
+ bars_sorted : sorted bars;
+ bars_begin_OQ : nth_base bars O = OQ;
+ bars_tail_OQ : nth_height bars (pred (len bars)) = OQ
+}.
+
+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.
 
-alias symbol "leq" = "Q less or equal than".
-lemma sum_bases_O:
-  ∀l.∀x.sum_bases l x ≤ OQ → x = O.
-intros; cases x in H; [intros; reflexivity] intro; cases (?:False);
-cases (q_le_cases ?? H); 
-[1: apply (q_lt_corefl OQ); rewrite < H1 in ⊢ (?? %); 
-|2: apply (q_lt_antisym ??? H1);] clear H H1; cases l;
-simplify; apply q_lt_plus_trans;
-try apply q_pos_lt_OQ; 
-try apply (sum_bases_ge_OQ []);
-apply (sum_bases_ge_OQ l1);
+inductive non_empty_list (A:Type) : list A → Type := 
+| show_head: ∀x,l. non_empty_list A (x::l).
+
+lemma bars_not_nil: ∀f:q_f.non_empty_list ? (bars f).
+intro f; generalize in match (bars_begin_OQ f); cases (bars f); 
+[1: intro X; normalize in X; destruct X;
+|2: intros; constructor 1;]
 qed.
 
+lemma sorted_tail: ∀x,l.sorted (x::l) → sorted l.
+intros; inversion H; intros; [destruct H1;|destruct H1;constructor 1;]
+destruct H4; assumption;
+qed.
 
-lemma sum_bases_increasing:
-  ∀l.∀n1,n2:nat.n1<n2→sum_bases l n1 < sum_bases l n2.                           
-intro; elim l 0;
-[1: intros 2; apply (cic:/matita/dama/nat_ordered_set/nat_elim2.con ???? n1 n2);
-    [1: intro; cases n;
-        [1: intro X; cases (not_le_Sn_O ? X);
-        |2: simplify; intros; apply q_lt_plus_trans;
-            [1: apply sum_bases_ge_OQ;|2: apply (q_pos_lt_OQ one)]]
-    |2: simplify; intros;  cases (not_le_Sn_O ? H);
-    |3: simplify; intros; apply q_lt_inj_plus_r;
-        apply H; apply le_S_S_to_le; apply H1;]
-|2:  intros 5; apply (cic:/matita/dama/nat_ordered_set/nat_elim2.con ???? n1 n2);
-    [1: simplify; intros; cases n in H1; intros;
-        [1: cases (not_le_Sn_O ? H1);
-        |2: simplify; apply q_lt_plus_trans;
-            [1: apply sum_bases_ge_OQ;|2: apply q_pos_lt_OQ]]
-    |2: simplify; intros; cases (not_le_Sn_O ? H1);
-    |3: simplify; intros; apply q_lt_inj_plus_r; apply H;
-        apply le_S_S_to_le; apply H2;]]
+lemma sorted_skip: ∀x,y,l. sorted (x::y::l) → sorted (x::l).
+intros; inversion H; intros; [1,2: destruct H1]
+destruct H4; inversion H2; intros; [destruct H4]
+[1: destruct H4; constructor 2;
+|2: destruct H7; constructor 3; [apply (q_lt_trans ??? H1 H4);]
+    apply (sorted_tail ?? H2);]
 qed.
 
-lemma sum_bases_n_m:
-  ∀n,m,l.
-    sum_bases l n < sum_bases l (S m) →
-    sum_bases l m < sum_bases l (S n) →
-    n = m.
-intros 2; apply (nat_elim2 ???? n m);
-[1: intro X; cases X; intros; [reflexivity] cases (?:False);
-    cases l in H H1; simplify; intros;
-    apply (q_lt_le_incompat ??? (sum_bases_ge_OQ ? n1));
-    apply (q_lt_canc_plus_r ??? H1); 
-|2: intros 2; cases l; simplify; intros; cases (?:False); 
-    apply (q_lt_le_incompat ??? (sum_bases_ge_OQ ? n1));
-    apply (q_lt_canc_plus_r ??? H); (* magia ... *) 
-|3: intros 4; cases l; simplify; intros; 
-    [1: rewrite > (H []); [reflexivity]
-        apply (q_lt_canc_plus_r ??(Qpos one)); assumption;
-    |2: rewrite > (H l1); [reflexivity]
-        apply (q_lt_canc_plus_r ??(Qpos (\fst b))); assumption;]]
+lemma sorted_tail_bigger : ∀x,l.sorted (x::l) → ∀i. i < len l → \fst x < nth_base l i.
+intros 2; elim l; [ cases (not_le_Sn_O i H1);]
+cases i in H2;
+[2: intros; apply (H ? 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.
 
-definition eject1 ≝
+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 eject1.
-definition inject1 ≝ λP.λp:nat × ℚ.λh:P p. ex_introT ? P p h.
-coercion inject1 with 0 1 nocomposites.
-
-definition value : 
-  ∀f:q_f.∀i:ℚ.∃p:nat × ℚ. 
-   Or4
-    (And3 (i < start f) (\fst p = O) (\snd p = OQ))
-    (And3 
-     (start f + sum_bases (bars f) (len (bars f)) ≤ i) 
-     (\fst p = O) (\snd p = OQ))
-    (And3 (bars f = []) (\fst p = O) (\snd p = OQ)) 
-    (And4 
-     (And3 (bars f ≠ []) (start f ≤ i) (i < start f + sum_bases (bars f) (len (bars f))))
-     (\fst p ≤ (len (bars f))) 
-     (\snd p = \snd (nth (bars f) ▭ (\fst p)))
-     (sum_bases (bars f) (\fst p) ≤ ⅆ[i,start f] ∧
-       (ⅆ[i, start f] < sum_bases (bars f) (S (\fst p))))).
+coercion eject_NxQ.
+definition inject_NxQ ≝ λP.λp:nat × ℚ.λh:P p. ex_introT ? P p h.
+coercion inject_NxQ with 0 1 nocomposites.
+
+definition value_spec : q_f → ℚ → nat × ℚ → Prop ≝
+  λf,i,q. nth_height (bars f) (\fst q) = \snd q ∧  
+        (nth_base (bars f) (\fst q)  < i ∧
+         ∀n.\fst q < n → n < len (bars f) → i ≤ nth_base (bars f) n). 
+
+definition value :  ∀f:q_f.∀i:ratio.∃p:ℚ.∃j.value_spec f (Qpos i) 〈j,p〉.
 intros;
+alias symbol "pi2" = "pair pi2".
+alias symbol "pi1" = "pair pi1".
+alias symbol "lt" (instance 7) = "Q less than".
+alias symbol "leq" = "Q less or equal than".
+letin value_spec_aux ≝ (
+  λ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 (p: ℚ) (l : list bar) on l ≝
+  let rec value (acc: nat × ℚ) (l : list bar) on l : nat × ℚ ≝
     match l with
-    [ nil ⇒ 〈nat_of_q p,OQ〉
+    [ nil ⇒ acc
     | cons x tl ⇒
-        match q_cmp p (Qpos (\fst x)) with
-        [ q_lt _ ⇒ 〈O, \snd x〉
-        | _ ⇒
-           let rc ≝ value (p - Qpos (\fst x)) tl in
-           〈S (\fst rc),\snd rc〉]]
+        match q_cmp (\fst x) (Qpos i) with
+        [ q_leq _ ⇒ value 〈S (\fst acc), \snd x〉 tl  
+        | q_gt _ ⇒ acc]]
   in value :
-  ∀acc,l.∃p:nat × ℚ.OQ ≤ acc →
-     Or 
-      (And3 (l = []) (\fst p = nat_of_q acc) (\snd p = OQ))
-      (And3
-       (sum_bases l (\fst p) ≤ acc) 
-       (acc < sum_bases l (S (\fst p))) 
-       (\snd p = \snd (nth l ▭ (\fst p)))));
+  ∀acc,l.∃p:nat × ℚ.
+    ∀story. story @ l = bars f → S (\fst acc) = len story →
+    value_spec_aux story (Qpos i) acc → 
+    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 (\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 (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));
     [2: exists [apply 〈O,OQ〉] simplify; constructor 1; split; try assumption; 
@@ -257,15 +357,14 @@ cases H3; simplify; clear H3; cases H4; clear H4;
   |3: apply (H H3);]
 |4: cases H7; clear H7; cases w in H3 H4 H5 H6 H8; simplify; intros;
     constructor 1; assumption;]
-qed.   
+qed.
 
 definition same_values ≝
   λl1,l2:q_f.
    ∀input.\snd (\fst (value l1 input)) = \snd (\fst (value l2 input)). 
 
 definition same_bases ≝ 
-  λl1,l2:q_f.
-    (∀i.\fst (nth (bars l1) ▭ i) = \fst (nth (bars l2) ▭ i)).
+  λl1,l2:list bar. (∀i.\fst (nth l1 ▭ i) = \fst (nth l2 ▭ i)).
 
 alias symbol "lt" = "Q less than".
 lemma unpos: ∀x:ℚ.OQ < x → ∃r:ratio.Qpos r = x.