]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/models/q_function.ma
simplified coercDb implementation with additional info about the position of
[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_shift.ma".
17
18 alias symbol "pi2" = "pair pi2".
19 alias symbol "pi1" = "pair pi1".
20 definition rebase_spec ≝ 
21  ∀l1,l2:q_f.∃p:q_f × q_f.
22    And4
23     (start (\fst p) = start (\snd p))
24     (same_bases (bars (\fst p)) (bars (\snd p)))
25     (same_values l1 (\fst p)) 
26     (same_values l2 (\snd p)).
27
28 definition rebase_spec_simpl ≝ 
29  λstart.λl1,l2:list bar.λp:(list bar) × (list bar).
30    And3
31     (same_bases (\fst p) (\snd p))
32     (same_values (mk_q_f start l1) (mk_q_f start (\fst p))) 
33     (same_values (mk_q_f start l2) (mk_q_f start (\snd p))).
34
35 (* a local letin makes russell fail *)
36 definition cb0h : list bar → list bar ≝ 
37   λl.mk_list (λi.〈\fst (nth l ▭ i),OQ〉) (len l).
38
39 definition eject ≝
40   λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
41 coercion eject.
42 definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h.
43 coercion inject with 0 1 nocomposites.
44
45 axiom devil : False.
46         
47 definition rebase: rebase_spec.
48 intros 2 (f1 f2); cases f1 (s1 l1); cases f2 (s2 l2); clear f1 f2;
49 letin spec ≝ (
50   λs.λl1,l2.λm.λz.len l1 + len l2 < m → rebase_spec_simpl s l1 l2 z);
51 alias symbol "pi1" (instance 34) = "exT \fst".
52 alias symbol "pi1" (instance 21) = "exT \fst".
53 letin aux ≝ ( 
54 let rec aux (l1,l2:list bar) (n:nat) on n : (list bar) × (list bar) ≝
55 match n with
56 [ O ⇒ 〈 nil ? , nil ? 〉
57 | S m ⇒ 
58   match l1 with
59   [ nil ⇒ 〈cb0h l2, l2〉
60   | cons he1 tl1 ⇒
61      match l2 with
62      [ nil ⇒ 〈l1, cb0h l1〉
63      | cons he2 tl2 ⇒  
64          let base1 ≝ Qpos (\fst he1) in
65          let base2 ≝ Qpos (\fst he2) in
66          let height1 ≝ (\snd he1) in
67          let height2 ≝ (\snd he2) in
68          match q_cmp base1 base2 with
69          [ q_eq _ ⇒ 
70              let rc ≝ aux tl1 tl2 m in 
71              〈he1 :: \fst rc,he2 :: \snd rc〉 
72          | q_lt Hp ⇒
73              let rest ≝ base2 - base1 in
74              let rc ≝ aux tl1 (〈\fst (unpos rest ?),height2〉 :: tl2) m in
75              〈〈\fst he1,height1〉 :: \fst rc,〈\fst he1,height2〉 :: \snd rc〉 
76          | q_gt Hp ⇒ 
77              let rest ≝ base1 - base2 in
78              let rc ≝ aux (〈\fst (unpos rest ?),height1〉 :: tl1) tl2 m in
79              〈〈\fst he2,height1〉 :: \fst rc,〈\fst he2,height2〉 :: \snd rc〉
80 ]]]]
81 in aux : ∀l1,l2,m.∃z.∀s.spec s l1 l2 m z); unfold spec;
82 [9: clearbody aux; unfold spec in aux; clear spec;
83     cases (q_cmp s1 s2);
84     [1: cases (aux l1 l2 (S (len l1 + len l2)));
85         cases (H1 s1 (le_n ?)); clear H1;
86         exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s2 (\snd w)〉] split;
87         [1,2: assumption;
88         |3: intro; apply (H3 input);
89         |4: intro; rewrite > H in H4; 
90             rewrite > (H4 input) in ⊢ (? ? % ?); reflexivity;]
91     |2: letin l2' ≝ (〈\fst (unpos (s2-s1) ?),OQ〉::l2);[
92           apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
93           assumption]
94         cases (aux l1 l2' (S (len l1 + len l2')));
95         cases (H1 s1 (le_n ?)); clear H1 aux;
96         exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s1 (\snd w)〉] split;
97         [1: reflexivity
98         |2: assumption;
99         |3: assumption;
100         |4: intro; 
101             rewrite > (initial_shift_same_values (mk_q_f s2 l2) s1 H input) in ⊢ (? ? % ?);
102             rewrite < (H4 input)in ⊢ (? ? ? %); reflexivity;]
103     |3: letin l1' ≝ (〈\fst (unpos (s1-s2) ?),OQ〉::l1);[
104           apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
105           assumption]
106         cases (aux l1' l2 (S (len l1' + len l2)));
107         cases (H1 s2 (le_n ?)); clear H1 aux;
108         exists [apply 〈mk_q_f s2 (\fst w), mk_q_f s2 (\snd w)〉] split;
109         [1: reflexivity
110         |2: assumption;
111         |4: assumption;
112         |3: intro; simplify in ⊢ (? ? ? (? ? ? (? ? ? (? % ?))));
113             rewrite > (initial_shift_same_values (mk_q_f s1 l1) s2 H input) in ⊢ (? ? % ?);
114             rewrite < (H3 input) in ⊢ (? ? ? %); reflexivity;]]
115 |1,2: unfold rest; apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
116       assumption;        
117 |8: intros; cases (?:False); apply (not_le_Sn_O ? H1);
118 |3: intros; generalize in match (unpos ??); intro X; cases X; clear X;
119     simplify in ⊢ (???? (??? (??? (??? (?? (? (?? (??? % ?) ?) ??)))) ?));
120     simplify in ⊢ (???? (???? (??? (??? (?? (? (?? (??? % ?) ?) ??)))))); 
121     clear H4; cases (aux (〈w,\snd b〉::l4) l5 n1); clear aux;
122     cut (len (〈w,\snd b〉::l4) + len l5 < n1) as K;[2:
123       simplify in H5; simplify; rewrite > sym_plus in H5; simplify in H5;
124       rewrite > sym_plus in H5; apply le_S_S_to_le; apply H5;] 
125     split;
126     [1: simplify in ⊢ (? % ?); simplify in ⊢ (? ? %); 
127         cases (H4 s K); clear K H4; intro input; cases input; [reflexivity]
128         simplify; apply H7; 
129     |2: simplify in ⊢ (? ? %); cases (H4 s K); clear H4 K H5 spec;
130         intro;
131         (* input < s + b1 || input >= s + b1 *)
132     |3: simplify in ⊢ (? ? %);]   
133 |4: intros; generalize in match (unpos ??); intro X; cases X; clear X;
134     (* duale del 3 *)
135 |5: intros; (* triviale, caso in cui non fa nulla *)
136 |6,7: (* casi base in cui allunga la lista più corta *) 
137 ]
138 elim devil;
139 qed.
140
141 include "Q/q/qtimes.ma".
142
143 let rec area (l:list bar) on l ≝
144   match l with 
145   [ nil ⇒ OQ
146   | cons he tl ⇒ area tl + Qpos (\fst he) * ⅆ[OQ,\snd he]].
147
148 alias symbol "pi1" = "exT \fst".
149 alias symbol "minus" = "Q minus".
150 alias symbol "exists" = "CProp exists".
151 definition minus_spec_bar ≝
152  λf,g,h:list bar.
153    same_bases f g → len f = len g →
154      ∀s,i:ℚ. \snd (\fst (value (mk_q_f s h) i)) = 
155        \snd (\fst (value (mk_q_f s f) i)) - \snd (\fst (value (mk_q_f s g) i)). 
156
157 definition minus_spec ≝
158  λf,g:q_f.
159    ∃h:q_f. 
160      ∀i:ℚ. \snd (\fst (value h i)) = 
161        \snd (\fst (value f i)) - \snd (\fst (value g i)). 
162
163 definition eject_bar : ∀P:list bar → CProp.(∃l:list bar.P l) → list bar ≝
164  λP.λp.match p with [ex_introT x _ ⇒ x].
165 definition inject_bar ≝ ex_introT (list bar).
166
167 coercion inject_bar with 0 1 nocomposites.
168 coercion eject_bar with 0 0 nocomposites.
169
170 lemma minus_q_f : ∀f,g. minus_spec f g.
171 intros;
172 letin aux ≝ (
173   let rec aux (l1, l2 : list bar) on l1 ≝
174     match l1 with
175     [ nil ⇒ []
176     | cons he1 tl1 ⇒
177         match l2 with
178         [ nil ⇒ []
179         | cons he2 tl2 ⇒ 〈\fst he1, \snd he1 - \snd he2〉 :: aux tl1 tl2]]
180   in aux : ∀l1,l2 : list bar.∃h.minus_spec_bar l1 l2 h);
181 [2: intros 4; simplify in H3; destruct H3;
182 |3: intros 4; simplify in H3; cases l1 in H2; [2: intro X; simplify in X; destruct X]    
183     intros; rewrite > (value_OQ_e (mk_q_f s []) i); [2: reflexivity]
184     rewrite > q_elim_minus; rewrite > q_plus_OQ; reflexivity;
185 |1: cases (aux l2 l3); unfold in H2; intros 4;
186     simplify in ⊢ (? ? (? ? ? (? ? ? (? % ?))) ?);
187     cases (q_cmp i (s + Qpos (\fst b)));
188     
189
190
191 definition excess ≝ 
192   λf,g.∃i.\snd (\fst (value f i)) < \snd (\fst (value g i)).
193