]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_teqo_vector.ma
milestone update in basic_2, update in ground and static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpxs_teqo_vector.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 "static_2/syntax/teqo_simple_vector.ma".
16 include "static_2/relocation/lifts_vector.ma".
17 include "basic_2/rt_computation/cpxs_teqo.ma".
18
19 (* EXTENDED CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS *************)
20
21 (* Vector form of forward lemmas with outer equivalence for terms ***********)
22
23 lemma cpxs_fwd_sort_vector (G) (L):
24       ∀s,Vs,X2. ❪G,L❫ ⊢ ⒶVs.⋆s ⬈* X2 → ⒶVs.⋆s ⩳ X2.
25 #G #L #s #Vs elim Vs -Vs /2 width=4 by cpxs_fwd_sort/
26 #V #Vs #IHVs #X2 #H
27 elim (cpxs_inv_appl1 … H) -H *
28 [ -IHVs #V1 #T1 #_ #_ #H destruct /2 width=1 by teqo_pair/
29 | #p #W1 #T1 #HT1 #HU
30   lapply (IHVs … HT1) -IHVs -HT1 #HT1
31   elim (teqo_inv_applv_bind_simple … HT1) //
32 | #p #V1 #V2 #V3 #T1 #HV01 #HV12 #HT1 #HU
33   lapply (IHVs … HT1) -IHVs -HT1 #HT1
34   elim (teqo_inv_applv_bind_simple … HT1) //
35 ]
36 qed-.
37
38 (* Basic_2A1: was: cpxs_fwd_delta_vector *)
39 lemma cpxs_fwd_delta_drops_vector (I) (G) (L) (K):
40       ∀V1,i. ⇩[i] L ≘ K.ⓑ[I]V1 →
41       ∀V2. ⇧[↑i] V1 ≘ V2 →
42       ∀Vs,X2. ❪G,L❫ ⊢ ⒶVs.#i ⬈* X2 →
43       ∨∨ ⒶVs.#i ⩳ X2 | ❪G,L❫ ⊢ ⒶVs.V2 ⬈* X2.
44 #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs
45 elim Vs -Vs /2 width=5 by cpxs_fwd_delta_drops/
46 #V #Vs #IHVs #X2 #H -K -V1
47 elim (cpxs_inv_appl1 … H) -H *
48 [ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by teqo_pair, or_introl/
49 | #q #W0 #T0 #HT0 #HU
50   elim (IHVs … HT0) -IHVs -HT0 #HT0
51   [ elim (teqo_inv_applv_bind_simple … HT0) //
52   | @or_intror -i (**) (* explicit constructor *)
53     @(cpxs_trans … HU) -X2
54     @(cpxs_strap1 … (ⓐV.ⓛ[q]W0.T0)) /3 width=1 by cpxs_flat_dx, cpx_beta/
55   ]
56 | #q #V0 #V1 #V3 #T0 #HV0 #HV01 #HT0 #HU
57   elim (IHVs … HT0) -IHVs -HT0 #HT0
58   [ elim (teqo_inv_applv_bind_simple … HT0) //
59   | @or_intror -i (**) (* explicit constructor *)
60     @(cpxs_trans … HU) -X2
61     @(cpxs_strap1 … (ⓐV0.ⓓ[q]V3.T0)) /3 width=3 by cpxs_flat, cpx_theta/
62   ]
63 ]
64 qed-.
65
66 (* Basic_1: was just: pr3_iso_appls_beta *)
67 lemma cpxs_fwd_beta_vector (p) (G) (L):
68       ∀Vs,V,W,T,X2. ❪G,L❫ ⊢ ⒶVs.ⓐV.ⓛ[p]W.T ⬈* X2 →
69       ∨∨ ⒶVs.ⓐV.ⓛ[p]W. T ⩳ X2 | ❪G,L❫ ⊢ ⒶVs.ⓓ[p]ⓝW.V.T ⬈* X2.
70 #p #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_beta/
71 #V0 #Vs #IHVs #V #W #T #X2 #H
72 elim (cpxs_inv_appl1 … H) -H *
73 [ -IHVs #V1 #T1 #_ #_ #H destruct /2 width=1 by teqo_pair, or_introl/
74 | #q #W1 #T1 #HT1 #HU
75   elim (IHVs … HT1) -IHVs -HT1 #HT1
76   [ elim (teqo_inv_applv_bind_simple … HT1) //
77   | @or_intror (**) (* explicit constructor *)
78     @(cpxs_trans … HU) -X2
79     @(cpxs_strap1 … (ⓐV0.ⓛ[q]W1.T1)) /3 width=1 by cpxs_flat_dx, cpx_beta/
80   ]
81 | #q #V1 #V2 #V3 #T1 #HV01 #HV12 #HT1 #HU
82   elim (IHVs … HT1) -IHVs -HT1 #HT1
83   [ elim (teqo_inv_applv_bind_simple … HT1) //
84   | @or_intror (**) (* explicit constructor *)
85     @(cpxs_trans … HU) -X2
86     @(cpxs_strap1 … (ⓐV1.ⓓ[q]V3.T1)) /3 width=3 by cpxs_flat, cpx_theta/
87   ]
88 ]
89 qed-.
90
91 (* Basic_1: was just: pr3_iso_appls_abbr *)
92 lemma cpxs_fwd_theta_vector (G) (L):
93       ∀V1b,V2b. ⇧[1] V1b ≘ V2b →
94       ∀p,V,T,X2. ❪G,L❫ ⊢ ⒶV1b.ⓓ[p]V.T ⬈* X2 →
95       ∨∨ ⒶV1b.ⓓ[p]V.T ⩳ X2 | ❪G,L❫ ⊢ ⓓ[p]V.ⒶV2b.T ⬈* X2.
96 #G #L #V1b #V2b * -V1b -V2b /3 width=1 by or_intror/
97 #V1b #V2b #V1a #V2a #HV12a #HV12b #p
98 generalize in match HV12a; -HV12a
99 generalize in match V2a; -V2a
100 generalize in match V1a; -V1a
101 elim HV12b -V1b -V2b /2 width=1 by cpxs_fwd_theta/
102 #V1b #V2b #V1b #V2b #HV12b #_ #IHV12b #V1a #V2a #HV12a #V #T #X2 #H
103 elim (cpxs_inv_appl1 … H) -H *
104 [ -IHV12b -HV12a -HV12b #V0 #T0 #_ #_ #H destruct /2 width=1 by teqo_pair, or_introl/
105 | #q #W0 #T0 #HT0 #HU
106   elim (IHV12b … HV12b … HT0) -IHV12b -HT0 #HT0
107   [ -HV12a -HV12b -HU
108     elim (teqo_inv_pair1 … HT0) #V1 #T1 #H destruct
109   | @or_intror -V1b (**) (* explicit constructor *)
110     @(cpxs_trans … HU) -X2
111     elim (cpxs_inv_abbr1_dx … HT0) -HT0 *
112     [ -HV12a #V1 #T1 #_ #_ #H destruct
113     | -V1b #X #HT1 #H #H0 destruct
114       elim (lifts_inv_bind1 … H) -H #W1 #T1 #HW01 #HT01 #H destruct
115       @(cpxs_trans … (+ⓓV.ⓐV2a.ⓛ[q]W1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2b
116       @(cpxs_strap2 … (ⓐV1a.ⓛ[q]W0.T0))
117       /4 width=7 by cpxs_beta_dx, cpx_zeta, lifts_bind, lifts_flat/
118     ]
119   ]
120 | #q #V0a #Va #V0 #T0 #HV10a #HV0a #HT0 #HU
121   elim (IHV12b … HV12b … HT0) -HV12b -IHV12b -HT0 #HT0
122   [ -HV12a -HV10a -HV0a -HU
123     elim (teqo_inv_pair1 … HT0) #V1 #T1 #H destruct
124   | @or_intror -V1b -V1b (**) (* explicit constructor *)
125     @(cpxs_trans … HU) -X2
126     elim (cpxs_inv_abbr1_dx … HT0) -HT0 *
127     [ #V1 #T1 #HV1 #HT1 #H destruct
128       lapply (cpxs_lifts_bi … HV10a (Ⓣ) … (L.ⓓV) … HV12a … HV0a) -V1a -V0a /3 width=1 by drops_refl, drops_drop/ #HV2a
129       @(cpxs_trans … (ⓓ[p]V.ⓐV2a.T1)) /3 width=1 by cpxs_bind, cpxs_pair_sn, cpxs_flat_dx, cpxs_bind_dx/
130     | #X #HT1 #H #H0 destruct
131       elim (lifts_inv_bind1 … H) -H #V1 #T1 #HW01 #HT01 #H destruct
132       lapply (cpxs_lifts_bi … HV10a (Ⓣ) … (L.ⓓV0) … HV12a … HV0a) -V0a /3 width=1 by drops_refl, drops_drop/ #HV2a
133       @(cpxs_trans … (+ⓓV.ⓐV2a.ⓓ[q]V1.T1)) [ /3 width=1 by cpxs_flat_dx, cpxs_bind_dx/ ] -T -V2b -V2b
134       @(cpxs_strap2 … (ⓐV1a.ⓓ[q]V0.T0)) [ /4 width=7 by cpx_zeta, lifts_bind, lifts_flat/ ] -V -V1 -T1
135       @(cpxs_strap2 … (ⓓ[q]V0.ⓐV2a.T0)) /3 width=3 by cpxs_pair_sn, cpxs_bind_dx, cpx_theta/
136     ]
137   ]
138 ]
139 qed-.
140
141 (* Basic_1: was just: pr3_iso_appls_cast *)
142 lemma cpxs_fwd_cast_vector (G) (L):
143       ∀Vs,W,T,X2. ❪G,L❫ ⊢ ⒶVs.ⓝW.T ⬈* X2 →
144       ∨∨ ⒶVs. ⓝW. T ⩳ X2
145        | ❪G,L❫ ⊢ ⒶVs.T ⬈* X2
146        | ❪G,L❫ ⊢ ⒶVs.W ⬈* X2.
147 #G #L #Vs elim Vs -Vs /2 width=1 by cpxs_fwd_cast/
148 #V #Vs #IHVs #W #T #X2 #H
149 elim (cpxs_inv_appl1 … H) -H *
150 [ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by teqo_pair, or3_intro0/
151 | #q #W0 #T0 #HT0 #HU elim (IHVs … HT0) -IHVs -HT0 #HT0
152   [ elim (teqo_inv_applv_bind_simple … HT0) //
153   | @or3_intro1 -W (**) (* explicit constructor *)
154     @(cpxs_trans … HU) -X2
155     @(cpxs_strap1 … (ⓐV.ⓛ[q]W0.T0)) /2 width=1 by cpxs_flat_dx, cpx_beta/
156   | @or3_intro2 -T (**) (* explicit constructor *)
157     @(cpxs_trans … HU) -X2
158     @(cpxs_strap1 … (ⓐV.ⓛ[q]W0.T0)) /2 width=1 by cpxs_flat_dx, cpx_beta/
159   ]
160 | #q #V0 #V1 #V2 #T0 #HV0 #HV01 #HT0 #HU
161   elim (IHVs … HT0) -IHVs -HT0 #HT0
162   [ elim (teqo_inv_applv_bind_simple … HT0) //
163   | @or3_intro1 -W (**) (* explicit constructor *)
164     @(cpxs_trans … HU) -X2
165     @(cpxs_strap1 … (ⓐV0.ⓓ[q]V2.T0)) /2 width=3 by cpxs_flat, cpx_theta/
166   | @or3_intro2 -T (**) (* explicit constructor *)
167     @(cpxs_trans … HU) -X2
168     @(cpxs_strap1 … (ⓐV0.ⓓ[q]V2.T0)) /2 width=3 by cpxs_flat, cpx_theta/
169   ]
170 ]
171 qed-.
172
173 (* Basic_1: was just: nf2_iso_appls_lref *)
174 lemma cpxs_fwd_cnx_vector (G) (L):
175       ∀T. 𝐒❪T❫ → ❪G,L❫ ⊢ ⬈𝐍 T →
176       ∀Vs,X2. ❪G,L❫ ⊢ ⒶVs.T ⬈* X2 → ⒶVs.T ⩳ X2.
177 #G #L #T #H1T #H2T #Vs elim Vs -Vs [ @(cpxs_fwd_cnx … H2T) ] (**) (* /2 width=3 by cpxs_fwd_cnx/ does not work *)
178 #V #Vs #IHVs #X2 #H
179 elim (cpxs_inv_appl1 … H) -H *
180 [ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1 by teqo_pair/
181 | #p #W0 #T0 #HT0 #HU
182   lapply (IHVs … HT0) -IHVs -HT0 #HT0
183   elim (teqo_inv_applv_bind_simple … HT0) //
184 | #p #V1 #V2 #V0 #T0 #HV1 #HV12 #HT0 #HU
185   lapply (IHVs … HT0) -IHVs -HT0 #HT0
186   elim (teqo_inv_applv_bind_simple … HT0) //
187 ]
188 qed-.