]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/models/q_rebase.ma
Preparing for 0.5.9 release.
[helm.git] / helm / software / matita / contribs / dama / dama / models / q_rebase.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 "dama/russell_support.ma".
16 include "models/q_copy.ma".
17 (*
18 definition rebase_spec ≝ 
19  λl1,l2:q_f.λp:q_f × q_f. 
20    And3
21     (same_bases (bars (\fst p)) (bars (\snd p)))
22     (same_values l1 (\fst p)) 
23     (same_values l2 (\snd p)).
24
25 inductive rebase_cases : list bar → list bar → (list bar) × (list bar) → Prop ≝
26 | rb_fst_full  : ∀b,h1,xs.
27    rebase_cases (〈b,h1〉::xs) [] 〈〈b,h1〉::xs,〈b,〈OQ,OQ〉〉::copy xs〉
28 | rb_snd_full  : ∀b,h1,ys.
29    rebase_cases [] (〈b,h1〉::ys) 〈〈b,〈OQ,OQ〉〉::copy ys,〈b,h1〉::ys〉  
30 | rb_all_full  : ∀b,h1,h2,h3,h4,xs,ys,r1,r2.
31    \snd(\last ▭ (〈b,h1〉::xs)) = \snd(\last ▭ (〈b,h3〉::r1)) →
32    \snd(\last ▭ (〈b,h2〉::ys)) = \snd(\last ▭ (〈b,h4〉::r2)) →
33    rebase_cases (〈b,h1〉::xs) (〈b,h2〉::ys) 〈〈b,h3〉::r1,〈b,h4〉::r2〉
34 | rb_all_full_l  : ∀b1,b2,h1,h2,xs,ys,r1,r2.
35    \snd(\last ▭ (〈b1,h1〉::xs)) = \snd(\last ▭ (〈b1,h1〉::r1)) →
36    \snd(\last ▭ (〈b2,h2〉::ys)) = \snd(\last ▭ (〈b1,h2〉::r2)) →
37    b1 < b2 →
38    rebase_cases (〈b1,h1〉::xs) (〈b2,h2〉::ys) 〈〈b1,h1〉::r1,〈b1,〈OQ,OQ〉〉::r2〉 
39 | rb_all_full_r  : ∀b1,b2,h1,h2,xs,ys,r1,r2.
40    \snd(\last ▭ (〈b1,h1〉::xs)) = \snd(\last ▭ (〈b2,h1〉::r1)) →
41    \snd(\last ▭ (〈b2,h2〉::ys)) = \snd(\last ▭ (〈b2,h2〉::r2)) →
42    b2 < b1 →
43    rebase_cases (〈b1,h1〉::xs) (〈b2,h2〉::ys) 〈〈b2,〈OQ,OQ〉〉::r1,〈b2,h2〉::r2〉     
44 | rb_all_empty : rebase_cases [] [] 〈[],[]〉.
45
46 alias symbol "pi2" = "pair pi2".
47 alias symbol "pi1" = "pair pi1".
48 alias symbol "leq" = "natural 'less or equal to'".
49 inductive rebase_spec_aux_p (l1, l2:list bar) (p:(list bar) × (list bar)) : Prop ≝
50 | prove_rebase_spec_aux:
51    rebase_cases l1 l2 p →
52    (sorted q2_lt (\fst p)) →
53    (sorted q2_lt (\snd p)) → 
54    (same_bases (\fst p) (\snd p)) →
55    (same_values_simpl l1 (\fst p)) → 
56    (same_values_simpl l2 (\snd p)) →  
57    rebase_spec_aux_p l1 l2 p.   
58
59 lemma aux_preserves_sorting:
60  ∀b,b3,l2,l3,w. rebase_cases l2 l3 w → 
61   sorted q2_lt (b::l2) → sorted q2_lt (b3::l3) → \fst b3 = \fst b →
62   sorted q2_lt (\fst w) → sorted q2_lt (\snd w) → 
63   same_bases (\fst w) (\snd w) → 
64     sorted q2_lt (b :: \fst w).
65 intros 6; cases H; simplify; intros; clear H;
66 [ apply (sorted_cons q2_lt); [2:assumption] apply (inversion_sorted2 ??? H1);
67 | apply (sorted_cons q2_lt); [2:assumption]
68   whd; rewrite < H3; apply (inversion_sorted2 ??? H2);
69 | apply (sorted_cons q2_lt); [2:assumption] apply (inversion_sorted2 ??? H3);
70 | apply (sorted_cons q2_lt); [2:assumption] apply (inversion_sorted2 ??? H4);
71 | apply (sorted_cons q2_lt); [2:assumption] 
72   whd; rewrite < H6; apply (inversion_sorted2 ??? H5);
73 | apply (sorted_one q2_lt);]
74 qed.   
75
76 lemma aux_preserves_sorting2:
77  ∀b,b3,l2,l3,w. rebase_cases l2 l3 w → 
78   sorted q2_lt (b::l2) → sorted q2_lt (b3::l3) → \fst b3 = \fst b →
79   sorted q2_lt (\fst w) → sorted q2_lt (\snd w) → same_bases (\fst w) (\snd w) → 
80   sorted q2_lt (b :: \snd w).
81 intros 6; cases H; simplify; intros; clear H;
82 [ apply (sorted_cons q2_lt); [2:assumption] apply (inversion_sorted2 ??? H1);   
83 | apply (sorted_cons q2_lt); [2:assumption] 
84   whd; rewrite < H3; apply (inversion_sorted2 ??? H2);
85 | apply (sorted_cons q2_lt); [2: assumption] apply (inversion_sorted2 ??? H3);
86 | apply (sorted_cons q2_lt); [2: assumption] apply (inversion_sorted2 ??? H4);
87 | apply (sorted_cons q2_lt); [2: assumption] 
88   whd; rewrite < H6; apply (inversion_sorted2 ??? H5);
89 | apply (sorted_one q2_lt);]
90 qed.
91 *)
92
93
94
95 definition rebase_spec_aux ≝ 
96  λl1,l2
97  :list bar.λp:(list bar) × (list bar).
98  sorted q2_lt l1 → (\snd (\last ▭ l1) = 〈OQ,OQ〉) →
99  sorted q2_lt l2 → (\snd (\last ▭ l2) = 〈OQ,OQ〉) → 
100    rebase_spec_aux_p l1 l2 p.
101
102 alias symbol "lt" = "Q less than".
103 alias symbol "Q" = "Rationals".
104 axiom q_unlimited: ∀x:ℚ.∃y:ratio.x<Qpos y.                
105 axiom q_halving: ∀x,y:ℚ.∃z:ℚ.x<z ∧ z<y.
106 alias symbol "not" = "logical not".
107 axiom q_not_OQ_lt_Qneg: ∀r. ¬ (OQ < Qneg r).
108 lemma same_values_unit_OQ: 
109   ∀b1,b2,h1,l. OQ < b2 → b2 < b1 → sorted q2_lt (〈b1,h1〉::l) →
110     sorted q2_lt  [〈b2,〈OQ,OQ〉〉] → 
111     same_values_simpl (〈b1,h1〉::l) [〈b2,〈OQ,OQ〉〉]  → h1 = 〈OQ,OQ〉.
112 intros 5 (b1 b2 h1 l POS); cases l;
113 [1: intros; cases (q_unlimited b1); cut (b2 < Qpos w); [2:apply (q_lt_trans ??? H H4);] 
114     lapply (H3 H1 ? H2 ? w H4 Hcut) as K; simplify; [1,2: autobatch]
115     rewrite > (value_unit 〈b1,h1〉) in K;
116     rewrite > (value_unit 〈b2,〈OQ,OQ〉〉) in K; assumption;
117 |2: intros; unfold in H3; lapply depth=0 (H3 H1 ? H2 ?) as K; [1,2:simplify; autobatch]
118     clear H3; cases (q_halving b1 (b1 + \fst p)) (w Pw); cases w in Pw; intros;
119     [cases (q_lt_le_incompat ?? POS); apply q_lt_to_le; cases H3;
120      apply (q_lt_trans ???? H4); assumption;
121     |3: elim H3; lapply (q_lt_trans ??? H H4); lapply (q_lt_trans ??? POS Hletin);
122         cases (q_not_OQ_lt_Qneg ? Hletin1);
123     | cases H3; lapply (K r);
124       [2: simplify; assumption
125       |3: simplify; apply (q_lt_trans ???? H4); assumption;
126       |rewrite > (value_head b1,h1 .. q) in Hletin;
127         
128
129
130  (* MANCA che le basi sono positive, 
131                poi con halving prendi tra b1 e \fst p e hai h1=OQ,OQ*)
132
133
134 definition eject ≝
135   λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
136 coercion eject.
137 definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h.
138 coercion inject with 0 1 nocomposites.
139
140 definition rebase: ∀l1,l2:q_f.∃p:q_f × q_f.rebase_spec l1 l2 p.
141 intros 2 (f1 f2); cases f1 (b1 Hs1 Hb1 He1); cases f2 (b2 Hs2 Hb2 He2); clear f1 f2;
142 alias symbol "leq" = "natural 'less or equal to'".
143 alias symbol "minus" = "Q minus".
144 letin aux ≝ ( 
145 let rec aux (l1,l2:list bar) (n : nat) on n : (list bar) × (list bar) ≝
146 match n with
147 [ O ⇒ 〈[], []〉
148 | S m ⇒
149   match l1 with
150   [ nil ⇒ 〈copy l2, l2〉
151   | cons he1 tl1 ⇒
152      match l2 with
153      [ nil ⇒ 〈l1, copy l1〉
154      | cons he2 tl2 ⇒  
155          let base1 ≝ \fst he1 in
156          let base2 ≝ \fst he2 in
157          let height1 ≝ \snd he1 in
158          let height2 ≝ \snd he2 in
159          match q_cmp base1 base2 with
160          [ q_leq Hp1 ⇒ 
161              match q_cmp base2 base1 with
162              [ q_leq Hp2 ⇒
163                  let rc ≝ aux tl1 tl2 m in 
164                  〈he1 :: \fst rc,he2 :: \snd rc〉
165              | q_gt Hp ⇒ 
166                  let rest ≝ base2 - base1 in
167                  let rc ≝ aux tl1 (〈rest,height2〉 :: tl2) m in
168                  〈〈base1,height1〉 :: \fst rc,〈base1,height2〉 :: \snd rc〉] 
169          | q_gt Hp ⇒ 
170              let rest ≝ base1 - base2 in
171              let rc ≝ aux (〈rest,height1〉 :: tl1) tl2 m in
172              〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉]]]]
173 in aux : ∀l1,l2,m.∃z.\len l1 + \len l2 ≤ m → rebase_spec_aux l1 l2 z);
174 [7: clearbody aux; cases (aux b1 b2 (\len b1 + \len b2)) (res Hres);
175     exists; [split; constructor 1; [apply (\fst res)|5:apply (\snd res)]]
176     [1,4: apply hide; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); assumption;
177     |2,5: apply hide; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); clear Hres aux;
178           lapply (H3 O) as K; clear H1 H2 H3 H4 H5; unfold nth_base; 
179           cases H in K He1 He2 Hb1 Hb2; simplify; intros; assumption;
180     |3,6: apply hide; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); clear Hres aux;
181           cases H in He1 He2; simplify; intros;
182           [1,6,8,12: assumption;
183           |2,7: rewrite > len_copy; generalize in match (\len ?); intro X;
184                 cases X; [1,3: reflexivity] simplify;
185                 [apply (copy_OQ ys n);|apply (copy_OQ xs n);]
186           |3,4: rewrite < H6; assumption;
187           |5: cases r1 in H6; simplify; intros; [reflexivity] rewrite < H6; assumption;
188           |9,11: rewrite < H7; assumption;
189           |10: cases r2 in H7; simplify; intros; [reflexivity] rewrite < H7; assumption]]
190     split; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); clear Hres; unfold same_values; intros; 
191     [1: assumption
192     |2,3: simplify in match (\snd 〈?,?〉); simplify in match (\fst 〈?,?〉);
193         apply same_values_simpl_to_same_values; assumption]
194 |3: cut (\fst b3 = \fst b) as E; [2: apply q_le_to_le_to_eq; assumption]
195     clear H6 H5 H4 H3 He2 Hb2 Hs2 b2 He1 Hb1 Hs1 b1; cases (aux l2 l3 n1) (rc Hrc); 
196     clear aux; intro K; simplify in K; rewrite <plus_n_Sm in K; 
197     lapply le_S_S_to_le to K as W; lapply lt_to_le to W as R; 
198     simplify in match (? ≪rc,Hrc≫); intros (Hsbl2 Hendbl2 Hsb3l3 Hendb3l3);
199     change in Hendbl2 with (\snd (\last ▭ (b::l2)) = 〈OQ,OQ〉);
200     change in Hendb3l3 with (\snd (\last ▭ (b3::l3)) = 〈OQ,OQ〉);
201     cases (Hrc R) (RC S1 S2 SB SV1 SV2); clear Hrc R W K; 
202       [2,4: apply (sorted_tail q2_lt);[apply b|3:apply b3]assumption;
203       |3: cases l2 in Hendbl2; simplify; intros; [reflexivity] assumption; 
204       |5: cases l3 in Hendb3l3; simplify; intros; [reflexivity] assumption;]
205     constructor 1; simplify in match (\fst 〈?,?〉); simplify in match (\snd 〈?,?〉);
206     [1: cases b in E Hsbl2 Hendbl2; cases b3 in Hsb3l3 Hendb3l3; intros (Hsbl3 Hendbl2 E Hsb3l2 Hendb3l3); 
207         simplify in E; destruct E; constructor 3;
208         [1: clear Hendbl2 Hsbl3 SV2 SB S2;
209             cases RC in S1 SV1 Hsb3l2 Hendb3l3; intros; 
210             [1,6: reflexivity;
211             |3,4: assumption;
212             |5: simplify in H6:(??%) ⊢ %; rewrite > H3; cases r1 in H6; intros [2:reflexivity]
213                 use same_values_unit_OQ; 
214
215             |2: simplify in H3:(??%) ⊢ %; rewrite > H3; rewrite > len_copy; elim (\len ys); [reflexivity]
216           symmetry; apply (copy_OQ ys n2);
217         | cases H8 in H5 H7; simplify; intros; [2,6:reflexivity|3,4,5: assumption] 
218           simplify; rewrite > H5; rewrite > len_copy; elim (\len xs); [reflexivity]
219           symmetry; apply (copy_OQ xs n2);]
220     |2: apply (aux_preserves_sorting ? b3 ??? H8); assumption;
221     |3: apply (aux_preserves_sorting2 ? b3 ??? H8); try assumption;
222         try reflexivity; cases (inversion_sorted ?? H4);[2:rewrite >H3; apply (sorted_one q2_lt);]
223         cases l2 in H3 H4; intros; [apply (sorted_one q2_lt)]
224         apply (sorted_cons q2_lt);[2:apply (sorted_tail q2_lt ?? H3);] whd;
225         rewrite > E; assumption;
226     |4: apply I 
227     |5: apply I
228     |6: intro; elim i; intros; simplify; solve [symmetry;assumption|apply H13]
229     |7: unfold; intros; clear H9 H10 H11 H12 H13; simplify in Hi1 Hi2 H16 H18;
230        cases H8 in H14 H15 H17 H3 H16 H18 H5 H6; 
231        simplify in match (\fst 〈?,?〉); simplify in match (\snd 〈?,?〉); intros;
232        [1: reflexivity;
233        |2: simplify in H3; rewrite > (value_unit b); rewrite > H3; symmetry;
234            cases b in H3 H12 Hi2; intros 2; simplify in H12; rewrite > H12;
235            intros; change in ⊢ (? ? (? % ? ? ? ?) ?) with (copy (〈q,〈OQ,OQ〉〉::〈b1,〈OQ,OQ〉〉::ys));
236            apply (value_copy (〈q,〈OQ,OQ〉〉::〈b1,〈OQ,OQ〉〉::ys));
237        |3: apply (same_value_tail b b1 h1 h3 xs r1 input); assumption;
238        |4: apply (same_value_tail b b1 h1 h1 xs r1 input); assumption;
239        |5: simplify in H9; STOP
240              
241        |6: reflexivity;]
242                           
243             ]
244     |8: 
245     
246
247 include "Q/q/qtimes.ma".
248
249 let rec area (l:list bar) on l ≝
250   match l with 
251   [ nil ⇒ OQ
252   | cons he tl ⇒ area tl + Qpos (\fst he) * ⅆ[OQ,\snd he]].
253
254 alias symbol "pi1" = "exT \fst".
255 alias symbol "minus" = "Q minus".
256 alias symbol "exists" = "CProp exists".
257 definition minus_spec_bar ≝
258  λf,g,h:list bar.
259    same_bases f g → len f = len g →
260      ∀s,i:ℚ. \snd (\fst (value (mk_q_f s h) i)) = 
261        \snd (\fst (value (mk_q_f s f) i)) - \snd (\fst (value (mk_q_f s g) i)). 
262
263 definition minus_spec ≝
264  λf,g:q_f.
265    ∃h:q_f. 
266      ∀i:ℚ. \snd (\fst (value h i)) = 
267        \snd (\fst (value f i)) - \snd (\fst (value g i)). 
268
269 definition eject_bar : ∀P:list bar → CProp.(∃l:list bar.P l) → list bar ≝
270  λP.λp.match p with [ex_introT x _ ⇒ x].
271 definition inject_bar ≝ ex_introT (list bar).
272
273 coercion inject_bar with 0 1 nocomposites.
274 coercion eject_bar with 0 0 nocomposites.
275
276 lemma minus_q_f : ∀f,g. minus_spec f g.
277 intros;
278 letin aux ≝ (
279   let rec aux (l1, l2 : list bar) on l1 ≝
280     match l1 with
281     [ nil ⇒ []
282     | cons he1 tl1 ⇒
283         match l2 with
284         [ nil ⇒ []
285         | cons he2 tl2 ⇒ 〈\fst he1, \snd he1 - \snd he2〉 :: aux tl1 tl2]]
286   in aux : ∀l1,l2 : list bar.∃h.minus_spec_bar l1 l2 h);
287 [2: intros 4; simplify in H3; destruct H3;
288 |3: intros 4; simplify in H3; cases l1 in H2; [2: intro X; simplify in X; destruct X]    
289     intros; rewrite > (value_OQ_e (mk_q_f s []) i); [2: reflexivity]
290     rewrite > q_elim_minus; rewrite > q_plus_OQ; reflexivity;
291 |1: cases (aux l2 l3); unfold in H2; intros 4;
292     simplify in ⊢ (? ? (? ? ? (? ? ? (? % ?))) ?);
293     cases (q_cmp i (s + Qpos (\fst b)));
294     
295
296
297 definition excess ≝ 
298   λf,g.∃i.\snd (\fst (value f i)) < \snd (\fst (value g i)).
299