]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/models/q_function.ma
Major reordering of theorems in the appropriate files.
[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 "russell_support.ma".
16 include "models/q_bars.ma".
17
18 (* move in nat/minus *)
19 lemma minus_lt : ∀i,j. i < j → j - i = S (j - S i). 
20 intros 2;
21 apply (nat_elim2 ???? i j); simplify; intros;
22 [1: cases n in H; intros; rewrite < minus_n_O; [cases (not_le_Sn_O ? H);]
23     simplify; rewrite < minus_n_O; reflexivity;
24 |2: cases (not_le_Sn_O ? H);
25 |3: apply H; apply le_S_S_to_le; assumption;]
26 qed.
27
28 definition copy ≝
29  λl:list bar.make_list ? (λn.〈nth_base l (\len l - S n),〈OQ,OQ〉〉) (\len l).
30
31 lemma sorted_copy:
32   ∀l:list bar.sorted q2_lt l → sorted q2_lt (copy l).
33 intros 2; unfold copy; generalize in match (le_n (\len l)); 
34 elim (\len l) in ⊢ (?%?→? ? (? ? ? %));
35 simplify; [apply (sorted_nil q2_lt);] cases n in H1 H2;
36 simplify; intros; [apply (sorted_one q2_lt);]
37 apply (sorted_cons q2_lt);
38 [2: apply H1; apply lt_to_le; apply H2;
39 |1: elim l in H2 H; simplify; [simplify in H2; cases (not_le_Sn_O ? H2)]    
40     simplify in H3; unfold nth_base;
41     unfold canonical_q_lt; unfold q2_trel; unfold q2_lt; simplify;
42     change with (q2_lt (\nth (a::l1) ▭ (\len l1-S n1)) (\nth (a::l1) ▭ (\len l1-n1)));
43     cut (∃w.w = \len l1 - S n1); [2: exists[apply (\len l1 - S n1)] reflexivity]
44     cases Hcut; rewrite < H4; rewrite < (?:S w = \len l1 - n1);
45     [1: apply (sorted_near q2_lt (a::l1) H2); rewrite > H4;
46         simplify; apply le_S_S; elim (\len l1) in H3; simplify;
47         [ cases (not_le_Sn_O ? (le_S_S_to_le ?? H3));
48         | lapply le_S_S_to_le to H5 as H6;
49           lapply le_S_S_to_le to H6 as H7; clear H5 H6;
50           cases H7 in H3; intros; [rewrite < minus_n_n; apply le_S_S; apply le_O_n]
51           simplify in H5; apply le_S_S; apply (trans_le ???? (H5 ?));
52           [2: apply le_S_S; apply le_S_S; assumption;
53           |1: apply (lt_minus_S_n_to_le_minus_n n1 (S m) (S (minus m n1)) ?).
54               apply (not_le_to_lt (S (minus m n1)) (minus (S m) (S n1)) ?).
55               apply (not_le_Sn_n (minus (S m) (S n1))).]]
56     |2: rewrite > H4; lapply le_S_S_to_le to H3 as K;
57         clear H4 Hcut H3 H H1 H2; generalize in match K; clear K;
58         apply (nat_elim2 ???? n1 (\len l1)); simplify; intros;
59         [1: rewrite < minus_n_O; cases n2 in H; [intro; cases (not_le_Sn_O ? H)]
60             intros; cases n3; simplify; reflexivity;
61         |2: cases (not_le_Sn_O ? H);
62         |3: apply H; apply le_S_S_to_le; apply H1;]]]
63 qed.
64
65 lemma len_copy: ∀l. \len (copy l) = \len l. 
66 intro; unfold copy; apply len_mk_list;
67 qed.
68
69 lemma same_bases_cons: ∀a,b,l1,l2.
70   same_bases l1 l2 → \fst a = \fst b → same_bases (a::l1) (b::l2).
71 intros; intro; cases i; simplify; [assumption;] apply (H n);
72 qed.
73
74 lemma copy_same_bases: ∀l. same_bases l (copy l).
75 intro; unfold copy; elim l using list_elim_with_len; [1: intro;reflexivity]
76 simplify; rewrite < minus_n_n;
77 simplify in ⊢ (? ? (? ? (? ? ? % ?) ?));
78 apply same_bases_cons; [2: reflexivity]
79 cases l1 in H; [intros 2; reflexivity]
80 simplify in ⊢ (? ? (? ? (λ_:?.? ? ? (? ? %) ?) ?)→?);
81 simplify in ⊢ (?→? ? (? ? (λ_:?.? ? ? (? ? (? % ?)) ?) ?));
82 intro; rewrite > (mk_list_ext ?? (λn:nat.〈nth_base (b::l2) (\len l2-n),〈OQ,OQ〉〉));[assumption]
83 intro; elim x; [simplify; rewrite < minus_n_O; reflexivity]
84 simplify in ⊢ (? ? (? ? ? (? ? %) ?) ?);
85 simplify in H2:(? ? %); rewrite > minus_lt; [reflexivity] apply le_S_S_to_le;
86 assumption;
87 qed.
88
89 lemma prepend_sorted_with_same_head:
90  ∀r,x,l1,l2,d1,d2.
91    sorted r (x::l1) → sorted r l2 → 
92    (r x (\nth l1 d1 O) → r x (\nth l2 d2 O)) → (l1 = [] → r x d1) →
93    sorted r (x::l2).
94 intros 8 (R x l1 l2 d1 d2 Sl1 Sl2);  inversion Sl1; inversion Sl2;
95 intros; destruct; try assumption; [3: apply (sorted_one R);]
96 [1: apply sorted_cons;[2:assumption] apply H2; apply H3; reflexivity;
97 |2: apply sorted_cons;[2: assumption] apply H5; apply H6; reflexivity;
98 |3: apply sorted_cons;[2: assumption] apply H5; assumption;
99 |4: apply sorted_cons;[2: assumption] apply H8; apply H4;]
100 qed.
101
102 lemma move_head_sorted: ∀x,l1,l2. 
103   sorted q2_lt (x::l1) → sorted q2_lt l2 → nth_base l2 O = nth_base l1 O →
104     l1 ≠ [] → sorted q2_lt (x::l2).
105 intros; apply (prepend_sorted_with_same_head q2_lt x l1 l2 ▭ ▭);
106 try assumption; intros; unfold nth_base in H2; whd in H4;
107 [1: rewrite < H2 in H4; assumption;
108 |2: cases (H3 H4);]
109 qed.
110        
111 definition rebase_spec ≝ 
112  λl1,l2:q_f.λp:q_f × q_f. 
113    And3
114     (same_bases (bars (\fst p)) (bars (\snd p)))
115     (same_values l1 (\fst p)) 
116     (same_values l2 (\snd p)).
117
118 definition last ≝
119  λT:Type.λd.λl:list T. \nth l d (pred (\len l)).
120
121 notation > "\last" non associative with precedence 90 for @{'last}.
122 notation < "\last d l" non associative with precedence 70 for @{'last2 $d $l}.
123 interpretation "list last" 'last = (last _). 
124 interpretation "list last applied" 'last2 d l = (last _ d l).
125
126 definition head ≝
127  λT:Type.λd.λl:list T.\nth l d O.
128
129 notation > "\hd" non associative with precedence 90 for @{'hd}.
130 notation < "\hd d l" non associative with precedence 70 for @{'hd2 $d $l}.
131 interpretation "list head" 'hd = (head _).
132 interpretation "list head applied" 'hd2 d l = (head _ d l).
133
134 alias symbol "lt" = "bar lt".
135 lemma inversion_sorted:
136   ∀a,l. sorted q2_lt (a::l) → a < \hd ▭ l ∨ l = [].
137 intros 2; elim l; [right;reflexivity] left; inversion H1; intros;
138 [1,2:destruct H2| destruct H5; assumption]
139 qed.
140
141 lemma inversion_sorted2:
142   ∀a,b,l. sorted q2_lt (a::b::l) → a < b.
143 intros; inversion H; intros; [1,2:destruct H1] destruct H4; assumption; 
144 qed.
145
146 definition same_values_simpl ≝
147  λl1,l2.∀p1,p2,p3,p4,p5,p6.
148    same_values (mk_q_f l1 p1 p2 p3) (mk_q_f l2 p4 p5 p6).
149
150 inductive rebase_cases : list bar → list bar → (list bar) × (list bar) → Prop ≝
151 | rb_fst_full  : ∀b,h1,h2,xs,ys,r1,r2. 〈b,h1〉 < \hd ▭ ys → 
152    \snd(\last ▭ (〈b,h1〉::xs)) = \snd(\last ▭ (〈b,h1〉::r1)) →
153    \snd(\last ▭ (〈b,h1〉::xs)) = \snd(\last ▭ (〈b,h2〉::r2)) →
154    rebase_cases (〈b,h1〉::xs) ys 〈〈b,h1〉::r1,〈b,h2〉::r2〉
155 | rb_snd_full  : ∀b,h1,h2,xs,ys,r1,r2. 〈b,h1〉 < \hd ▭ xs →
156    \snd(\last ▭ (〈b,h1〉::ys)) = \snd(\last ▭ (〈b,h1〉::r2)) →
157    \snd(\last ▭ (〈b,h1〉::ys)) = \snd(\last ▭ (〈b,h2〉::r1)) →
158    rebase_cases xs (〈b,h1〉::ys) 〈〈b,h2〉::r1,〈b,h1〉::r2〉  
159 | rb_all_full  : ∀b,h1,h2,h3,h4,xs,ys,r1,r2.
160    \snd(\last ▭ (〈b,h1〉::xs)) = \snd(\last ▭ (〈b,h3〉::r1)) →
161    \snd(\last ▭ (〈b,h2〉::ys)) = \snd(\last ▭ (〈b,h4〉::r2)) →
162    rebase_cases (〈b,h1〉::xs) (〈b,h2〉::ys) 〈〈b,h3〉::r1,〈b,h4〉::r2〉 
163 | rb_all_empty : rebase_cases [] [] 〈[],[]〉.
164
165 alias symbol "pi2" = "pair pi2".
166 alias symbol "pi1" = "pair pi1".
167 alias symbol "leq" = "natural 'less or equal to'".
168 inductive rebase_spec_aux_p (l1, l2:list bar) (p:(list bar) × (list bar)) : Prop ≝
169 | prove_rebase_spec_aux:
170    rebase_cases l1 l2 p →
171    (sorted q2_lt (\fst p)) →
172    (sorted q2_lt (\snd p)) → True → True → (*
173    ((l1 ≠ [] ∧ \snd (\last ▭ (\fst p)) = 〈OQ,OQ〉) ∨
174     (l1 = [] ∧ ∀i.\snd (\nth (\fst p) ▭ i) = 〈OQ,OQ〉)) → 
175    ((l2 ≠ [] ∧ \snd (\last ▭ (\snd p)) = 〈OQ,OQ〉) ∨
176     (l2 = [] ∧ ∀i.\snd (\nth (\snd p) ▭ i) = 〈OQ,OQ〉)) →*)
177    (same_bases (\fst p) (\snd p)) →
178    (same_values_simpl l1 (\fst p)) → 
179    (same_values_simpl l2 (\snd p)) →  
180    (*\len (\fst p) ≤ \len l1 + \len l2 → *)
181    (*\len (\fst p) = \len (\snd p) → *)
182    rebase_spec_aux_p l1 l2 p.   
183
184 lemma sort_q2lt_same_base:
185   ∀b,h1,h2,l. sorted q2_lt (〈b,h1〉::l) → sorted q2_lt (〈b,h2〉::l).
186 intros; cases (inversion_sorted ?? H); [2: rewrite > H1; apply (sorted_one q2_lt)]
187 lapply (sorted_tail q2_lt ?? H) as K; clear H; cases l in H1 K; simplify; intros;
188 [apply (sorted_one q2_lt);|apply (sorted_cons q2_lt);[2: assumption] apply H]
189 qed.
190
191 lemma aux_preserves_sorting:
192  ∀b,b3,l2,l3,w. rebase_cases l2 l3 w → 
193   sorted q2_lt (b::l2) → sorted q2_lt (b3::l3) → \fst b3 = \fst b →
194   sorted q2_lt (\fst w) → sorted q2_lt (\snd w) → same_bases (\fst w) (\snd w) → 
195   sorted q2_lt (b :: \fst w).
196 intros 6; cases H; simplify; intros;
197 [ apply (sorted_cons q2_lt); [2:assumption] apply (inversion_sorted2 ??? H4);   
198 | apply (sorted_cons q2_lt); [2:apply (sort_q2lt_same_base ???? H7);]
199   whd; rewrite < H6; apply (inversion_sorted2 ??? H5);
200 | apply (sorted_cons q2_lt); [2: assumption] apply (inversion_sorted2 ??? H3);
201 | apply (sorted_one q2_lt);]
202 qed.   
203
204 lemma aux_preserves_sorting2:
205  ∀b,b3,l2,l3,w. rebase_cases l2 l3 w → 
206   sorted q2_lt (b::l2) → sorted q2_lt (b3::l3) → \fst b3 = \fst b →
207   sorted q2_lt (\fst w) → sorted q2_lt (\snd w) → same_bases (\fst w) (\snd w) → 
208   sorted q2_lt (b :: \snd w).
209 intros 6; cases H; simplify; intros;
210 [ apply (sorted_cons q2_lt); [2:assumption] apply (inversion_sorted2 ??? H4);   
211 | apply (sorted_cons q2_lt); [2:assumption] 
212   whd; rewrite < H6; apply (inversion_sorted2 ??? H5);
213 | apply (sorted_cons q2_lt); [2: assumption] apply (inversion_sorted2 ??? H3);
214 | apply (sorted_one q2_lt);]
215 qed.   
216
217 definition rebase_spec_aux ≝ 
218  λl1,l2:list bar.λp:(list bar) × (list bar).
219  sorted q2_lt l1 → (\snd (\last ▭ l1) = 〈OQ,OQ〉) →
220  sorted q2_lt l2 → (\snd (\last ▭ l2) = 〈OQ,OQ〉) → 
221    rebase_spec_aux_p l1 l2 p.
222     
223 definition eject ≝
224   λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
225 coercion eject.
226 definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h.
227 coercion inject with 0 1 nocomposites.
228
229 definition rebase: ∀l1,l2:q_f.∃p:q_f × q_f.rebase_spec l1 l2 p.
230 intros 2 (f1 f2); cases f1 (b1 Hs1 Hb1 He1); cases f2 (b2 Hs2 Hb2 He2); clear f1 f2;
231 alias symbol "leq" = "natural 'less or equal to'".
232 alias symbol "minus" = "Q minus".
233 letin aux ≝ ( 
234 let rec aux (l1,l2:list bar) (n : nat) on n : (list bar) × (list bar) ≝
235 match n with
236 [ O ⇒ 〈[], []〉
237 | S m ⇒
238   match l1 with
239   [ nil ⇒ 〈copy l2, l2〉
240   | cons he1 tl1 ⇒
241      match l2 with
242      [ nil ⇒ 〈l1, copy l1〉
243      | cons he2 tl2 ⇒  
244          let base1 ≝ \fst he1 in
245          let base2 ≝ \fst he2 in
246          let height1 ≝ \snd he1 in
247          let height2 ≝ \snd he2 in
248          match q_cmp base1 base2 with
249          [ q_leq Hp1 ⇒ 
250              match q_cmp base2 base1 with
251              [ q_leq Hp2 ⇒
252                  let rc ≝ aux tl1 tl2 m in 
253                  〈he1 :: \fst rc,he2 :: \snd rc〉
254              | q_gt Hp ⇒ 
255                  let rest ≝ base2 - base1 in
256                  let rc ≝ aux tl1 (〈rest,height2〉 :: tl2) m in
257                  〈〈base1,height1〉 :: \fst rc,〈base1,height2〉 :: \snd rc〉] 
258          | q_gt Hp ⇒ 
259              let rest ≝ base1 - base2 in
260              let rc ≝ aux (〈rest,height1〉 :: tl1) tl2 m in
261              〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉]]]]
262 in aux : ∀l1,l2,m.∃z.\len l1 + \len l2 ≤ m → rebase_spec_aux l1 l2 z);
263 [7: clearbody aux; cases (aux b1 b2 (\len b1 + \len b2)) (res Hres);
264     exists; [split; constructor 1; [apply (\fst res)|5:apply (\snd res)]]
265     [1,4: apply hide; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); assumption;
266     |2,5: apply hide; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); lapply (H5 O) as K;
267           clear H1 H2 H3 H4 H5 H6 H7 Hres aux; unfold nth_base;
268           cases H in K He1 He2 Hb1 Hb2; simplify; intros; assumption;
269     |3,6: apply hide; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); 
270           cases H in He1 He2; simplify; intros;
271           [4,8: assumption;
272           |1,6,7: rewrite < H9; assumption;
273           |2,3,5: rewrite < H10; [2: symmetry] assumption;]]
274     split; cases (Hres (le_n ?) Hs1 He1 Hs2 He2); unfold same_values; intros; 
275     [1: assumption
276     |2: simplify in match (\snd 〈?,?〉); simplify in match (\fst 〈?,?〉); apply H6;
277     |3: simplify in match (\snd 〈?,?〉); simplify in match (\fst 〈?,?〉); apply H7]
278 |3: cut (\fst b3 = \fst b) as E; [2: apply q_le_to_le_to_eq; assumption]
279     clear H6 H5 H4 H3 He2 Hb2 Hs2 b2 He1 Hb1 Hs1 b1; cases (aux l2 l3 n1); 
280     clear aux; intro K; simplify in K; rewrite <plus_n_Sm in K; 
281     lapply le_S_S_to_le to K as W; lapply lt_to_le to W as R; 
282     simplify in match (? ≪w,H3≫); intros; cases (H3 R); clear H3 R W K;
283       [2,4: apply (sorted_tail q2_lt);[apply b|3:apply b3]assumption;
284       |3: cases l2 in H5; simplify; intros; try reflexivity; assumption; 
285       |5: cases l3 in H7; simplify; intros; try reflexivity; assumption;]
286     constructor 1; simplify in match (\fst 〈?,?〉); simplify in match (\snd 〈?,?〉);
287     [1: cases b in E H5 H7 H11 H14; cases b3; intros (E H5 H7 H11 H14); simplify in E; 
288         destruct E; constructor 3;
289         [ cases H8 in H5 H7; intros; [1,3:assumption] simplify;
290           [ rewrite > H16; rewrite < H7; symmetry; apply H17; | reflexivity]
291         | cases H8 in H5 H7; simplify; intros; [2,3: assumption] 
292           [ rewrite < H7; rewrite > H16; apply H17; | reflexivity]]
293     |2: apply (aux_preserves_sorting ? b3 ??? H8); assumption;
294     |3: apply (aux_preserves_sorting2 ? b3 ??? H8); try assumption;
295         try reflexivity; cases (inversion_sorted ?? H4);[2:rewrite >H3; apply (sorted_one q2_lt);]
296         cases l2 in H3 H4; intros; [apply (sorted_one q2_lt)]
297         apply (sorted_cons q2_lt);[2:apply (sorted_tail q2_lt ?? H3);] whd;
298         rewrite > E; assumption;
299     |4: apply I 
300     |5: apply I
301     |6: intro; elim i; intros; simplify; solve [symmetry;assumption|apply H13]
302     |7: unfold; intros; apply H14;
303     |8: 
304     
305     clear H15 H14 H11 H12 H7 H5; cases H8; clear H8; cases H3; clear H3;
306         [1: apply (move_head ?? H4 H5 ?? H9); symmetry; assumption;
307         |2: apply (move_head ??? H5 ?? H9); [apply (soted_q2lt_rewrite_hd ??? E  H6)]
308             symmetry; rewrite > (H13 O); assumption;
309         |3: apply (soted_q2lt_rewrite_hd ??? E); apply (move_head ?? H6 H7); [symmetry;] assumption;  
310         |4: rewrite > H8; apply (sorted_one q2_lt);]
311     
312     
313      cases l2 in H4 H8 H16; cases l3 in H6;
314         intros; cases H5; clear H5; cases H7; clear H7;
315         [1,2: cases (q_lt_corefl ? H5);
316         |3: rewrite >(?:\fst w = []); [apply (sorted_one q2_lt)]
317             simplify in H6; lapply (le_n_O_to_eq ? H6) as K;
318             cases (\fst w) in K; simplify; intro X [reflexivity|destruct X]
319         |4:
320         
321          rewrite >H8 in H5; cases (\fst w) in H9 H5; intros [apply (sorted_one q2_lt)]
322                 
323
324
325
326
327
328
329
330
331     
332 definition same_values_simpl ≝
333  λl1,l2.∀p1,p2,p3,p4,p5,p6.same_values (mk_q_f l1 p1 p2 p3) (mk_q_f l2 p4 p5 p6).
334
335 alias symbol "pi2" = "pair pi2".
336 alias symbol "pi1" = "pair pi1".
337 definition rebase_spec_aux ≝ 
338  λl1,l2:list bar.λp:(list bar) × (list bar).
339    sorted q2_lt l1 → sorted q2_lt l2 →
340    (l1 ≠ [] → \snd (\nth l1 ▭ (pred (\len l1))) = 〈OQ,OQ〉) →
341    (l2 ≠ [] → \snd (\nth l2 ▭ (pred (\len l2))) = 〈OQ,OQ〉) →
342  And4
343    (nth_base l1 O = nth_base (\fst p) O ∨
344     nth_base l2 O = nth_base (\fst p) O) 
345    (sorted q2_lt (\fst p) ∧ sorted q2_lt (\snd p))
346    ((l1 ≠ [] → \snd (\nth (\fst p) ▭ (pred (\len (\fst p)))) = 〈OQ,OQ〉) ∧ 
347     (l2 ≠ [] → \snd (\nth (\snd p) ▭ (pred (\len (\snd p)))) = 〈OQ,OQ〉))  
348    (And3
349       (same_bases (\fst p) (\snd p))
350       (same_values_simpl l1 (\fst p)) 
351       (same_values_simpl l2 (\snd p))).   
352
353 lemma copy_rebases: 
354   ∀l1.rebase_spec_aux l1 [] 〈l1, copy l1〉.
355 intros; elim l1; intros 4;
356 [1: split; [left; reflexivity]; split; try assumption; unfold; intros;
357     unfold same_values; intros; reflexivity;
358 |2: rewrite > H3; [2: intro X; destruct X]
359     split; [left; reflexivity] split; 
360     unfold same_values_simpl; unfold same_values; intros; try reflexivity;
361     try assumption; [4: normalize in p2; destruct p2|2: cases H5; reflexivity;]
362     [1: apply (sorted_copy ? H1);
363     |2: apply (copy_same_bases (a::l));]]
364 qed.
365
366 lemma copy_rebases_r: 
367   ∀l1.rebase_spec_aux [] l1 〈copy l1, l1〉.
368 intros; elim l1; intros 4;
369 [1: split; [left; reflexivity]; split; try assumption; unfold; intros;
370     unfold same_values; intros; reflexivity;
371 |2: rewrite > H4; [2: intro X; destruct X]
372     split; [right; simplify; rewrite < minus_n_n; reflexivity] split; 
373     unfold same_values_simpl; unfold same_values; intros; try reflexivity;
374     try assumption; [4: normalize in p2; destruct p2|2: cases H5; reflexivity;]
375     [1: apply (sorted_copy ? H2);
376     |2: intro; symmetry; apply (copy_same_bases (a::l));]]
377 qed.
378
379
380 definition eject ≝
381   λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
382 coercion eject.
383 definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h.
384 coercion inject with 0 1 nocomposites.
385
386 definition rebase: ∀l1,l2:q_f.∃p:q_f × q_f.rebase_spec l1 l2 p.
387 intros 2 (f1 f2); cases f1 (b1 Hs1 Hb1 He1); cases f2 (b2 Hs2 Hb2 He2); clear f1 f2;
388 alias symbol "leq" = "natural 'less or equal to'".
389 alias symbol "minus" = "Q minus".
390 letin aux ≝ ( 
391 let rec aux (l1,l2:list bar) (n : nat) on n : (list bar) × (list bar) ≝
392 match n with
393 [ O ⇒ 〈[], []〉
394 | S m ⇒
395   match l1 with
396   [ nil ⇒ 〈copy l2, l2〉
397   | cons he1 tl1 ⇒
398      match l2 with
399      [ nil ⇒ 〈l1, copy l1〉
400      | cons he2 tl2 ⇒  
401          let base1 ≝ \fst he1 in
402          let base2 ≝ \fst he2 in
403          let height1 ≝ \snd he1 in
404          let height2 ≝ \snd he2 in
405          match q_cmp base1 base2 with
406          [ q_leq Hp1 ⇒ 
407              match q_cmp base2 base1 with
408              [ q_leq Hp2 ⇒
409                  let rc ≝ aux tl1 tl2 m in 
410                  〈he1 :: \fst rc,he2 :: \snd rc〉
411              | q_gt Hp ⇒ 
412                  let rest ≝ base2 - base1 in
413                  let rc ≝ aux tl1 (〈rest,height2〉 :: tl2) m in
414                  〈〈base1,height1〉 :: \fst rc,〈base1,height2〉 :: \snd rc〉] 
415          | q_gt Hp ⇒ 
416              let rest ≝ base1 - base2 in
417              let rc ≝ aux (〈rest,height1〉 :: tl1) tl2 m in
418              〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉]]]]
419 in aux : ∀l1,l2,m.∃z.\len l1 + \len l2 ≤ m → rebase_spec_aux l1 l2 z);
420 [7: clearbody aux; cases (aux b1 b2 (\len b1 + \len b2)) (w Hw); clear aux;
421     cases (Hw (le_n ?) Hs1 Hs2 (λ_.He1) (λ_.He2)); clear Hw; cases H1; cases H2; cases H3; clear H3 H1 H2;
422     exists [constructor 1;constructor 1;[apply (\fst w)|5:apply (\snd w)]] try assumption;
423     [1,3: apply hide; cases H (X X); try rewrite < (H8 O); try rewrite < X; assumption
424     |2,4: apply hide;[apply H6|apply H7]intro X;[rewrite > X in Hb1|rewrite > X in Hb2]
425          normalize in Hb1 Hb2; [destruct Hb1|destruct Hb2]]
426     unfold; unfold same_values; simplify in ⊢ (? (? % %) ? ?); 
427     simplify in match (\snd 〈?,?〉); simplify in match (\fst 〈?,?〉);
428     split; [assumption; |apply H9;|apply H10]
429 |6: intro ABS; unfold; intros 4; clear H1 H2;
430     cases l in ABS H3; intros 1; [2: simplify in H1; cases (not_le_Sn_O ? H1)]
431     cases l1 in H4 H1; intros; [2: simplify in H2; cases (not_le_Sn_O ? H2)]
432     split; [ left; reflexivity|split; apply (sorted_nil q2_lt);|split; assumption;]
433     split; unfold; intros; unfold same_values; intros; reflexivity;
434 |5: intros; apply copy_rebases_r;
435 |4: intros; rewrite < H1; apply copy_rebases;
436 |3: cut (\fst b = \fst b3) as K; [2: apply q_le_to_le_to_eq; assumption] clear H6 H5 H4 H3;
437     intros; cases (aux l2 l3 n1); cases w in H4 (w1 w2); clear w; 
438     intros 5; 
439     simplify in match (\fst 〈?,?〉); simplify in match (\snd 〈?,?〉);
440     cases H5; 
441       [2: apply le_S_S_to_le; apply (trans_le ???? H3); simplify;
442           rewrite < plus_n_Sm; apply le_S; apply le_n;
443       |3,4: apply (sorted_tail q2_lt); [2: apply H4|4:apply H6]
444       |5: intro; cases l2 in H7 H9; intros; [cases H9; reflexivity]
445           simplify in H7 ⊢ %; apply H7; intro; destruct H10;
446       |6: intro; cases l3 in H8 H9; intros; [cases H9; reflexivity]
447           simplify in H8 ⊢ %; apply H8; intro; destruct H10;]
448     clear aux H5; 
449     simplify in match (\fst 〈?,?〉) in H9 H10 H11 H12; 
450     simplify in match (\snd 〈?,?〉) in H9 H10 H11 H12;
451     split; 
452     [1: left; reflexivity;
453     |2: cases H10; cases H12; clear H15 H16 H12 H7 H8 H11 H10;
454         cases H9; clear H9;
455         [1: lapply (H14 O) as K1; clear H14; change in K1 with (nth_base w1 O = nth_base w2 O);
456             split; 
457             [1: apply (move_head_sorted ??? H4 H5 H7); STOP
458
459     
460      
461  unfold rebase_spec_aux; intros; cases l1 in H2 H4 H6; intros; [ simplify in H2; destruct H2;]
462     lapply H6 as H7; [2: intro X; destruct X] clear H6 H5;
463     rewrite > H7; split; [right; simplify;
464     
465      split; [left;reflexivity]
466     split; 
467
468 ,2: unfold rest; apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
469       assumption;        
470 |8: intros; cases (?:False); apply (not_le_Sn_O ? H1);
471 |3: intros; generalize in match (unpos ??); intro X; cases X; clear X;
472     simplify in ⊢ (???? (??? (??? (??? (?? (? (?? (??? % ?) ?) ??)))) ?));
473     simplify in ⊢ (???? (???? (??? (??? (?? (? (?? (??? % ?) ?) ??)))))); 
474     clear H4; cases (aux (〈w,\snd b〉::l4) l5 n1); clear aux;
475     cut (len (〈w,\snd b〉::l4) + len l5 < n1) as K;[2:
476       simplify in H5; simplify; rewrite > sym_plus in H5; simplify in H5;
477       rewrite > sym_plus in H5; apply le_S_S_to_le; apply H5;] 
478     split;
479     [1: simplify in ⊢ (? % ?); simplify in ⊢ (? ? %); 
480         cases (H4 s K); clear K H4; intro input; cases input; [reflexivity]
481         simplify; apply H7; 
482     |2: simplify in ⊢ (? ? %); cases (H4 s K); clear H4 K H5 spec;
483         intro;
484         (* input < s + b1 || input >= s + b1 *)
485     |3: simplify in ⊢ (? ? %);]   
486 |4: intros; generalize in match (unpos ??); intro X; cases X; clear X;
487     (* duale del 3 *)
488 |5: intros; (* triviale, caso in cui non fa nulla *)
489 |6,7: (* casi base in cui allunga la lista più corta *) 
490 ]
491 elim devil;
492 qed.
493
494 include "Q/q/qtimes.ma".
495
496 let rec area (l:list bar) on l ≝
497   match l with 
498   [ nil ⇒ OQ
499   | cons he tl ⇒ area tl + Qpos (\fst he) * ⅆ[OQ,\snd he]].
500
501 alias symbol "pi1" = "exT \fst".
502 alias symbol "minus" = "Q minus".
503 alias symbol "exists" = "CProp exists".
504 definition minus_spec_bar ≝
505  λf,g,h:list bar.
506    same_bases f g → len f = len g →
507      ∀s,i:ℚ. \snd (\fst (value (mk_q_f s h) i)) = 
508        \snd (\fst (value (mk_q_f s f) i)) - \snd (\fst (value (mk_q_f s g) i)). 
509
510 definition minus_spec ≝
511  λf,g:q_f.
512    ∃h:q_f. 
513      ∀i:ℚ. \snd (\fst (value h i)) = 
514        \snd (\fst (value f i)) - \snd (\fst (value g i)). 
515
516 definition eject_bar : ∀P:list bar → CProp.(∃l:list bar.P l) → list bar ≝
517  λP.λp.match p with [ex_introT x _ ⇒ x].
518 definition inject_bar ≝ ex_introT (list bar).
519
520 coercion inject_bar with 0 1 nocomposites.
521 coercion eject_bar with 0 0 nocomposites.
522
523 lemma minus_q_f : ∀f,g. minus_spec f g.
524 intros;
525 letin aux ≝ (
526   let rec aux (l1, l2 : list bar) on l1 ≝
527     match l1 with
528     [ nil ⇒ []
529     | cons he1 tl1 ⇒
530         match l2 with
531         [ nil ⇒ []
532         | cons he2 tl2 ⇒ 〈\fst he1, \snd he1 - \snd he2〉 :: aux tl1 tl2]]
533   in aux : ∀l1,l2 : list bar.∃h.minus_spec_bar l1 l2 h);
534 [2: intros 4; simplify in H3; destruct H3;
535 |3: intros 4; simplify in H3; cases l1 in H2; [2: intro X; simplify in X; destruct X]    
536     intros; rewrite > (value_OQ_e (mk_q_f s []) i); [2: reflexivity]
537     rewrite > q_elim_minus; rewrite > q_plus_OQ; reflexivity;
538 |1: cases (aux l2 l3); unfold in H2; intros 4;
539     simplify in ⊢ (? ? (? ? ? (? ? ? (? % ?))) ?);
540     cases (q_cmp i (s + Qpos (\fst b)));
541     
542
543
544 definition excess ≝ 
545   λf,g.∃i.\snd (\fst (value f i)) < \snd (\fst (value g i)).
546