]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/models/q_function.ma
shift almost done
[helm.git] / helm / software / matita / contribs / dama / dama / models / q_function.ma
index 48d3012ec66974d8c8993e5db02bd72e3b488229..3275f86f04134d6c1743ca397651ae1457aa923f 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. 
-
-record q_f : Type ≝ {
- start : ℚ;
- bars: list (ℚ × ℚ) (* base, height *)
-}.
-
-axiom qp : ℚ → ℚ → ℚ.
-
-interpretation "Q plus" 'plus x y = (qp x y).
-
-axiom qm : ℚ → ℚ → ℚ.
-
-interpretation "Q minus" 'minus x y = (qm x y).
-
-axiom qlt : ℚ → ℚ → CProp.
-
-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" 'le x y = (qle x y).
-
-notation "'nth'" with precedence 90 for @{'nth}.
-notation < "'nth' \nbsp l \nbsp d \nbsp i" with precedence 71 
-for @{'nth_appl $l $d $i}.
-interpretation "list nth" 'nth = (cic:/matita/list/list/nth.con _).
-interpretation "list nth" 'nth_appl l d i = (cic:/matita/list/list/nth.con _ l d i).
-
-notation < "\rationals \sup 2" non associative with precedence 90 for @{'q2}.
-interpretation "Q x Q" 'q2 = (Prod Q Q).
-
-let rec make_list (A:Type) (def:nat→A) (n:nat) on n ≝
-  match n with
-  [ O ⇒ nil ?
-  | S m ⇒ def m :: make_list A def m].
-
-notation "'mk_list'" with precedence 90 for @{'mk_list}.
-interpretation "'mk_list'" 'mk_list = (make_list _).   
-notation < "'mk_list' \nbsp f \nbsp n" 
-with precedence 71 for @{'mk_list_appl $f $n}.
-interpretation "'mk_list' appl" 'mk_list_appl f n = (make_list _ f n).
-
-definition q0 : ℚ × ℚ ≝ 〈OQ,OQ〉.
-notation < "0 \sub \rationals" with precedence 90 for @{'q0}.
-interpretation "q0" 'q0 = q0.
-
-notation < "[ \rationals \sup 2]" with precedence 90 for @{'lq2}.
-interpretation "lq2" 'lq2 = (list (Prod Q Q)).
-notation < "[ \rationals \sup 2] \sup 2" with precedence 90 for @{'lq22}.
-interpretation "lq22" 'lq22 = (Prod (list (Prod Q Q)) (list (Prod Q Q))).
-
-
-notation "'len'" with precedence 90 for @{'len}.
-interpretation "len" 'len = length.
-notation < "'len' \nbsp l" with precedence 70 for @{'len_appl $l}.
-interpretation "len appl" 'len_appl l = (length _ l).
-
-definition eject ≝
-  λP.λp:∃x:(list (ℚ × ℚ)) × (list (ℚ × ℚ)).P x.match p with [ex_introT p _ ⇒ p].
-coercion cic:/matita/dama/models/q_function/eject.con.
-definition inject ≝
-  λP.λp:(list (ℚ × ℚ)) × (list (ℚ × ℚ)).λh:P p. ex_introT ? P p h.
-(*coercion inject with 0 1 nocomposites.*)
-coercion cic:/matita/dama/models/q_function/inject.con 0 1 nocomposites.
-
-definition cb0h ≝ (λl.mk_list (λi.〈\fst (nth l q0 i),OQ〉) (length ? l)).
-
+include "nat_ordered_set.ma".
+include "models/q_bars.ma".
+
+lemma key:
+  ∀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;]]
+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 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)]
+    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;] 
+    rewrite > (value_OQ_r ?? K); simplify; symmetry; assumption;
+|3: simplify in Hi1; destruct Hi1;
+|4: cases (q_cmp input (start l1));
+    [2: 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; 
+        (* OK *)
+    |1,3: 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,3: rewrite > (value_OQ_e ?? Hb); rewrite > Hv12; rewrite > Hb in Hv11 ⊢ %;
+            simplify in Hv11 ⊢ %; cases (\fst v1) in Hv11; [1,3:intros; reflexivity]
+            cases n; [1,3: intros; reflexivity] intro X; cases (not_le_Sn_O ? (le_S_S_to_le ?? X));
+        |2,4: cases (value_ok l1 input);
+            [1,5: rewrite > Hv12; rewrite > Hb; clear Hv12; simplify;
+                rewrite > H10; rewrite > Hb;
+                cut (O < \fst v1);[2,4: cases (\fst v1) in H9; intros; [2,4: 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);
+                  [apply q_eq_to_le;symmetry|apply q_lt_to_le] assumption;]
+                cases (\fst v1) in H8 H9 Hcut; [1,3:intros (_ _ X); cases (not_le_Sn_O ? X)]
+                intros; clear H13; simplify;
+                rewrite > (key n n1 (b::l)); [1,4: reflexivity] rewrite < Hb;
+                [2,4: 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,4: 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,4: [apply q_eq_to_le; symmetry|apply q_lt_to_le] assumption]
+                      apply q_eq_to_le; reflexivity;
+                |*: apply (q_le_lt_trans ??? H11);
+                    rewrite > q_d_sym; rewrite > q_d_noabs;
+                      [2,4: [apply q_eq_to_le; symmetry|apply q_lt_to_le] assumption]
+                    generalize in match H9; rewrite > q_d_sym; rewrite > q_d_noabs;
+                      [2,4: 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;]
+            |2,6: rewrite > Hb; intro W; destruct W;
+            |3,7: [apply q_eq_to_le;symmetry|apply q_lt_to_le] assumption;
+            |4,8: 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;]]]
+qed.
+
+            
+        
 alias symbol "pi2" = "pair pi2".
 alias symbol "pi1" = "pair pi1".
-definition rebase: 
-  q_f → q_f → 
-    ∃p:q_f × q_f.∀i.
-     \fst (nth (bars (\fst p)) q0 i) = 
-     \fst (nth (bars (\snd p)) q0 i).
-intros (f1 f2); cases f1 (s1 l1); cases f2 (s2 l2); clear f1 f2;
-letin spec ≝ (λl1,l2:list (ℚ × ℚ).λm:nat.λz:(list (ℚ × ℚ)) × (list (ℚ × ℚ)).True);
+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".
 letin aux ≝ ( 
-let rec aux (l1,l2:list (ℚ × ℚ)) (n:nat) on n : (list (ℚ × ℚ)) × (list (ℚ × ℚ)) ≝
+let rec aux (l1,l2:list bar) (n:nat) on n : (list bar) × (list bar) ≝
 match n with
 [ O ⇒ 〈 nil ? , nil ? 〉
 | S m ⇒ 
@@ -113,22 +170,59 @@ match n with
      match l2 with
      [ nil ⇒ 〈l1, cb0h l1〉
      | cons he2 tl2 ⇒  
-         let base1 ≝ (\fst he1) in
-         let base2 ≝ (\fst he2) in
+         let base1 ≝ Qpos (\fst he1) in
+         let base2 ≝ Qpos (\fst he2) in
          let height1 ≝ (\snd he1) in
          let height2 ≝ (\snd he2) in
          match q_cmp base1 base2 with
-         [ q_eq _ ⇒
+         [ q_eq _ ⇒ 
              let rc ≝ aux tl1 tl2 m in 
-             〈he1 :: \fst rc,he2 :: \snd rc〉
-         | q_lt _ ⇒ 
+             〈he1 :: \fst rc,he2 :: \snd rc〉 
+         | q_lt 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 _ ⇒ 
+             let rc ≝ aux tl1 (〈\fst (unpos rest ?),height2〉 :: tl2) m in
+             〈〈\fst he1,height1〉 :: \fst rc,〈\fst he1,height2〉 :: \snd rc〉 
+         | q_gt Hp ⇒ 
              let rest ≝ base1 - base2 in
-             let rc ≝ aux (〈rest,height1〉 :: tl1) tl2 m in
-             〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉
+             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.spec l1 l2 m z); 
-qed.
\ No newline at end of file
+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;
+      assumption;        
+|3:(* TODO *)
+|4:(* TODO *)
+|5:(* TODO *)
+|6:(* TODO *)
+|7:(* TODO *)
+|8: intros; cases (?:False); apply (not_le_Sn_O ? H1);]
+qed.