]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/dama/dama/models/q_function.ma
bb3ba55fc8c49ff33bacbef1b8f5c2c6b8fd81f1
[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 "models/q_bars.ma".
16
17 lemma initial_shift_same_values:
18   ∀l1:q_f.∀init.init < start l1 →
19    same_values l1 
20      (mk_q_f init (〈\fst (unpos (start l1 - init) ?),OQ〉:: bars l1)).  
21 [apply hide; apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ; assumption]
22 intros; generalize in ⊢ (? ? (? ? (? ? (? ? ? (? ? ? (? ? %)) ?) ?))); intro;
23 cases (unpos (start l1-init) H1); intro input;
24 simplify in ⊢ (? ? ? (? ? ? (? ? ? (? (? ? (? ? (? ? ? % ?) ?)) ?))));
25 cases (value (mk_q_f init (〈w,OQ〉::bars l1)) input);
26 simplify in ⊢ (? ? ? (? ? ? %));
27 cases (q_cmp input (start (mk_q_f init (〈w,OQ〉::bars l1)))) in H3;
28 whd in ⊢ (% → ?); simplify in H3;
29 [1: intro; cases H4; clear H4; rewrite > H3;
30     cases (value l1 init); simplify; cases (q_cmp init (start l1)) in H4;
31     [1: cases (?:False); apply (q_lt_corefl init); rewrite > H4 in ⊢ (?? %); apply H;
32     |3: cases (?:False); apply (q_lt_antisym init (start l1)); assumption;
33     |2: whd in ⊢ (% → ?); intro; rewrite > H8; clear H8 H4;
34         rewrite > H7; clear H7; rewrite > (?:\fst w1 = O); [reflexivity]
35         symmetry; apply le_n_O_to_eq;
36         rewrite > (sum_bases_O (mk_q_f init (〈w,OQ〉::bars l1)) (\fst w1)); [apply le_n]   
37         clear H6 w2; simplify in H5:(? ? (? ? %));  
38         destruct H3; rewrite > q_d_x_x in H5; assumption;]
39 |2: intros; cases (value l1 input); simplify in ⊢ (? ? (? ? ? %) ?);
40     cases (q_cmp input (start l1)) in H5; whd in ⊢ (% → ?);
41     [1: cases (?:False); clear w2 H4 w1 H2 w H1; 
42         apply (q_lt_antisym init (start l1)); [assumption] rewrite < H5; assumption
43     |2: intros; rewrite > H6; clear H6; rewrite > H4; reflexivity;
44     |3: cases (?:False); apply (q_lt_antisym input (start l1)); [2: assumption]
45         apply (q_lt_trans ??? H3 H);]
46 |3: intro; cases H4; clear H4;   
47     cases (value l1 input); simplify; cases (q_cmp input (start l1)) in H4; whd in ⊢ (% → ?);
48     [1: intro; cases H8; clear H8; rewrite > H11; rewrite > H7; clear H11 H7;
49         simplify in ⊢ (? ? ? (? ? ? (? ? % ? ?)));
50         cut (\fst w1 = S (\fst w2)) as Key; [rewrite > Key; reflexivity;]
51         cut (\fst w2 = O); [2: clear H10;
52           symmetry; apply le_n_O_to_eq; rewrite > (sum_bases_O l1 (\fst w2)); [apply le_n]
53           apply (q_le_trans ??? H9); rewrite < H4; rewrite > q_d_x_x; 
54           apply q_eq_to_le; reflexivity;]
55         rewrite > Hcut; clear Hcut H10 H9; simplify in H5 H6;
56         cut (ⅆ[input,init] = Qpos w) as E; [2:
57           rewrite > H2; rewrite < H4; rewrite > q_d_sym; 
58           rewrite > q_d_noabs; [reflexivity] apply q_lt_to_le; assumption;]
59         cases (\fst w1) in H5 H6; intros;
60         [1: cases (?:False); clear H5; simplify in H6;
61             apply (q_lt_corefl ⅆ[input,init]);
62             rewrite > E in ⊢ (??%); rewrite < q_plus_OQ in ⊢ (??%);
63             rewrite > q_plus_sym; assumption;
64         |2: cases n in H5 H6; [intros; reflexivity] intros;
65             cases (?:False); clear H6; cases (bars l1) in H5; simplify; intros;
66             [apply (q_pos_OQ one);|apply (q_pos_OQ (\fst b));] 
67             apply (q_le_S ??? (sum_bases_ge_OQ (mk_q_f init ?) n1));[apply [];|3:apply l]
68             simplify in ⊢ (? (? (? % ?) ?) ?); rewrite < (q_plus_minus w);
69             apply q_le_minus_r; rewrite < q_minus_r; 
70             rewrite < E in ⊢ (??%); assumption;]
71     |2: intros; rewrite > H8; rewrite > H7; clear H8 H7;
72         simplify in H5 H6 ⊢ %; 
73         cases (\fst w1) in H5 H6; [intros; reflexivity]
74         cases (bars l1);
75         [1: intros; simplify; elim n [reflexivity] simplify; assumption;
76         |2: simplify; intros; cases (?:False);
77         
78 STOP            
79         
80 alias symbol "pi2" = "pair pi2".
81 alias symbol "pi1" = "pair pi1".
82 definition rebase_spec ≝ 
83  ∀l1,l2:q_f.∃p:q_f × q_f.
84    And4
85     (*len (bars (\fst p)) = len (bars (\snd p))*)
86     (start (\fst p) = start (\snd p))
87     (same_bases (\fst p) (\snd p))
88     (same_values l1 (\fst p)) 
89     (same_values l2 (\snd p)).
90
91 definition rebase_spec_simpl ≝ 
92  λstart.λl1,l2:list bar.λp:(list bar) × (list bar).
93    And3
94     (same_bases (mk_q_f start (\fst p)) (mk_q_f start (\snd p)))
95     (same_values (mk_q_f start l1) (mk_q_f start (\fst p))) 
96     (same_values (mk_q_f start l2) (mk_q_f start (\snd p))).
97
98 (* a local letin makes russell fail *)
99 definition cb0h : list bar → list bar ≝ 
100   λl.mk_list (λi.〈\fst (nth l ▭ i),OQ〉) (len l).
101
102 definition eject ≝
103   λP.λp:∃x:(list bar) × (list bar).P x.match p with [ex_introT p _ ⇒ p].
104 coercion eject.
105 definition inject ≝ λP.λp:(list bar) × (list bar).λh:P p. ex_introT ? P p h.
106 coercion inject with 0 1 nocomposites.
107         
108 definition rebase: rebase_spec.
109 intros 2 (f1 f2); cases f1 (s1 l1); cases f2 (s2 l2); clear f1 f2;
110 letin spec ≝ (
111   λs.λl1,l2.λm.λz.len l1 + len l2 < m → rebase_spec_simpl s l1 l2 z);
112 alias symbol "pi1" (instance 34) = "exT \fst".
113 alias symbol "pi1" (instance 21) = "exT \fst".
114 letin aux ≝ ( 
115 let rec aux (l1,l2:list bar) (n:nat) on n : (list bar) × (list bar) ≝
116 match n with
117 [ O ⇒ 〈 nil ? , nil ? 〉
118 | S m ⇒ 
119   match l1 with
120   [ nil ⇒ 〈cb0h l2, l2〉
121   | cons he1 tl1 ⇒
122      match l2 with
123      [ nil ⇒ 〈l1, cb0h l1〉
124      | cons he2 tl2 ⇒  
125          let base1 ≝ Qpos (\fst he1) in
126          let base2 ≝ Qpos (\fst he2) in
127          let height1 ≝ (\snd he1) in
128          let height2 ≝ (\snd he2) in
129          match q_cmp base1 base2 with
130          [ q_eq _ ⇒ 
131              let rc ≝ aux tl1 tl2 m in 
132              〈he1 :: \fst rc,he2 :: \snd rc〉 
133          | q_lt Hp ⇒
134              let rest ≝ base2 - base1 in
135              let rc ≝ aux tl1 (〈\fst (unpos rest ?),height2〉 :: tl2) m in
136              〈〈\fst he1,height1〉 :: \fst rc,〈\fst he1,height2〉 :: \snd rc〉 
137          | q_gt Hp ⇒ 
138              let rest ≝ base1 - base2 in
139              let rc ≝ aux (〈\fst (unpos rest ?),height1〉 :: tl1) tl2 m in
140              〈〈\fst he2,height1〉 :: \fst rc,〈\fst he2,height2〉 :: \snd rc〉
141 ]]]]
142 in aux : ∀l1,l2,m.∃z.∀s.spec s l1 l2 m z); unfold spec;
143 [9: clearbody aux; unfold spec in aux; clear spec;
144     cases (q_cmp s1 s2);
145     [1: cases (aux l1 l2 (S (len l1 + len l2)));
146         cases (H1 s1 (le_n ?)); clear H1;
147         exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s2 (\snd w)〉] split;
148         [1,2: assumption;
149         |3: intro; apply (H3 input);
150         |4: intro; rewrite > H in H4; 
151             rewrite > (H4 input); reflexivity;]
152     |2: letin l2' ≝ (〈\fst (unpos (s2-s1) ?),OQ〉::l2);[
153           apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
154           assumption]
155         cases (aux l1 l2' (S (len l1 + len l2')));
156         cases (H1 s1 (le_n ?)); clear H1 aux;
157         exists [apply 〈mk_q_f s1 (\fst w), mk_q_f s1 (\snd w)〉] split;
158         [1: reflexivity
159         |2: assumption;
160         |3: assumption;
161         |4: intro; rewrite < (H4 input); clear H3 H4 H2 w;
162             cases (value (mk_q_f s1 l2') input);
163             cases (q_cmp input (start (mk_q_f s1 l2'))) in H1;
164             whd in ⊢ (% → ?);
165             [1: intros; cases H2; clear H2; whd in ⊢ (??? %);
166                 cases (value (mk_q_f s2 l2) input);
167                 cases (q_cmp input (start (mk_q_f s2 l2))) in H2;
168                 whd in ⊢ (% → ?);
169                 [1: intros; cases H6; clear H6; change with (w1 = w);
170                           
171             (* TODO *) ]]    
172 |1,2: unfold rest; apply q_lt_minus; rewrite > q_plus_sym; rewrite > q_plus_OQ;
173       assumption;        
174 |3:(* TODO *)
175 |4:(* TODO *)
176 |5:(* TODO *)
177 |6:(* TODO *)
178 |7:(* TODO *)
179 |8: intros; cases (?:False); apply (not_le_Sn_O ? H1);]
180 qed.