]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/models/q_function.ma
new specification
[helm.git] / helm / software / matita / contribs / dama / dama / models / q_function.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "nat_ordered_set.ma".
16 include "models/q_bars.ma".
17
18 lemma initial_shift_same_values:
19   ∀l1:q_f.∀init.init < start l1 →
20    same_values l1 
21      (mk_q_f init (〈\fst (unpos (start l1 - init) ?),OQ〉:: bars l1)).  
22 [apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ; assumption]
23 intros; generalize in ⊢ (? ? (? ? (? ? (? ? ? (? ? ? (? ? %)) ?) ?))); intro;
24 cases (unpos (start l1-init) H1); intro input;
25 simplify in ⊢ (? ? ? (? ? ? (? ? ? (? (? ? (? ? (? ? ? % ?) ?)) ?))));
26 cases (value (mk_q_f init (〈w,OQ〉::bars l1)) input) (v1 Hv1); 
27 cases Hv1 (HV1 HV1 HV1 HV1); cases HV1 (Hi1 Hv11 Hv12); clear HV1 Hv1;
28 [1: cut (input < start l1) as K;[2: apply (q_lt_trans ??? Hi1 H)]
29     rewrite > (value_OQ_l ?? K); simplify; symmetry; assumption;
30 |2: cut (start l1 + sum_bases (bars l1) (len (bars l1)) ≤ input) as K;[2: 
31       simplify in Hi1; apply (q_le_trans ???? Hi1); rewrite > H2;
32       rewrite > q_plus_sym in ⊢ (? ? (? ? %));
33       rewrite > q_plus_assoc; rewrite > q_elim_minus;
34       rewrite > q_plus_sym in ⊢ (? ? (? (? ? %) ?));
35       rewrite > q_plus_assoc; rewrite < q_elim_minus;
36       rewrite > q_plus_minus; rewrite > q_plus_sym in ⊢ (? ? (? % ?));
37       rewrite > q_plus_OQ; apply q_eq_to_le; reflexivity;] 
38     rewrite > (value_OQ_r ?? K); simplify; symmetry; assumption;
39 |3: simplify in Hi1; destruct Hi1;
40 |4:
41
42 STOP               
43
44 qed.
45
46             
47         
48 alias symbol "pi2" = "pair pi2".
49 alias symbol "pi1" = "pair pi1".
50 definition rebase_spec ≝ 
51  ∀l1,l2:q_f.∃p:q_f × q_f.
52    And4
53     (*len (bars (\fst p)) = len (bars (\snd p))*)
54     (start (\fst p) = start (\snd p))
55     (same_bases (\fst p) (\snd p))
56     (same_values l1 (\fst p)) 
57     (same_values l2 (\snd p)).
58
59 definition rebase_spec_simpl ≝ 
60  λstart.λl1,l2:list bar.λp:(list bar) × (list bar).
61    And3
62     (same_bases (mk_q_f start (\fst p)) (mk_q_f start (\snd p)))
63     (same_values (mk_q_f start l1) (mk_q_f start (\fst p))) 
64     (same_values (mk_q_f start l2) (mk_q_f start (\snd p))).
65
66 (* a local letin makes russell fail *)
67 definition cb0h : list bar → list bar ≝ 
68   λl.mk_list (λi.〈\fst (nth l ▭ i),OQ〉) (len l).
69
70 definition eject ≝
71   λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
72 coercion eject.
73 definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h.
74 coercion inject with 0 1 nocomposites.
75         
76 definition rebase: rebase_spec.
77 intros 2 (f1 f2); cases f1 (s1 l1); cases f2 (s2 l2); clear f1 f2;
78 letin spec ≝ (
79   λs.λl1,l2.λm.λz.len l1 + len l2 < m → rebase_spec_simpl s l1 l2 z);
80 alias symbol "pi1" (instance 34) = "exT \fst".
81 alias symbol "pi1" (instance 21) = "exT \fst".
82 letin aux ≝ ( 
83 let rec aux (l1,l2:list bar) (n:nat) on n : (list bar) × (list bar) ≝
84 match n with
85 [ O ⇒ 〈 nil ? , nil ? 〉
86 | S m ⇒ 
87   match l1 with
88   [ nil ⇒ 〈cb0h l2, l2〉
89   | cons he1 tl1 ⇒
90      match l2 with
91      [ nil ⇒ 〈l1, cb0h l1〉
92      | cons he2 tl2 ⇒  
93          let base1 ≝ Qpos (\fst he1) in
94          let base2 ≝ Qpos (\fst he2) in
95          let height1 ≝ (\snd he1) in
96          let height2 ≝ (\snd he2) in
97          match q_cmp base1 base2 with
98          [ q_eq _ ⇒ 
99              let rc ≝ aux tl1 tl2 m in 
100              〈he1 :: \fst rc,he2 :: \snd rc〉 
101          | q_lt Hp ⇒
102              let rest ≝ base2 - base1 in
103              let rc ≝ aux tl1 (〈\fst (unpos rest ?),height2〉 :: tl2) m in
104              〈〈\fst he1,height1〉 :: \fst rc,〈\fst he1,height2〉 :: \snd rc〉 
105          | q_gt Hp ⇒ 
106              let rest ≝ base1 - base2 in
107              let rc ≝ aux (〈\fst (unpos rest ?),height1〉 :: tl1) tl2 m in
108              〈〈\fst he2,height1〉 :: \fst rc,〈\fst he2,height2〉 :: \snd rc〉
109 ]]]]
110 in aux : ∀l1,l2,m.∃z.∀s.spec s l1 l2 m z); unfold spec;
111 [9: clearbody aux; unfold spec in aux; clear spec;
112     cases (q_cmp s1 s2);
113     [1: cases (aux l1 l2 (S (len l1 + len l2)));
114         cases (H1 s1 (le_n ?)); clear H1;
115         exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s2 (\snd w)〉] split;
116         [1,2: assumption;
117         |3: intro; apply (H3 input);
118         |4: intro; rewrite > H in H4; 
119             rewrite > (H4 input); reflexivity;]
120     |2: letin l2' ≝ (〈\fst (unpos (s2-s1) ?),OQ〉::l2);[
121           apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
122           assumption]
123         cases (aux l1 l2' (S (len l1 + len l2')));
124         cases (H1 s1 (le_n ?)); clear H1 aux;
125         exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s1 (\snd w)〉] split;
126         [1: reflexivity
127         |2: assumption;
128         |3: assumption;
129         |4: intro; rewrite < (H4 input); clear H3 H4 H2 w;
130             cases (value (mk_q_f s1 l2') input);
131             cases (q_cmp input (start (mk_q_f s1 l2'))) in H1;
132             whd in ⊢ (% → ?);
133             [1: intros; cases H2; clear H2; whd in ⊢ (??? %);
134                 cases (value (mk_q_f s2 l2) input);
135                 cases (q_cmp input (start (mk_q_f s2 l2))) in H2;
136                 whd in ⊢ (% → ?);
137                 [1: intros; cases H6; clear H6; change with (w1 = w);
138                           
139             (* TODO *) ]]    
140 |1,2: unfold rest; apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
141       assumption;        
142 |3:(* TODO *)
143 |4:(* TODO *)
144 |5:(* TODO *)
145 |6:(* TODO *)
146 |7:(* TODO *)
147 |8: intros; cases (?:False); apply (not_le_Sn_O ? H1);]
148 qed.