]> matita.cs.unibo.it Git - helm.git/commitdiff
some more work
authorEnrico Tassi <enrico.tassi@inria.fr>
Wed, 25 Jun 2008 14:54:47 +0000 (14:54 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Wed, 25 Jun 2008 14:54:47 +0000 (14:54 +0000)
helm/software/matita/contribs/dama/dama/models/q_function.ma

index 0d34f2cebd904ca53d4daaeed59676ac0bf41904..c458d418d3c46cf91efcd84b13b1ce2e3604c22e 100644 (file)
@@ -20,21 +20,15 @@ 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 *)
-}.
+definition bar ≝ ratio × ℚ. (* base (Qpos) , height *)
+record q_f : Type ≝ { start : ℚ; bars: list bar }.
 
 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 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 ≝
@@ -46,63 +40,175 @@ 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).
+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 dist : ℚ → ℚ → ℚ.
+
+
+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 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 < "'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).
 
-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].
+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.
 
-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).
+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}.
 
-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))).
+definition empty_bar : bar ≝ 〈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).
 
 notation "'len'" with precedence 90 for @{'len}.
-interpretation "len" 'len = length.
-notation < "'len' \nbsp l" with precedence 70 for @{'len_appl $l}.
+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 (ℚ × ℚ)) × (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.
-
 (* a local letin makes russell fail *)
-definition cb0h ≝ (λl.mk_list (λi.〈\fst (nth l q0 i),OQ〉) (length ? l)).
+definition cb0h ≝ (λl.mk_list (λi.〈\fst (nth l ▭ i),OQ〉) (len l)).
+
+lemma len_mk_list : ∀T:Type.∀f:nat→T.∀n.len (mk_list f n) = n.
+intros; elim n; [reflexivity] simplify; rewrite > H; reflexivity;
+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)]].
 
+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".
-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);
+letin value ≝ (
+  let rec value (p: ℚ) (l : list bar) on l ≝
+    match l with
+    [ nil ⇒ 〈O,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; 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; clear H2;
+          simplify; apply q_le_minus; assumption;
+    |2,5: cases (value (q-Qpos (\fst b)) l1); cases H3; 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 H3; clear H3 value H2;
+        assumption;]
+|2: clear value H2; simplify; split;
+    [1: 
+    
+          
+definition same_shape ≝
+  λl1,l2:q_f.
+   ∀input.∃col.
+   
+    And3 
+      (sum_bases (bars l2) j ≤ offset - start l2)
+      (offset - start l2 ≤ sum_bases (bars l2) (S j))
+      (\snd (nth (bars l2)) q0 j) = \snd (nth (bars l1) q0 i).
+
+…┐─┌┐…
+\ldots\boxdl\boxh\boxdr\boxdl\ldots
+
+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))
+    (∀i.\fst (nth (bars (\fst p)) q0 i) = \fst (nth (bars (\snd p)) q0 i))
+    (∀i,offset.
+      sum_bases (bars l1) i ≤ offset - start l1 →
+      offset - start l1 ≤ sum_bases (bars l1) (S i) →
+      ∃j.      
+        And3 
+         (sum_bases (bars (\fst p)) j ≤ offset - start (\fst p))
+         (offset - start (\fst p) ≤ sum_bases (bars (\fst p)) (S j))
+         (\snd (nth (bars (\fst p)) q0 j) = \snd (nth (bars l1) q0 i)) ∧
+        And3 
+         (sum_bases (bars (\snd p)) j ≤ offset - start (\snd p))
+         (offset - start (\snd p) ≤ sum_bases (bars (\snd p)) (S j))
+         (\snd (nth (bars (\snd p)) q0 j) = \snd (nth (bars l2) q0 i))).
+
+definition rebase_spec_simpl ≝ 
+ λl1,l2:list (ℚ × ℚ).λp:(list (ℚ × ℚ)) × (list (ℚ × ℚ)).
+    len ( (\fst p)) = len ( (\snd p)) ∧
+    (∀i.
+     \fst (nth ( (\fst p)) q0 i) = 
+     \fst (nth ( (\snd p)) q0 i)) ∧
+    ∀i,offset.
+      sum_bases ( l1) i ≤ offset ∧
+      offset ≤ sum_bases ( l1) (S i)
+       →
+      ∃j.  
+        sum_bases ( (\fst p)) j ≤ offset ∧
+        offset ≤ sum_bases ((\fst p)) (S j) ∧
+        \snd (nth ( (\fst p)) q0 j) = 
+        \snd (nth ( l1) q0 i).
+
+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 ≝ (
+  λl1,l2:list (ℚ × ℚ).λm:nat.λz:(list (ℚ × ℚ)) × (list (ℚ × ℚ)).
+    len l1 + len l2 < m → rebase_spec_simpl l1 l2 z);
 letin aux ≝ ( 
 let rec aux (l1,l2:list (ℚ × ℚ)) (n:nat) on n : (list (ℚ × ℚ)) × (list (ℚ × ℚ)) ≝
 match n with
@@ -119,17 +225,38 @@ match n with
          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 _ ⇒
              let rest ≝ base2 - base1 in
              let rc ≝ aux tl1 (〈rest,height2〉 :: tl2) m in
-             〈〈base1,height1〉 :: \fst rc,〈base1,height2〉 :: \snd rc〉 *)
-         | q_gt _ ⇒ 〈[],[]〉 (*
+             〈〈base1,height1〉 :: \fst rc,〈base1,height2〉 :: \snd rc〉 
+         | q_gt _ ⇒ 
              let rest ≝ base1 - base2 in
              let rc ≝ aux (〈rest,height1〉 :: tl1) tl2 m in
-             〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉 *)
+             〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉
 ]]]]
-in aux : ∀l1,l2,m.∃z.spec l1 l2 m z); 
+in aux : ∀l1,l2,m.∃z.spec l1 l2 m z); unfold spec;
+[7: 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 (le_n ?)); clear H1;
+        exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s2 (\snd w)〉] repeat split;
+        [1: cases H2; assumption;
+        |2: assumption;
+        |3: cases H2; assumption;
+        |4: intros; cases (H3 i (offset - s1)); 
+            [2: 
+        
+        
+|1,2: simplify; generalize in ⊢ (? ? (? (? ? (? ? ? (? ? %)))) (? (? ? (? ? ? (? ? %))))); intro X; 
+      cases X (rc OK); clear X; simplify; apply eq_f; assumption;
+|3: cases (aux l4 l5 n1) (rc OK); simplify; apply eq_f; assumption;
+|4,5: simplify; unfold cb0h; rewrite > len_mk_list; reflexivity;
+|6: reflexivity]
+clearbody aux; unfold spec in aux; clear spec; 
+
+
+
 qed.
\ No newline at end of file