]> matita.cs.unibo.it Git - helm.git/commitdiff
new q_function representation
authorEnrico Tassi <enrico.tassi@inria.fr>
Tue, 15 Jul 2008 09:18:23 +0000 (09:18 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Tue, 15 Jul 2008 09:18:23 +0000 (09:18 +0000)
helm/software/matita/contribs/dama/dama/depends
helm/software/matita/contribs/dama/dama/depends.png
helm/software/matita/contribs/dama/dama/models/list_support.ma
helm/software/matita/contribs/dama/dama/models/q_bars.ma
helm/software/matita/contribs/dama/dama/models/q_function.ma
helm/software/matita/contribs/dama/dama/models/q_shift.ma
helm/software/matita/contribs/dama/dama/models/q_value_skip.ma

index b466b2e07a038854d9d9982a25fd57b1a4b82224..419a803694615fe35dc9b61011a3c54f4867f109 100644 (file)
@@ -1,29 +1,29 @@
-sandwich.ma ordered_uniform.ma
-property_sigma.ma ordered_uniform.ma russell_support.ma
-uniform.ma supremum.ma
-bishop_set.ma ordered_set.ma
-sequence.ma nat/nat.ma
 ordered_uniform.ma uniform.ma
-supremum.ma datatypes/constructors.ma nat/plus.ma nat_ordered_set.ma sequence.ma
-property_exhaustivity.ma ordered_uniform.ma property_sigma.ma
-bishop_set_rewrite.ma bishop_set.ma
 cprop_connectives.ma datatypes/constructors.ma logic/equality.ma
-nat_ordered_set.ma bishop_set.ma nat/compare.ma
+russell_support.ma cprop_connectives.ma nat/nat.ma
 lebesgue.ma property_exhaustivity.ma sandwich.ma
 ordered_set.ma cprop_connectives.ma
-russell_support.ma cprop_connectives.ma nat/nat.ma
-models/nat_lebesgue.ma lebesgue.ma models/nat_order_continuous.ma
+bishop_set.ma ordered_set.ma
+bishop_set_rewrite.ma bishop_set.ma
+sandwich.ma ordered_uniform.ma
+nat_ordered_set.ma bishop_set.ma nat/compare.ma
+property_sigma.ma ordered_uniform.ma russell_support.ma
+sequence.ma nat/nat.ma
+property_exhaustivity.ma ordered_uniform.ma property_sigma.ma
+supremum.ma datatypes/constructors.ma nat/plus.ma nat_ordered_set.ma sequence.ma
+uniform.ma supremum.ma
+models/nat_order_continuous.ma models/nat_dedekind_sigma_complete.ma models/nat_ordered_uniform.ma
 models/nat_ordered_uniform.ma bishop_set_rewrite.ma models/nat_uniform.ma ordered_uniform.ma
-models/q_support.ma Q/q/qplus.ma Q/q/qtimes.ma cprop_connectives.ma
+models/q_function.ma Q/q/qtimes.ma models/q_bars.ma
 models/discrete_uniformity.ma bishop_set_rewrite.ma uniform.ma
+models/list_support.ma cprop_connectives.ma list/list.ma nat/le_arith.ma nat/minus.ma
 models/q_bars.ma cprop_connectives.ma models/list_support.ma models/q_support.ma nat_ordered_set.ma
-models/q_function.ma Q/q/qtimes.ma models/q_shift.ma nat_ordered_set.ma
-models/nat_uniform.ma models/discrete_uniformity.ma nat_ordered_set.ma
+models/q_support.ma Q/q/qplus.ma Q/q/qtimes.ma cprop_connectives.ma
+models/nat_lebesgue.ma lebesgue.ma models/nat_order_continuous.ma
+models/q_shift.ma 
 models/nat_dedekind_sigma_complete.ma models/nat_uniform.ma nat/le_arith.ma russell_support.ma supremum.ma
-models/q_value_skip.ma models/q_shift.ma
-models/q_shift.ma models/q_bars.ma
-models/list_support.ma list/list.ma
-models/nat_order_continuous.ma models/nat_dedekind_sigma_complete.ma models/nat_ordered_uniform.ma
+models/nat_uniform.ma models/discrete_uniformity.ma nat_ordered_set.ma
+models/q_value_skip.ma 
 Q/q/qplus.ma 
 Q/q/qtimes.ma 
 datatypes/constructors.ma 
@@ -31,5 +31,6 @@ list/list.ma
 logic/equality.ma 
 nat/compare.ma 
 nat/le_arith.ma 
+nat/minus.ma 
 nat/nat.ma 
 nat/plus.ma 
index 95829b4db95a05ae707ba728c11128866fbbf9ff..4fb8e3fd71e93076cfc0327fb6c91e1b028251ef 100644 (file)
Binary files a/helm/software/matita/contribs/dama/dama/depends.png and b/helm/software/matita/contribs/dama/dama/depends.png differ
index 8dad0c436b9b224bcb73e7527e86bba5fde54e1b..4a335e01edeb8b941dc749250ba16e3880763861 100644 (file)
@@ -68,7 +68,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 +97,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 +147,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,6 +155,22 @@ lemma case_b : ∀A:Type.∀f:A → bool. ∀x.cases_bool (f x) (f x).
 intros; cases (f x);[left;|right] reflexivity;
 qed. 
 
+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 "cprop_connectives.ma".
 
 definition eject_N ≝
@@ -178,7 +179,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  
index 65066590f4baef3f754305ba709df1353d149bae..7279fe8c04bad616587f9e57e970d551e8248a6d 100644 (file)
@@ -17,12 +17,12 @@ include "models/q_support.ma".
 include "models/list_support.ma". 
 include "cprop_connectives.ma". 
 
-definition 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 ≝ 〈Qpos one,OQ〉.
+definition empty_bar : bar ≝ 〈Qpos one,〈OQ,OQ〉〉.
 notation "\rect" with precedence 90 for @{'empty_bar}.
 interpretation "q0" 'empty_bar = empty_bar.
 
@@ -55,7 +55,7 @@ record q_f : Type ≝ {
  bars: list bar; 
  bars_sorted : sorted q2_lt bars;
  bars_begin_OQ : nth_base bars O = OQ;
- bars_end_OQ : nth_height bars (pred (\len bars)) = OQ
+ bars_end_OQ : nth_height bars (pred (\len bars)) = 〈OQ,OQ〉
 }.
 
 lemma len_bases_gt_O: ∀f.O < \len (bars f).
@@ -75,62 +75,17 @@ cases (cmp_nat (\len l) i);
     apply (H2 n1); simplify in H3; 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 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. 
-*)
-
-
-inductive value_spec (f : q_f) (i : ℚ) : ℚ → nat → CProp ≝
- value_of : ∀q,j. 
-   nth_height (bars f) j = q →  
-   nth_base (bars f) j < i →
-   (∀n.j < n → n < \len (bars f) → i ≤ nth_base (bars f) n) → value_spec f i q j. 
-     
-     
-inductive 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.
-
-definition value :  ∀f:q_f.∀i:ratio.∃p:ℚ.∃j.value_spec f (Qpos i) p j.
+coinductive value_spec (f : q_f) (i : ℚ) : ℚ × ℚ → CProp ≝
+| value_of : ∀j,q. 
+    nth_height (bars f) j = q →  nth_base (bars f) j < i →
+    (∀n.j < n → n < \len (bars f) → i ≤ nth_base (bars f) n) → value_spec f i q. 
+         
+definition value_lemma :  ∀f:q_f.∀i:ratio.∃p:ℚ×ℚ.value_spec f (Qpos i) p.
 intros;
 letin P ≝ 
   (λx:bar.match q_cmp (Qpos i) (\fst x) with[ q_leq _ ⇒ true| q_gt _ ⇒ false]);
 exists [apply (nth_height (bars f) (pred (find ? P (bars f) ▭)));]
-exists [apply (pred (find ? P (bars f) ▭))] apply value_of;
+apply (value_of ?? (pred (find ? P (bars f) ▭)));
 [1: reflexivity
 |2: cases (cases_find bar P (bars f) ▭);
     [1: cases i1 in H H1 H2 H3; simplify; intros;
@@ -175,55 +130,21 @@ exists [apply (pred (find ? P (bars f) ▭))] apply value_of;
         [ apply le_O_n; | assumption]]]
 qed.    
 
-lemma value_OQ_l:
-  ∀l,i.i < start l → \snd (\fst (value l i)) = OQ.
-intros; cases (value l i) (q Hq); cases Hq; clear Hq; simplify; cases H1; clear H1;
-try assumption; cases H2; cases (?:False); apply (q_lt_le_incompat ?? H H6);
-qed.
-    
-lemma value_OQ_r:
-  ∀l,i.start l + sum_bases (bars l) (len (bars l)) ≤ i → \snd (\fst (value l i)) = OQ.
-intros; cases (value l i) (q Hq); cases Hq; clear Hq; simplify; cases H1; clear H1;
-try assumption; cases H2; cases (?:False); apply (q_lt_le_incompat ?? H7 H);
-qed.
-    
-lemma value_OQ_e:
-  ∀l,i.bars l = [] → \snd (\fst (value l i)) = OQ.
-intros; cases (value l i) (q Hq); cases Hq; clear Hq; simplify; cases H1; clear H1;
-try assumption; cases H2; cases (?:False); apply (H1 H);
-qed.
+lemma value : q_f → ratio → ℚ × ℚ.
+intros; cases (value_lemma q r); apply w; qed.
 
-inductive value_ok_spec (f : q_f) (i : ℚ) : nat × ℚ → Type ≝
- | value_ok : ∀n,q. n ≤ (len (bars f)) → 
-      q = \snd (nth (bars f) ▭ n) →
-      sum_bases (bars f) n ≤ ⅆ[i,start f] →
-           ⅆ[i, start f] < sum_bases (bars f) (S n) → value_ok_spec f i 〈n,q〉.
-  
-lemma value_ok:
-  ∀f,i.bars f ≠ [] → start f ≤ i → i < start f + sum_bases (bars f) (len (bars f)) → 
-    value_ok_spec f i (\fst (value f i)). 
-intros; cases (value f i); simplify;   
-cases H3; simplify; clear H3; cases H4; clear H4;
-[1,2,3: cases (?:False); 
-  [1: apply (q_lt_le_incompat ?? H3 H1);
-  |2: apply (q_lt_le_incompat ?? H2 H3);
-  |3: apply (H H3);]
-|4: cases H7; clear H7; cases w in H3 H4 H5 H6 H8; simplify; intros;
-    constructor 1; assumption;]
-qed.
+lemma cases_value : ∀f,i. value_spec f (Qpos i) (value f i).
+intros; unfold value; cases (value_lemma f i); assumption; qed.
 
-definition same_values ≝
-  λl1,l2:q_f.
-   ∀input.\snd (\fst (value l1 input)) = \snd (\fst (value l2 input)). 
+definition same_values ≝ λl1,l2:q_f.∀input. value l1 input = value l2 input. 
 
-definition same_bases ≝ 
-  λl1,l2:list bar. (∀i.\fst (nth l1 ▭ i) = \fst (nth l2 ▭ i)).
+definition same_bases ≝ λ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.
 intro; cases x; intros; [2:exists [apply r] reflexivity]
 cases (?:False);
-[ apply (q_lt_corefl ? H)|apply (q_neg_gt ? H)]
+[ apply (q_lt_corefl ? H)| cases (q_lt_le_incompat ?? (q_neg_gt ?) (q_lt_to_le ?? H))]
 qed.
 
 notation < "x \blacksquare" non associative with precedence 50 for @{'unpos $x}.
index b1ff639e44b959bd6a2a482c001df38d552316f9..346ee58437beab3ab5b53bef645f1a1dd99e211e 100644 (file)
@@ -12,8 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "nat_ordered_set.ma".
-include "models/q_shift.ma".
+include "models/q_bars.ma".
 
 alias symbol "pi2" = "pair pi2".
 alias symbol "pi1" = "pair pi1".
index a630666595aaee63616e187c98b7bedb54917980..e69de29bb2d1d6434b8b29ae775ad8c2e48c5391 100644 (file)
@@ -1,106 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||M||                                                             *)
-(*      ||A||       A project by Andrea Asperti                           *)
-(*      ||T||                                                             *)
-(*      ||I||       Developers:                                           *)
-(*      ||T||         The HELM team.                                      *)
-(*      ||A||         http://helm.cs.unibo.it                             *)
-(*      \   /                                                             *)
-(*       \ /        This file is distributed under the terms of the       *)
-(*        v         GNU General Public License Version 2                  *)
-(*                                                                        *)
-(**************************************************************************)
-
-include "models/q_bars.ma".
-
-lemma initial_shift_same_values:
-  ∀l1:q_f.∀init.init < start l1 →
-   same_values l1 
-     (mk_q_f init (〈\fst (unpos (start l1 - init) ?),OQ〉:: bars l1)).  
-[apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ; assumption]
-intros; generalize in ⊢ (? ? (? ? (? ? (? ? ? (? ? ? (? ? %)) ?) ?))); intro;
-cases (unpos (start l1-init) H1); intro input;
-simplify in ⊢ (? ? ? (? ? ? (? ? ? (? (? ? (? ? (? ? ? % ?) ?)) ?))));
-cases (value (mk_q_f init (〈w,OQ〉::bars l1)) input) (v1 Hv1); 
-cases Hv1 (HV1 HV1 HV1 HV1); cases HV1 (Hi1 Hv11 Hv12); clear HV1 Hv1;
-[1: cut (input < start l1) as K;[2: apply (q_lt_trans ??? Hi1 H)]
-    simplify in ⊢ (? ? ? (? ? ? %));
-    rewrite > (value_OQ_l ?? K); simplify; symmetry; assumption;
-|2: cut (start l1 + sum_bases (bars l1) (len (bars l1)) ≤ input) as K;[2: 
-      simplify in Hi1; apply (q_le_trans ???? Hi1); rewrite > H2;
-      rewrite > q_plus_sym in ⊢ (? ? (? ? %));
-      rewrite > q_plus_assoc; rewrite > q_elim_minus;
-      rewrite > q_plus_sym in ⊢ (? ? (? (? ? %) ?));
-      rewrite > q_plus_assoc; rewrite < q_elim_minus;
-      rewrite > q_plus_minus; rewrite > q_plus_sym in ⊢ (? ? (? % ?));
-      rewrite > q_plus_OQ; apply q_eq_to_le; reflexivity;]
-    simplify in ⊢ (? ? ? (? ? ? %));
-    rewrite > (value_OQ_r ?? K); simplify; symmetry; assumption;
-|3: simplify in Hi1; destruct Hi1;
-|4: cut (start l1 ≤ input → \snd (\fst (value l1 input))=\snd v1) as solution;[2:
-      intro H4; cases Hi1; clear Hi1; cases H3; clear H3;
-      simplify in H5 H6 H8 H9 H7:(? ? (? % %)) ⊢ (? ? ? (? ? ? %));     
-      generalize in match (refl_eq ? (bars l1):bars l1 = bars l1);
-      generalize in ⊢ (???% → ?); intro X; cases X; clear X; intro Hb;
-      [1: rewrite > (value_OQ_e ?? Hb); rewrite > Hv12; rewrite > Hb in Hv11 ⊢ %;
-          simplify in Hv11 ⊢ %; cases (\fst v1) in Hv11; [intros; reflexivity]
-          cases n; [intros; reflexivity] intro X; cases (not_le_Sn_O ? (le_S_S_to_le ?? X));
-      |2: cases (value_ok l1 input);
-          [2: rewrite > Hb; intro W; destruct W;
-          |3: assumption;
-          |4: apply (q_lt_le_trans ??? H7); rewrite > H2;
-                rewrite > q_plus_sym; rewrite < q_plus_assoc;
-                rewrite > q_plus_sym; apply q_le_inj_plus_r;
-                apply q_le_minus; apply q_eq_to_le; reflexivity;
-          |1: rewrite > Hv12; rewrite > Hb; clear Hv12; simplify;
-              rewrite > H10; rewrite > Hb;
-              cut (O < \fst v1);[2: cases (\fst v1) in H9; intros; [2: autobatch]
-              cases (?:False); generalize in match H9;
-                rewrite > q_d_sym; rewrite > q_d_noabs; [2,4: assumption]
-                rewrite > H2; simplify; rewrite > q_plus_sym; rewrite > q_plus_OQ;
-                repeat rewrite > q_elim_minus;
-                intro X; lapply (q_lt_canc_plus_r ??? X) as Y;
-                apply (q_lt_le_incompat ?? Y); assumption;]
-              cases (\fst v1) in H8 H9 Hcut; [1:intros (_ _ X); cases (not_le_Sn_O ? X)]
-              intros; clear H13; simplify;
-              rewrite > (sum_bases_n_m n n1 (b::l)); [1,4: reflexivity] rewrite < Hb;
-              [2: simplify in H8; apply (q_le_lt_trans ??? (q_le_plus_r ??? H8));
-                    apply (q_le_lt_trans ???? H12); rewrite > H2;
-                    rewrite > q_d_sym; rewrite > q_d_noabs; [2: assumption]
-                    rewrite > (q_elim_minus (start l1) init); rewrite > q_minus_distrib;
-                    rewrite > q_elim_opp; repeat rewrite > q_elim_minus;
-                    rewrite < q_plus_assoc; rewrite > (q_plus_sym ? init);
-                    rewrite > q_plus_assoc;rewrite < q_plus_assoc in ⊢ (? (? % ?) ?);
-                    rewrite > (q_plus_sym ? init); do 2 rewrite < q_elim_minus;
-                    rewrite > q_plus_minus; rewrite > q_plus_OQ;
-                    rewrite > q_d_sym; rewrite > q_d_noabs; [2:assumption] 
-                    apply q_eq_to_le; reflexivity;
-              |*: apply (q_le_lt_trans ??? H11);
-                  rewrite > q_d_sym; rewrite > q_d_noabs; [2:assumption;]
-                  generalize in match H9; rewrite > q_d_sym; rewrite > q_d_noabs; [2: assumption]  
-                  rewrite > H2; intro X; 
-                  lapply (q_lt_inj_plus_r ?? (Qopp (start l1-init)) X) as Y; clear X;
-                  rewrite < q_plus_assoc in Y; repeat rewrite < q_elim_minus in Y;
-                  rewrite > q_plus_minus in Y; rewrite > q_plus_OQ in Y;
-                  apply (q_le_lt_trans ???? Y); 
-                  rewrite > (q_elim_minus (start l1) init); rewrite > q_minus_distrib;
-                  rewrite > q_elim_opp; repeat rewrite > q_elim_minus;
-                  rewrite < q_plus_assoc; rewrite > (q_plus_sym ? init);
-                  rewrite > q_plus_assoc;rewrite < q_plus_assoc in ⊢ (? ? (? % ?));
-                  rewrite > (q_plus_sym ? init); rewrite < (q_elim_minus init);
-                  rewrite > q_plus_minus; rewrite > q_plus_OQ;
-                  apply q_eq_to_le; reflexivity;]]]]]
-    cases (q_cmp input (start l1));
-      [2: simplify in ⊢ (? ? ? (? ? ? %)); rewrite > (value_OQ_l ?? H4); 
-        change with (OQ = \snd v1); rewrite > Hv12;
-        cases H3; clear H3; simplify in H5; cases (\fst v1) in H5;[intros;reflexivity]
-        simplify; rewrite > q_d_sym; rewrite > q_d_noabs; [2:cases Hi1; apply H5]
-        rewrite > H2; do 2 rewrite > q_elim_minus;rewrite > q_plus_assoc;
-        intro X; lapply (q_le_canc_plus_r ??? X) as Y; clear X; 
-        cases (?:False); apply (q_lt_le_incompat input (start l1)); try assumption;
-        apply (q_le_S ???? Y); try assumption; apply sum_bases_ge_OQ; 
-      |1: apply solution; apply (q_eq_to_le ?? (sym_eq ??? H4));
-      |3: apply solution; apply (q_lt_to_le ?? H4);]
-qed.        
-
index e639582e1d7fc856baf39bdd67735b260b4d1d99..e69de29bb2d1d6434b8b29ae775ad8c2e48c5391 100644 (file)
@@ -1,34 +0,0 @@
-(**************************************************************************)
-(*       ___                                                              *)
-(*      ||M||                                                             *)
-(*      ||A||       A project by Andrea Asperti                           *)
-(*      ||T||                                                             *)
-(*      ||I||       Developers:                                           *)
-(*      ||T||         The HELM team.                                      *)
-(*      ||A||         http://helm.cs.unibo.it                             *)
-(*      \   /                                                             *)
-(*       \ /        This file is distributed under the terms of the       *)
-(*        v         GNU General Public License Version 2                  *)
-(*                                                                        *)
-(**************************************************************************)
-
-include "models/q_shift.ma".
-
-lemma q_cmp2:
-  ∀a,b:ℚ.a < b ∨ b ≤ a.
-intros; cases (q_cmp a b);
-[1: right; apply q_eq_to_le; symmetry; assumption; 
-|2: left; assumption;
-|3: right; apply q_lt_to_le; assumption;]
-qed.
-
-
-lemma value_skip:
- ∀f,i,b,tl. bars f = b::tl → start f + Qpos (\fst b) < i →
-  \fst (value f i) = \fst (value (mk_q_f (start f + Qpos (\fst b)) tl) i).
-intros; cases (value f i); cases H2; clear H2;
-[1: cases (?:False); apply (q_lt_corefl (start f)); cases H3; clear H3 H4 H5;
-    apply (q_lt_trans ???? H2);
-|2:
-simplify in ⊢ (? ? % ?);