]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/dama/dama/models/q_function.ma
more work
[helm.git] / helm / software / matita / contribs / dama / dama / models / q_function.ma
index df368bf9065b169e76bf598163fcdf6201b7e39d..32122371004982cad24e4deccf33ec5f71dad0e1 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'" left associative with precedence 70 for @{'nth}.
-notation < "\nth \nbsp l \nbsp d \nbsp i" left associative with precedence 70 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 40 for @{'q2}.
-interpretation "Q x Q" 'q2 = (product Q Q).
-
-let rec mk_list (A:Type) (def:nat→A) (n:nat) on n ≝
-  match n with
-  [ O ⇒ []
-  | S m ⇒ def m :: mk_list A def m].
-
-interpretation "mk_list appl" 'mk_list f n = (mk_list f n).
-interpretation "mk_list" 'mk_list = mk_list.   
-notation < "\mk_list \nbsp f \nbsp n" left associative with precedence 70 for @{'mk_list_appl $f $n}.
-notation "'mk_list'" left associative with precedence 70 for @{'mk_list}.
-
-alias symbol "pair" = "pair".
-definition q00 : ℚ × ℚ ≝ 〈OQ,OQ〉.
-
+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 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; cases (value l1 input); simplify in ⊢ (? ? (? ? ? %) ?);
+    cases (q_cmp input (start l1)) in H5; whd in ⊢ (% → ?);
+    [1: cases (?:False); clear w2 H4 w1 H2 w H1; 
+        apply (q_lt_antisym init (start l1)); [assumption] rewrite < H5; assumption
+    |2: intros; rewrite > H6; clear H6; rewrite > H4; reflexivity;
+    |3: cases (?:False); apply (q_lt_antisym input (start l1)); [2: assumption]
+        apply (q_lt_trans ??? H3 H);]
+|3: intro; cases H4; clear H4;   
+    cases (value l1 input); simplify; cases (q_cmp input (start l1)) in H4; whd in ⊢ (% → ?);
+    [1: intro; cases H8; clear H8; rewrite > H11; rewrite > H7; clear H11 H7;
+        simplify in ⊢ (? ? ? (? ? ? (? ? % ? ?)));
+        cut (\fst w1 = S (\fst w2)) as Key; [rewrite > Key; reflexivity;]
+        cut (\fst w2 = O); [2: clear H10;
+          symmetry; apply le_n_O_to_eq; rewrite > (sum_bases_O l1 (\fst w2)); [apply le_n]
+          apply (q_le_trans ??? H9); rewrite < H4; rewrite > q_d_x_x; 
+          apply q_eq_to_le; reflexivity;]
+        rewrite > Hcut; clear Hcut H10 H9; simplify in H5 H6;
+        cut (ⅆ[input,init] = Qpos w) as E; [2:
+          rewrite > H2; rewrite < H4; rewrite > q_d_sym; 
+          rewrite > q_d_noabs; [reflexivity] apply q_lt_to_le; assumption;]
+        cases (\fst w1) in H5 H6; intros;
+        [1: cases (?:False); clear H5; simplify in H6;
+            apply (q_lt_corefl ⅆ[input,init]);
+            rewrite > E in ⊢ (??%); rewrite < q_plus_OQ in ⊢ (??%);
+            rewrite > q_plus_sym; assumption;
+        |2: cases n in H5 H6; [intros; reflexivity] intros;
+            cases (?:False); clear H6; cases (bars l1) in H5; simplify; intros;
+            [apply (q_pos_OQ one);|apply (q_pos_OQ (\fst b));] 
+            apply (q_le_S ??? (sum_bases_ge_OQ ? n1));[apply []|3:apply l]
+            simplify in ⊢ (? (? (? % ?) ?) ?); rewrite < (q_plus_minus (Qpos w));
+            rewrite > q_elim_minus; apply q_le_minus_r; 
+            rewrite > q_elim_opp; rewrite < E in ⊢ (??%); assumption;]
+    |2: intros; rewrite > H8; rewrite > H7; clear H8 H7;
+        simplify in H5 H6 ⊢ %; 
+        cases (\fst w1) in H5 H6; [intros; reflexivity]
+        cases (bars l1);
+        [1: intros; simplify; elim n [reflexivity] simplify; assumption;
+        |2: simplify; intros; cases (?:False); clear H6;
+            apply (q_lt_le_incompat (input - init) (Qpos w) );
+            [1: rewrite > H2; do 2 rewrite > q_elim_minus;
+                apply q_lt_plus; rewrite > q_elim_minus;
+                rewrite < q_plus_assoc; rewrite < q_elim_minus;
+                rewrite > q_plus_minus;
+                rewrite > q_plus_OQ; assumption;
+            |2: rewrite < q_d_noabs; [2: apply q_lt_to_le; assumption]
+                rewrite > q_d_sym; apply (q_le_S ???? H5);
+                apply sum_bases_ge_OQ;]]
+    |3:
+    
+        
+STOP            
+        
 alias symbol "pi2" = "pair pi2".
 alias symbol "pi1" = "pair pi1".
-alias symbol "pair" = "pair".
-definition rebase: 
-  q_f → q_f → 
-    ∃p:q_f × q_f.∀i.
-     fst (nth (bars (fst p)) q00 i) = 
-     fst (nth (bars (snd p)) q00 i).   
-intros (f1 f2); cases f1 (s1 l1); cases f2 (s2 l2); clear f1 f2;
-letin aux ≝ (
-let rec aux (l1,l2:list (ℚ × ℚ)) on l1 : (list (ℚ × ℚ)) × (list (ℚ × ℚ)) ≝
+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 bar) (n:nat) on n : (list bar) × (list bar) ≝
+match n with
+[ O ⇒ 〈 nil ? , nil ? 〉
+| S m ⇒ 
   match l1 with
-  [ nil ⇒ 〈mk_list (λi.〈fst (nth l2 q00 i),OQ〉) (length ? l2),l2〉
-  | cons he tl ⇒ 〈[],[]〉] in aux); 
-  
-cases (q_cmp s1 s2);
-[1: apply (mk_q_f s1);
-|2: apply (mk_q_f s1); cases l2;
-    [1: letin l2' ≝ (
-[1: (* offset: the smallest one *)
-    cases 
+  [ nil ⇒ 〈cb0h l2, l2〉
+  | cons he1 tl1 ⇒
+     match l2 with
+     [ nil ⇒ 〈l1, cb0h 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
+         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_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;
+      assumption;        
+|3:(* TODO *)
+|4:(* TODO *)
+|5:(* TODO *)
+|6:(* TODO *)
+|7:(* TODO *)
+|8: intros; cases (?:False); apply (not_le_Sn_O ? H1);]
+qed.