]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/models/q_function.ma
...
[helm.git] / helm / software / matita / contribs / dama / dama / models / q_function.ma
index 21e51808a10189754f4e0e7e80736a562c9168df..a38cf6e4cda648377be8ea314956b673a2f2becd 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "Q/q/q.ma".
-include "list/list.ma".
-include "cprop_connectives.ma". 
-
-
-notation "\rationals" non associative with precedence 99 for @{'q}.
-interpretation "Q" 'q = Q. 
-
-definition bar ≝ ratio × ℚ. (* base (Qpos) , height *)
-record q_f : Type ≝ { start : ℚ; bars: list bar }.
-
-axiom qp : ℚ → ℚ → ℚ.
-axiom qm : ℚ → ℚ → ℚ.
-axiom qlt : ℚ → ℚ → CProp.
-
-interpretation "Q plus" 'plus x y = (qp x y).
-interpretation "Q minus" 'minus x y = (qm x y).
-interpretation "Q less than" 'lt x y = (qlt x y).
-
-inductive q_comparison (a,b:ℚ) : CProp ≝
- | q_eq : a = b → q_comparison a b 
- | q_lt : a < b → q_comparison a b
- | q_gt : b < a → q_comparison a b.
-
-axiom q_cmp:∀a,b:ℚ.q_comparison a b.
-
-definition qle ≝ λa,b:ℚ.a = b ∨ a < b.
-
-interpretation "Q less or equal than" 'leq x y = (qle x y).
-
-axiom q_le_minus: ∀a,b,c:ℚ. a ≤ c - b → a + b ≤ c.
-axiom q_lt_plus: ∀a,b,c:ℚ. a - b < c → a < c + b.
-axiom q_lt_minus: ∀a,b,c:ℚ. a + b < c → a < c - b.
-
-axiom q_dist : ℚ → ℚ → ℚ.
-
-notation "hbox(\dd [term 19 x, break term 19 y])" with precedence 90
-for @{'distance $x $y}.
-interpretation "ℚ distance" 'distance x y = (q_dist x y).
-
-axiom q_dist_ge_OQ : ∀x,y:ℚ. OQ ≤ ⅆ[x,y].
-
-axiom q_lt_to_le: ∀a,b:ℚ.a < b → a ≤ b.
-axiom q_le_to_diff_ge_OQ : ∀a,b.a ≤ b → OQ ≤ b-a.
-axiom q_plus_OQ: ∀x:ℚ.x + OQ = x.
-axiom q_plus_sym: ∀x,y:ℚ.x + y = y + x.
-axiom nat_of_q: ℚ → nat. 
-
-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" 
-with precedence 69 for @{'nth_appl $l $d $i}.
-
-notation < "\rationals \sup 2" non associative with precedence 90 for @{'q2}.
-interpretation "Q x Q" 'q2 = (Prod Q Q).
-
-definition make_list ≝
-  λA:Type.λdef:nat→A.
-    let rec make_list (n:nat) on n ≝
-      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" 
-with precedence 69 for @{'mk_list_appl $f $n}.
+include "russell_support.ma".
+include "models/q_bars.ma".
 
+definition rebase_spec ≝ 
+ λl1,l2:q_f.λp:q_f × q_f. 
+   And3
+    (same_bases (bars (\fst p)) (bars (\snd p)))
+    (same_values l1 (\fst p)) 
+    (same_values l2 (\snd p)).
 
-definition empty_bar : bar ≝ 〈one,OQ〉.
-notation "\rect" with precedence 90 for @{'empty_bar}.
-interpretation "q0" 'empty_bar = empty_bar.
+definition same_values_simpl ≝
+ λl1,l2.∀p1,p2,p3,p4,p5,p6.same_values (mk_q_f l1 p1 p2 p3) (mk_q_f l2 p4 p5 p6).
 
-notation < "\ldots\rect\square\EmptySmallSquare\ldots" with precedence 90 for @{'lq2}.
-interpretation "lq2" 'lq2 = (list bar).
+alias symbol "pi2" = "pair pi2".
+alias symbol "pi1" = "pair pi1".
+definition rebase_spec_aux ≝ 
+ λl1,l2:list bar.λp:(list bar) × (list bar).
+   sorted q2_lt l1 → sorted q2_lt l2 →
+   (l1 ≠ [] → \snd (\nth l1 ▭ (pred (\len l1))) = 〈OQ,OQ〉) →
+   (l2 ≠ [] → \snd (\nth l2 ▭ (pred (\len l2))) = 〈OQ,OQ〉) →
+ And4
+   (nth_base l1 O = nth_base (\fst p) O ∨
+    nth_base l2 O = nth_base (\fst p) O) 
+   (sorted q2_lt (\fst p) ∧ sorted q2_lt (\snd p))
+   ((l1 ≠ [] → \snd (\nth (\fst p) ▭ (pred (\len (\fst p)))) = 〈OQ,OQ〉) ∧ 
+    (l2 ≠ [] → \snd (\nth (\snd p) ▭ (pred (\len (\snd p)))) = 〈OQ,OQ〉))  
+   (And3
+      (same_bases (\fst p) (\snd p))
+      (same_values_simpl l1 (\fst p)) 
+      (same_values_simpl l2 (\snd p))).   
 
-notation "'len'" with precedence 90 for @{'len}.
-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).
+definition eject ≝
+  λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
+coercion eject.
+definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h.
+coercion inject with 0 1 nocomposites.
 
-lemma len_mk_list : ∀T:Type.∀f:nat→T.∀n.len (mk_list f n) = n.
-intros; elim n; [reflexivity] simplify; rewrite > H; reflexivity;
+axiom devil : False.
+        
+definition copy ≝
+ λl:list bar.make_list ? (λn.〈nth_base l (\len l - S n),〈OQ,OQ〉〉) (\len l).
+
+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.  
+lemma sorted_copy:
+  ∀l:list bar.sorted q2_lt l → sorted q2_lt (copy l).
+intros 2; unfold copy; generalize in match (le_n (\len l)); 
+elim (\len l) in ⊢ (?%?→? ? (? ? ? %));
+simplify; [apply (sorted_nil q2_lt);] cases n in H1 H2;
+simplify; intros; [apply (sorted_one q2_lt);]
+apply (sorted_cons q2_lt);
+[2: apply H1; apply lt_to_le; apply H2;
+|1: elim l in H2 H; simplify; [simplify in H2; cases (not_le_Sn_O ? H2)]    
+    simplify in H3; unfold nth_base;
+    unfold canonical_q_lt; unfold q2_trel; unfold q2_lt; simplify;
+    change with (q2_lt (\nth (a::l1) ▭ (\len l1-S n1)) (\nth (a::l1) ▭ (\len l1-n1)));
+    cut (∃w.w = \len l1 - S n1); [2: exists[apply (\len l1 - S n1)] reflexivity]
+    cases Hcut; rewrite < H4; rewrite < (?:S w = \len l1 - n1);
+    [1: apply (sorted_near q2_lt (a::l1) H2); rewrite > H4;
+        simplify; apply le_S_S; elim (\len l1) in H3; simplify;
+        [ cases (not_le_Sn_O ? (le_S_S_to_le ?? H3));
+        | lapply le_S_S_to_le to H5 as H6;
+          lapply le_S_S_to_le to H6 as H7; clear H5 H6;
+          cases H7 in H3; intros; [rewrite < minus_n_n; apply le_S_S; apply le_O_n]
+          simplify in H5; apply le_S_S; apply (trans_le ???? (H5 ?));
+          [2: apply le_S_S; apply le_S_S; assumption;
+          |1: apply (lt_minus_S_n_to_le_minus_n n1 (S m) (S (minus m n1)) ?).
+              apply (not_le_to_lt (S (minus m n1)) (minus (S m) (S n1)) ?).
+              apply (not_le_Sn_n (minus (S m) (S n1))).]]
+    |2: rewrite > H4; lapply le_S_S_to_le to H3 as K;
+        clear H4 Hcut H3 H H1 H2; generalize in match K; clear K;
+        apply (nat_elim2 ???? n1 (\len l1)); simplify; intros;
+        [1: rewrite < minus_n_O; cases n2 in H; [intro; cases (not_le_Sn_O ? H)]
+            intros; cases n3; simplify; reflexivity;
+        |2: cases (not_le_Sn_O ? H);
+        |3: apply H; apply le_S_S_to_le; apply H1;]]]
 qed.
 
-let rec sum_bases (l:list bar) (i:nat)on i ≝
-    match i with
-    [ O ⇒ OQ
-    | S m ⇒ 
-         match l with
-         [ nil ⇒ sum_bases l 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.
-
-definition eject1 ≝
-  λ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 × ℚ. 
-   match q_cmp i (start f) with
-   [ q_lt _ ⇒ \snd p = OQ
-   | _ ⇒ 
-        And3
-         (sum_bases (bars f) (\fst p) ≤ ⅆ[i,start f]) 
-         (ⅆ[i, start f] < sum_bases (bars f) (S (\fst p))) 
-         (\snd p = \snd (nth (bars f) ▭ (\fst p)))].
-intros;
-alias symbol "pi2" = "pair pi2".
-alias symbol "pi1" = "pair pi1".
-letin value ≝ (
-  let rec value (p: ℚ) (l : list bar) on l ≝
-    match l with
-    [ nil ⇒ 〈nat_of_q p,OQ〉
-    | 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〉]]
-  in value :
-  ∀acc,l.∃p:nat × ℚ. OQ ≤ acc →
-     And3
-       (sum_bases l (\fst p) ≤ acc) 
-       (acc < sum_bases l (S (\fst p))) 
-       (\snd p = \snd (nth l ▭ (\fst p))));
-[5: clearbody value; 
-    cases (q_cmp i (start f));
-    [2: exists [apply 〈O,OQ〉] simplify; reflexivity;
-    |*: cases (value ⅆ[i,start f] (bars f)) (p Hp);
-        cases (Hp (q_dist_ge_OQ ? ?)); clear Hp value;
-        exists[1,3:apply p]; simplify; split; assumption;]
-|1,3: intros; split;
-    [1,4: clear H2; cases (value (q-Qpos (\fst b)) l1);
-           cases (H2 (q_le_to_diff_ge_OQ ?? (? H1)));
-          [1,3: intros; [right|left;symmetry] assumption]
-          simplify; apply q_le_minus; assumption;
-    |2,5: cases (value (q-Qpos (\fst b)) l1); 
-          cases (H4 (q_le_to_diff_ge_OQ ?? (? H1)));
-          [1,3: intros; [right|left;symmetry] assumption]
-          clear H3 H2 value;
-          change with (q < sum_bases l1 (S (\fst w)) + Qpos (\fst b));
-          apply q_lt_plus; assumption;
-    |*: cases (value (q-Qpos (\fst b)) l1); simplify; 
-        cases (H4 (q_le_to_diff_ge_OQ ?? (? H1))); 
-        [1,3: intros; [right|left;symmetry] assumption]
-        assumption;]
-|2: clear value H2; simplify; intros; split; [assumption|3:reflexivity]
-    rewrite > q_plus_sym; rewrite > q_plus_OQ; assumption;
-|4: simplify; intros; split; 
-    [1: apply sum_bases_empty_nat_of_q_le_q;
-    |2: apply sum_bases_empty_nat_of_q_le_q_one;
-    |3: elim (nat_of_q q); [reflexivity] simplify; assumption]] 
+lemma make_list_ext: ∀T,f1,f2,n. (∀x.x<n → f1 x = f2 x) → make_list T f1 n = make_list T 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.
     
-          
-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)).
-
-axiom q_lt_corefl: ∀x:Q.x < x → False.
-axiom q_lt_antisym: ∀x,y:Q.x < y → y < x → False.
-axiom q_neg_gt: ∀r:ratio.OQ < Qneg r → False.
-axiom q_d_x_x: ∀x:Q.ⅆ[x,x] = OQ.
-axiom q_pos_OQ: ∀x.Qpos x ≤ OQ → False.
-axiom q_lt_plus_trans:
-  ∀x,y:Q.OQ ≤ x → OQ < y → OQ < x + y.
-axiom q_pos_lt_OQ: ∀x.OQ < Qpos x.
-axiom q_le_plus_trans:
-  ∀x,y:Q. OQ ≤ x → OQ ≤ y → OQ ≤ x + y. 
-
-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)]
+lemma len_copy: ∀l. \len l = \len (copy l).
+intro; elim l; [reflexivity] simplify; rewrite > H; clear H;
+apply eq_f; elim (\len (copy l1)) in ⊢ (??%(??(???%))); [reflexivity] simplify;
+rewrite > H in ⊢ (??%?); reflexivity;
+qed. 
+
+lemma same_bases_cons: ∀a,b,l1,l2.
+  same_bases l1 l2 → \fst a = \fst b → same_bases (a::l1) (b::l2).
+intros; intro; cases i; simplify; [assumption;] apply (H n);
 qed.
 
-notation < "\blacksquare" non associative with precedence 90 for @{'hide}.
-definition hide ≝ λT:Type.λx:T.x.
-interpretation "hide" 'hide = (hide _ _).
-
-lemma sum_bases_ge_OQ:
-  ∀l,n. OQ ≤ sum_bases (bars l) n.
-intro; elim (bars l); simplify; intros;
-[1: elim n; [left;reflexivity] simplify;
-    apply q_le_plus_trans; try assumption; apply q_lt_to_le; apply q_pos_lt_OQ;
-|2: cases n; [left;reflexivity] simplify; 
-    apply q_le_plus_trans; [apply H| apply q_lt_to_le; apply q_pos_lt_OQ;]]
+lemma minus_lt : ∀i,j. i < j → j - i = S (j - S i). 
+intros 2;
+apply (nat_elim2 ???? i j); simplify; intros;
+[1: cases n in H; intros; rewrite < minus_n_O; [cases (not_le_Sn_O ? H);]
+    simplify; rewrite < minus_n_O; reflexivity;
+|2: cases (not_le_Sn_O ? H);
+|3: apply H; apply le_S_S_to_le; assumption;]
 qed.
-
-lemma sum_bases_O:
-  ∀l:q_f.∀x.sum_bases (bars l) x ≤ OQ → x = O.
-intros; cases x in H; [intros; reflexivity] intro; cases (?:False);
-cases H; 
-[1: apply (q_lt_corefl OQ); rewrite < H1 in ⊢ (?? %); 
-|2: apply (q_lt_antisym ??? H1);] clear H H1; cases (bars l);
-simplify; apply q_lt_plus_trans;
-try apply q_pos_lt_OQ; 
-try apply (sum_bases_ge_OQ (mk_q_f OQ []));
-apply (sum_bases_ge_OQ (mk_q_f OQ l1));
+    
+lemma copy_same_bases: ∀l. same_bases l (copy l).
+intro; unfold copy; elim l using list_elim_with_len; [1: intro;reflexivity]
+simplify; rewrite < minus_n_n;
+simplify in ⊢ (? ? (? ? (? ? ? % ?) ?));
+apply same_bases_cons; [2: reflexivity]
+cases l1 in H; [intros 2; reflexivity]
+simplify in ⊢ (? ? (? ? (λ_:?.? ? ? (? ? %) ?) ?)→?);
+simplify in ⊢ (?→? ? (? ? (λ_:?.? ? ? (? ? (? % ?)) ?) ?));
+intro; rewrite > (make_list_ext ?? (λn:nat.〈nth_base (b::l2) (\len l2-n),〈OQ,OQ〉〉));[assumption]
+intro; elim x; [simplify; rewrite < minus_n_O; reflexivity]
+simplify in ⊢ (? ? (? ? ? (? ? %) ?) ?);
+simplify in H2:(? ? %); rewrite > minus_lt; [reflexivity] apply le_S_S_to_le;
+assumption;
 qed.
 
-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 hide; 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);
-simplify in ⊢ (? ? ? (? ? ? %));
-cases (q_cmp input (start (mk_q_f init (〈w,OQ〉::bars l1)))) in H3;
-whd in ⊢ (% → ?); simplify in H3;
-[1: intro; cases H4; clear H4; rewrite > H3;
-    cases (value l1 init); simplify; cases (q_cmp init (start l1)) in H4;
-    [1: cases (?:False); apply (q_lt_corefl init); rewrite > H4 in ⊢ (?? %); apply H;
-    |3: cases (?:False); apply (q_lt_antisym init (start l1)); assumption;
-    |2: whd in ⊢ (% → ?); intro; rewrite > H8; clear H8 H4;
-        rewrite > H7; clear H7; rewrite > (?:\fst w1 = O); [reflexivity]
-        symmetry; apply le_n_O_to_eq;
-        rewrite > (sum_bases_O (mk_q_f init (〈w,OQ〉::bars l1)) (\fst w1)); [apply le_n]   
-        clear H6 w2;
-        simplify in H5:(? ? (? ? %));  
-        destruct H3; rewrite > q_d_x_x in H5; assumption;]
-|2: intros;
-  
+lemma copy_rebases: 
+  ∀l1.rebase_spec_aux l1 [] 〈l1, copy l1〉.
+intros; elim l1; intros 4;
+[1: split; [left; reflexivity]; split; try assumption; unfold; intros;
+    unfold same_values; intros; reflexivity;
+|2: rewrite > H3; [2: intro X; destruct X]
+    split; [left; reflexivity] split; 
+    unfold same_values_simpl; unfold same_values; intros; try reflexivity;
+    try assumption; [4: normalize in p2; destruct p2|2: cases H5; reflexivity;]
+    [1: apply (sorted_copy ? H1);
+    |2: apply (copy_same_bases (a::l));]]
+qed.
+               
+lemma copy_rebases_r: 
+  ∀l1.rebase_spec_aux [] l1 〈copy l1, l1〉.
+intros; elim l1; intros 4;
+[1: split; [left; reflexivity]; split; try assumption; unfold; intros;
+    unfold same_values; intros; reflexivity;
+|2: rewrite > H4; [2: intro X; destruct X]
+    split; [right; simplify; rewrite < minus_n_n; reflexivity] split; 
+    unfold same_values_simpl; unfold same_values; intros; try reflexivity;
+    try assumption; [4: normalize in p2; destruct p2|2: cases H5; reflexivity;]
+    [1: apply (sorted_copy ? H2);
+    |2: intro; symmetry; apply (copy_same_bases (a::l));]]
+qed.
+               
+definition rebase: ∀l1,l2:q_f.∃p:q_f × q_f.rebase_spec l1 l2 p.
+intros 2 (f1 f2); cases f1 (b1 Hs1 Hb1 He1); cases f2 (b2 Hs2 Hb2 He2); clear f1 f2;
+alias symbol "plus" = "natural plus".
 alias symbol "pi2" = "pair pi2".
 alias symbol "pi1" = "pair pi1".
-definition rebase_spec ≝ 
- ∀l1,l2:q_f.∃p:q_f × q_f.
-   And4
-    (*len (bars (\fst p)) = len (bars (\snd p))*)
-    (start (\fst p) = start (\snd p))
-    (same_bases (\fst p) (\snd p))
-    (same_values l1 (\fst p)) 
-    (same_values l2 (\snd p)).
-
-definition rebase_spec_simpl ≝ 
- λstart.λl1,l2:list bar.λp:(list bar) × (list bar).
-   And3
-    (same_bases (mk_q_f start (\fst p)) (mk_q_f start (\snd p)))
-    (same_values (mk_q_f start l1) (mk_q_f start (\fst p))) 
-    (same_values (mk_q_f start l2) (mk_q_f start (\snd p))).
-
-(* a local letin makes russell fail *)
-definition cb0h : list bar → list bar ≝ 
-  λl.mk_list (λi.〈\fst (nth l ▭ i),OQ〉) (len l).
-
-definition eject ≝
-  λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
-coercion eject.
-definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h.
-coercion inject with 0 1 nocomposites.
-        
-definition rebase: rebase_spec.
-intros 2 (f1 f2); cases f1 (s1 l1); cases f2 (s2 l2); clear f1 f2;
-letin spec ≝ (
-  λs.λl1,l2.λm.λz.len l1 + len l2 < m → rebase_spec_simpl s l1 l2 z);
-alias symbol "pi1" (instance 34) = "exT \fst".
-alias symbol "pi1" (instance 21) = "exT \fst".
+alias symbol "minus" = "Q minus".
 letin aux ≝ ( 
-let rec aux (l1,l2:list bar) (n:nat) on n : (list bar) × (list bar) ≝
+let rec aux (l1,l2:list bar) (n : nat) on n : (list bar) × (list bar) ≝
 match n with
-[ O ⇒ 〈 nil ? , nil ? 
-| S m ⇒ 
+[ O ⇒ 〈[], []
+| S m ⇒
   match l1 with
-  [ nil ⇒ 〈cb0h l2, l2〉
+  [ nil ⇒ 〈copy l2, l2〉
   | cons he1 tl1 ⇒
      match l2 with
-     [ nil ⇒ 〈l1, cb0h l1〉
+     [ nil ⇒ 〈l1, copy l1〉
      | cons he2 tl2 ⇒  
-         let base1 ≝ Qpos (\fst he1) in
-         let base2 ≝ Qpos (\fst he2) in
-         let height1 ≝ (\snd he1) in
-         let height2 ≝ (\snd he2) in
+         let base1 ≝ \fst he1 in
+         let base2 ≝ \fst he2 in
+         let height1 ≝ \snd he1 in
+         let height2 ≝ \snd he2 in
          match q_cmp base1 base2 with
-         [ q_eq _ ⇒ 
-             let rc ≝ aux tl1 tl2 m in 
-             〈he1 :: \fst rc,he2 :: \snd rc〉 
-         | q_lt Hp ⇒
-             let rest ≝ base2 - base1 in
-             let rc ≝ aux tl1 (〈\fst (unpos rest ?),height2〉 :: tl2) m in
-             〈〈\fst he1,height1〉 :: \fst rc,〈\fst he1,height2〉 :: \snd rc〉 
+         [ q_leq Hp1 ⇒ 
+             match q_cmp base2 base1 with
+             [ q_leq Hp2 ⇒
+                 let rc ≝ aux tl1 tl2 m in 
+                 〈he1 :: \fst rc,he2 :: \snd rc〉
+             | q_gt Hp ⇒ 
+                 let rest ≝ base2 - base1 in
+                 let rc ≝ aux tl1 (〈rest,height2〉 :: tl2) m in
+                 〈〈base1,height1〉 :: \fst rc,〈base1,height2〉 :: \snd rc〉] 
          | q_gt Hp ⇒ 
              let rest ≝ base1 - base2 in
-             let rc ≝ aux (〈\fst (unpos rest ?),height1〉 :: tl1) tl2 m in
-             〈〈\fst he2,height1〉 :: \fst rc,〈\fst he2,height2〉 :: \snd rc〉
-]]]]
-in aux : ∀l1,l2,m.∃z.∀s.spec s l1 l2 m z); unfold spec;
-[9: clearbody aux; unfold spec in aux; clear spec;
-    cases (q_cmp s1 s2);
-    [1: cases (aux l1 l2 (S (len l1 + len l2)));
-        cases (H1 s1 (le_n ?)); clear H1;
-        exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s2 (\snd w)〉] split;
-        [1,2: assumption;
-        |3: intro; apply (H3 input);
-        |4: intro; rewrite > H in H4; 
-            rewrite > (H4 input); reflexivity;]
-    |2: letin l2' ≝ (〈\fst (unpos (s2-s1) ?),OQ〉::l2);[
-          apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
-          assumption]
-        cases (aux l1 l2' (S (len l1 + len l2')));
-        cases (H1 s1 (le_n ?)); clear H1 aux;
-        exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s1 (\snd w)〉] split;
-        [1: reflexivity
-        |2: assumption;
-        |3: assumption;
-        |4: intro; rewrite < (H4 input); clear H3 H4 H2 w;
-            cases (value (mk_q_f s1 l2') input);
-            cases (q_cmp input (start (mk_q_f s1 l2'))) in H1;
-            whd in ⊢ (% → ?);
-            [1: intros; cases H2; clear H2; whd in ⊢ (??? %);
-                cases (value (mk_q_f s2 l2) input);
-                cases (q_cmp input (start (mk_q_f s2 l2))) in H2;
-                whd in ⊢ (% → ?);
-                [1: intros; cases H6; clear H6; change with (w1 = w);
-                          
-            (* TODO *) ]]    
-|1,2: unfold rest; apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
+             let rc ≝ aux (〈rest,height1〉 :: tl1) tl2 m in
+             〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉]]]]
+in aux : ∀l1,l2,m.∃z.\len l1 + \len l2 ≤ m → rebase_spec_aux l1 l2 z);
+[7: clearbody aux; cases (aux b1 b2 (\len b1 + \len b2)) (w Hw); clear aux;
+    cases (Hw (le_n ?) Hs1 Hs2 (λ_.He1) (λ_.He2)); clear Hw; cases H1; cases H2; cases H3; clear H3 H1 H2;
+    exists [constructor 1;constructor 1;[apply (\fst w)|5:apply (\snd w)]] try assumption;
+    [1,3: apply hide; cases H (X X); try rewrite < (H8 O); try rewrite < X; assumption
+    |2,4: apply hide;[apply H6|apply H7]intro X;[rewrite > X in Hb1|rewrite > X in Hb2]
+         normalize in Hb1 Hb2; [destruct Hb1|destruct Hb2]]
+    unfold; unfold same_values; simplify in ⊢ (? (? % %) ? ?); 
+    simplify in match (\snd 〈?,?〉); simplify in match (\fst 〈?,?〉);
+    split; [assumption; |apply H9;|apply H10]
+|6: intro ABS; unfold; intros 4; clear H1 H2;
+    cases l in ABS H3; intros 1; [2: simplify in H1; cases (not_le_Sn_O ? H1)]
+    cases l1 in H4 H1; intros; [2: simplify in H2; cases (not_le_Sn_O ? H2)]
+    split; [ left; reflexivity|split; apply (sorted_nil q2_lt);|split; assumption;]
+    split; unfold; intros; unfold same_values; intros; reflexivity;
+|5: intros; apply copy_rebases_r;
+|4: intros; rewrite < H1; apply copy_rebases;
+|3: cut (\fst b = \fst b3) as K; [2: apply q_le_to_le_to_eq; assumption] clear H6 H5 H4 H3;
+    intros; cases (aux l2 l3 n1); intros 4; simplify in match (\fst ≪w,H≫);
+    simplify in match (\fst 〈?,?〉); simplify in match (\snd 〈?,?〉);
+    cases H4; 
+      [2: apply le_S_S_to_le; apply (trans_le ???? H3); simplify;
+          rewrite < plus_n_Sm; apply le_S; apply le_n;
+      |3,4: apply (sorted_tail q2_lt); [2: apply H5|4:apply H6]
+      |5: intro; cases l2 in H7 H9; intros; [cases H9; reflexivity]
+          simplify in H7 ⊢ %; apply H7; intro; destruct H10;
+      |6: intro; cases l3 in H8 H9; intros; [cases H9; reflexivity]
+          simplify in H8 ⊢ %; apply H8; intro; destruct H10;]
+    clear aux; split; 
+    [1: left; reflexivity;
+    |2: cases H10;  
+
+    
+     
+ unfold rebase_spec_aux; intros; cases l1 in H2 H4 H6; intros; [ simplify in H2; destruct H2;]
+    lapply H6 as H7; [2: intro X; destruct X] clear H6 H5;
+    rewrite > H7; split; [right; simplify;
+    
+     split; [left;reflexivity]
+    split; 
+
+,2: unfold rest; apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
       assumption;        
-|3:(* TODO *)
-|4:(* TODO *)
-|5:(* TODO *)
-|6:(* TODO *)
-|7:(* TODO *)
-|8: intros; cases (?:False); apply (not_le_Sn_O ? H1);]
-qed.
\ No newline at end of file
+|8: intros; cases (?:False); apply (not_le_Sn_O ? H1);
+|3: intros; generalize in match (unpos ??); intro X; cases X; clear X;
+    simplify in ⊢ (???? (??? (??? (??? (?? (? (?? (??? % ?) ?) ??)))) ?));
+    simplify in ⊢ (???? (???? (??? (??? (?? (? (?? (??? % ?) ?) ??)))))); 
+    clear H4; cases (aux (〈w,\snd b〉::l4) l5 n1); clear aux;
+    cut (len (〈w,\snd b〉::l4) + len l5 < n1) as K;[2:
+      simplify in H5; simplify; rewrite > sym_plus in H5; simplify in H5;
+      rewrite > sym_plus in H5; apply le_S_S_to_le; apply H5;] 
+    split;
+    [1: simplify in ⊢ (? % ?); simplify in ⊢ (? ? %); 
+        cases (H4 s K); clear K H4; intro input; cases input; [reflexivity]
+        simplify; apply H7; 
+    |2: simplify in ⊢ (? ? %); cases (H4 s K); clear H4 K H5 spec;
+        intro;
+        (* input < s + b1 || input >= s + b1 *)
+    |3: simplify in ⊢ (? ? %);]   
+|4: intros; generalize in match (unpos ??); intro X; cases X; clear X;
+    (* duale del 3 *)
+|5: intros; (* triviale, caso in cui non fa nulla *)
+|6,7: (* casi base in cui allunga la lista più corta *) 
+]
+elim devil;
+qed.
+
+include "Q/q/qtimes.ma".
+
+let rec area (l:list bar) on l ≝
+  match l with 
+  [ nil ⇒ OQ
+  | cons he tl ⇒ area tl + Qpos (\fst he) * ⅆ[OQ,\snd he]].
+
+alias symbol "pi1" = "exT \fst".
+alias symbol "minus" = "Q minus".
+alias symbol "exists" = "CProp exists".
+definition minus_spec_bar ≝
+ λf,g,h:list bar.
+   same_bases f g → len f = len g →
+     ∀s,i:ℚ. \snd (\fst (value (mk_q_f s h) i)) = 
+       \snd (\fst (value (mk_q_f s f) i)) - \snd (\fst (value (mk_q_f s g) i)). 
+
+definition minus_spec ≝
+ λf,g:q_f.
+   ∃h:q_f. 
+     ∀i:ℚ. \snd (\fst (value h i)) = 
+       \snd (\fst (value f i)) - \snd (\fst (value g i)). 
+
+definition eject_bar : ∀P:list bar → CProp.(∃l:list bar.P l) → list bar ≝
+ λP.λp.match p with [ex_introT x _ ⇒ x].
+definition inject_bar ≝ ex_introT (list bar).
+
+coercion inject_bar with 0 1 nocomposites.
+coercion eject_bar with 0 0 nocomposites.
+
+lemma minus_q_f : ∀f,g. minus_spec f g.
+intros;
+letin aux ≝ (
+  let rec aux (l1, l2 : list bar) on l1 ≝
+    match l1 with
+    [ nil ⇒ []
+    | cons he1 tl1 ⇒
+        match l2 with
+        [ nil ⇒ []
+        | cons he2 tl2 ⇒ 〈\fst he1, \snd he1 - \snd he2〉 :: aux tl1 tl2]]
+  in aux : ∀l1,l2 : list bar.∃h.minus_spec_bar l1 l2 h);
+[2: intros 4; simplify in H3; destruct H3;
+|3: intros 4; simplify in H3; cases l1 in H2; [2: intro X; simplify in X; destruct X]    
+    intros; rewrite > (value_OQ_e (mk_q_f s []) i); [2: reflexivity]
+    rewrite > q_elim_minus; rewrite > q_plus_OQ; reflexivity;
+|1: cases (aux l2 l3); unfold in H2; intros 4;
+    simplify in ⊢ (? ? (? ? ? (? ? ? (? % ?))) ?);
+    cases (q_cmp i (s + Qpos (\fst b)));
+    
+
+
+definition excess ≝ 
+  λf,g.∃i.\snd (\fst (value f i)) < \snd (\fst (value g i)).
+