]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/models/q_function.ma
...
[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 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 definition same_values_simpl ≝
26  λl1,l2.∀p1,p2,p3,p4,p5,p6.same_values (mk_q_f l1 p1 p2 p3) (mk_q_f l2 p4 p5 p6).
27
28 alias symbol "pi2" = "pair pi2".
29 alias symbol "pi1" = "pair pi1".
30 definition rebase_spec_aux ≝ 
31  λl1,l2:list bar.λp:(list bar) × (list bar).
32    sorted q2_lt l1 → sorted q2_lt l2 →
33    (l1 ≠ [] → \snd (\nth l1 ▭ (pred (\len l1))) = 〈OQ,OQ〉) →
34    (l2 ≠ [] → \snd (\nth l2 ▭ (pred (\len l2))) = 〈OQ,OQ〉) →
35  And4
36    (nth_base l1 O = nth_base (\fst p) O ∨
37     nth_base l2 O = nth_base (\fst p) O) 
38    (sorted q2_lt (\fst p) ∧ sorted q2_lt (\snd p))
39    ((l1 ≠ [] → \snd (\nth (\fst p) ▭ (pred (\len (\fst p)))) = 〈OQ,OQ〉) ∧ 
40     (l2 ≠ [] → \snd (\nth (\snd p) ▭ (pred (\len (\snd p)))) = 〈OQ,OQ〉))  
41    (And3
42       (same_bases (\fst p) (\snd p))
43       (same_values_simpl l1 (\fst p)) 
44       (same_values_simpl l2 (\snd p))).   
45
46 definition eject ≝
47   λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
48 coercion eject.
49 definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h.
50 coercion inject with 0 1 nocomposites.
51
52 axiom devil : False.
53         
54 definition copy ≝
55  λl:list bar.make_list ? (λn.〈nth_base l (n - \len l),〈OQ,OQ〉〉) (\len l).
56  
57 lemma copy_rebases: 
58   ∀l1.rebase_spec_aux l1 [] 〈l1, copy l1〉.
59 intros; cases l1; intros 4;
60 [1: split; [left; reflexivity]; split; try assumption; unfold; intros;
61     unfold same_values; intros; reflexivity;
62 |2: rewrite > H2; [2: intro X; destruct X] clear H2 H3;
63     split; [left; reflexivity] split; 
64     unfold same_values_simpl; unfold same_values; intros; try reflexivity;
65     try assumption; [4: normalize in p2; destruct p2|2: cases H2; reflexivity;]
66     simplify; clear H1;
67     [1: elim (\len l) in H; simplify; [apply (sorted_one q2_lt);]
68         
69         
70     
71         
72 definition rebase: ∀l1,l2:q_f.∃p:q_f × q_f.rebase_spec l1 l2 p.
73 intros 2 (f1 f2); cases f1 (b1 Hs1 Hb1 He1); cases f2 (b2 Hs2 Hb2 He2); clear f1 f2;
74 alias symbol "plus" = "natural plus".
75 alias symbol "pi2" = "pair pi2".
76 alias symbol "pi1" = "pair pi1".
77 alias symbol "minus" = "Q minus".
78 letin aux ≝ ( 
79 let rec aux (l1,l2:list bar) (n : nat) on n : (list bar) × (list bar) ≝
80 match n with
81 [ O ⇒ 〈[], []〉
82 | S m ⇒
83   match l1 with
84   [ nil ⇒ 〈copy l2, l2〉
85   | cons he1 tl1 ⇒
86      match l2 with
87      [ nil ⇒ 〈l1, copy l1〉
88      | cons he2 tl2 ⇒  
89          let base1 ≝ \fst he1 in
90          let base2 ≝ \fst he2 in
91          let height1 ≝ \snd he1 in
92          let height2 ≝ \snd he2 in
93          match q_cmp base1 base2 with
94          [ q_leq Hp1 ⇒ 
95              match q_cmp base2 base1 with
96              [ q_leq Hp2 ⇒
97                  let rc ≝ aux tl1 tl2 m in 
98                  〈he1 :: \fst rc,he2 :: \snd rc〉
99              | q_gt Hp ⇒ 
100                  let rest ≝ base2 - base1 in
101                  let rc ≝ aux tl1 (〈rest,height2〉 :: tl2) m in
102                  〈〈base1,height1〉 :: \fst rc,〈base1,height2〉 :: \snd rc〉] 
103          | q_gt Hp ⇒ 
104              let rest ≝ base1 - base2 in
105              let rc ≝ aux (〈rest,height1〉 :: tl1) tl2 m in
106              〈〈base2,height1〉 :: \fst rc,〈base2,height2〉 :: \snd rc〉]]]]
107 in aux : ∀l1,l2,m.∃z.m = \len l1 + \len l2 → rebase_spec_aux l1 l2 z);
108 [7: clearbody aux; cases (aux b1 b2 (\len b1 + \len b2)) (w Hw); clear aux;
109     cases (Hw (refl_eq ??) Hs1 Hs2 (λ_.He1) (λ_.He2)); clear Hw; cases H1; cases H2; cases H3; clear H3 H1 H2;
110     exists [constructor 1;constructor 1;[apply (\fst w)|5:apply (\snd w)]] try assumption;
111     [1,3: apply hide; cases H (X X); try rewrite < (H8 O); try rewrite < X; assumption
112     |2,4: apply hide;[apply H6|apply H7]intro X;[rewrite > X in Hb1|rewrite > X in Hb2]
113          normalize in Hb1 Hb2; [destruct Hb1|destruct Hb2]]
114     unfold; unfold same_values; simplify in ⊢ (? (? % %) ? ?); 
115     simplify in match (\snd 〈?,?〉); simplify in match (\fst 〈?,?〉);
116     split; [assumption; |apply H9;|apply H10]
117 |6: intro ABS; unfold; intros 4; clear H1 H2;
118     cases l in ABS H3; intros 1; [2: simplify in H1; destruct H1]
119     cases l1 in H4 H1; intros; [2: simplify in H2; destruct H2]
120     split; [left;reflexivity|split; apply (sorted_nil q2_lt);|split; assumption;]
121     split; unfold; intros; unfold same_values; intros; reflexivity;
122 |5: unfold rebase_spec_aux; intros; cases l1 in H2 H4 H6; intros; [ simplify in H2; destruct H2;]
123     lapply H6 as H7; [2: intro X; destruct X] clear H6 H5;
124     rewrite > H7; split; [right; simplify;
125     
126      split; [left;reflexivity]
127     split; 
128
129 ,2: unfold rest; apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
130       assumption;        
131 |8: intros; cases (?:False); apply (not_le_Sn_O ? H1);
132 |3: intros; generalize in match (unpos ??); intro X; cases X; clear X;
133     simplify in ⊢ (???? (??? (??? (??? (?? (? (?? (??? % ?) ?) ??)))) ?));
134     simplify in ⊢ (???? (???? (??? (??? (?? (? (?? (??? % ?) ?) ??)))))); 
135     clear H4; cases (aux (〈w,\snd b〉::l4) l5 n1); clear aux;
136     cut (len (〈w,\snd b〉::l4) + len l5 < n1) as K;[2:
137       simplify in H5; simplify; rewrite > sym_plus in H5; simplify in H5;
138       rewrite > sym_plus in H5; apply le_S_S_to_le; apply H5;] 
139     split;
140     [1: simplify in ⊢ (? % ?); simplify in ⊢ (? ? %); 
141         cases (H4 s K); clear K H4; intro input; cases input; [reflexivity]
142         simplify; apply H7; 
143     |2: simplify in ⊢ (? ? %); cases (H4 s K); clear H4 K H5 spec;
144         intro;
145         (* input < s + b1 || input >= s + b1 *)
146     |3: simplify in ⊢ (? ? %);]   
147 |4: intros; generalize in match (unpos ??); intro X; cases X; clear X;
148     (* duale del 3 *)
149 |5: intros; (* triviale, caso in cui non fa nulla *)
150 |6,7: (* casi base in cui allunga la lista più corta *) 
151 ]
152 elim devil;
153 qed.
154
155 include "Q/q/qtimes.ma".
156
157 let rec area (l:list bar) on l ≝
158   match l with 
159   [ nil ⇒ OQ
160   | cons he tl ⇒ area tl + Qpos (\fst he) * ⅆ[OQ,\snd he]].
161
162 alias symbol "pi1" = "exT \fst".
163 alias symbol "minus" = "Q minus".
164 alias symbol "exists" = "CProp exists".
165 definition minus_spec_bar ≝
166  λf,g,h:list bar.
167    same_bases f g → len f = len g →
168      ∀s,i:ℚ. \snd (\fst (value (mk_q_f s h) i)) = 
169        \snd (\fst (value (mk_q_f s f) i)) - \snd (\fst (value (mk_q_f s g) i)). 
170
171 definition minus_spec ≝
172  λf,g:q_f.
173    ∃h:q_f. 
174      ∀i:ℚ. \snd (\fst (value h i)) = 
175        \snd (\fst (value f i)) - \snd (\fst (value g i)). 
176
177 definition eject_bar : ∀P:list bar → CProp.(∃l:list bar.P l) → list bar ≝
178  λP.λp.match p with [ex_introT x _ ⇒ x].
179 definition inject_bar ≝ ex_introT (list bar).
180
181 coercion inject_bar with 0 1 nocomposites.
182 coercion eject_bar with 0 0 nocomposites.
183
184 lemma minus_q_f : ∀f,g. minus_spec f g.
185 intros;
186 letin aux ≝ (
187   let rec aux (l1, l2 : list bar) on l1 ≝
188     match l1 with
189     [ nil ⇒ []
190     | cons he1 tl1 ⇒
191         match l2 with
192         [ nil ⇒ []
193         | cons he2 tl2 ⇒ 〈\fst he1, \snd he1 - \snd he2〉 :: aux tl1 tl2]]
194   in aux : ∀l1,l2 : list bar.∃h.minus_spec_bar l1 l2 h);
195 [2: intros 4; simplify in H3; destruct H3;
196 |3: intros 4; simplify in H3; cases l1 in H2; [2: intro X; simplify in X; destruct X]    
197     intros; rewrite > (value_OQ_e (mk_q_f s []) i); [2: reflexivity]
198     rewrite > q_elim_minus; rewrite > q_plus_OQ; reflexivity;
199 |1: cases (aux l2 l3); unfold in H2; intros 4;
200     simplify in ⊢ (? ? (? ? ? (? ? ? (? % ?))) ?);
201     cases (q_cmp i (s + Qpos (\fst b)));
202     
203
204
205 definition excess ≝ 
206   λf,g.∃i.\snd (\fst (value f i)) < \snd (\fst (value g i)).
207