]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc_vector.ma
We are decapitalizing the contributions' names ...
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / cprs_tstc_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 "Basic_2/grammar/tstc_vector.ma".
16 include "Basic_2/substitution/lift_vector.ma".
17 include "Basic_2/computation/cprs_tstc.ma".
18
19 (* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************)
20
21 (* Vector form of forward lemmas involving same top term constructor ********)
22 (*
23 lemma cpr_fwd_beta_vector: ∀L,V,W,T,U,Vs. L ⊢ ⒶVs. ⓐV. ⓛW. T ➡ U →
24                            ⒶVs. ⓐV. ⓛW. T ≃ U ∨ L ⊢ ⒶVs. ⓓV. T ➡* U.
25 #L #V #W #T #U * /2 width=1 by cpr_fwd_beta/
26 #V0 #Vs #H
27 elim (cpr_inv_appl1_simple … H ?) -H
28 [ #V1 #T1 #_ #_ #H destruct /2 width=1/
29 | elim Vs -Vs //
30 ]
31 qed-.
32
33 lemma cpr_fwd_theta_vector: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s →
34                             ∀V,T,U. L ⊢ ⒶV1s. ⓓV. T ➡ U →
35                             ⒶV1s. ⓓV. T ≃ U ∨ L ⊢ ⓓV. ⒶV2s. T ➡* U.
36 #L #V1s #V2s * -V1s -V2s /3 width=1/
37 #V1s #V2s #V1a #V2a #HV12a * -V1s -V2s /2 width=1 by cpr_fwd_theta/ -HV12a
38 #V1s #V2s #V1b #V2b #_ #_ #V #U #T #H
39 elim (cpr_inv_appl1_simple … H ?) -H //
40 #V0 #T0 #_ #_ #H destruct /2 width=1/
41 qed-.
42 *)
43
44 (* Basic_1: was: pr3_iso_appls_cast *)
45 lemma cprs_fwd_tau_vector: ∀L,Vs,W,T,U. L ⊢ ⒶVs. ⓣW. T ➡* U →
46                            ⒶVs. ⓣW. T ≃ U ∨ L ⊢ ⒶVs. T ➡* U.
47 #L #Vs elim Vs -Vs /2 width=1 by cprs_fwd_tau/
48 #V #Vs #IHVs #W #T #U #H
49 elim (cprs_inv_appl1 … H) -H *
50 [ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1/
51 | #V0 #W0 #T0 #HV0 #HT0 #HU
52   elim (IHVs … HT0) -IHVs -HT0 #HT0
53   [ elim (tstc_inv_bind_appls_simple … HT0 ?) //
54   | @or_intror
55     @(cprs_trans … HU) -HU
56     @(cprs_strap1 … (ⓐV0.ⓛW0.T0)) /2 width=1/ -HV0 -HT0 /3 width=1/
57   ]
58 | #V0 #V1 #V2 #T0 #HV0 #HV01 #HT0 #HU
59   elim (IHVs … HT0) -IHVs -HT0 #HT0
60   [ elim (tstc_inv_bind_appls_simple … HT0 ?) //
61   | @or_intror
62     @(cprs_trans … HU) -HU
63     @(cprs_strap1 … (ⓐV0.ⓓV2.T0)) /2 width=1/ -HV0 -HT0 /3 width=3/
64 ]
65 qed-.
66
67 axiom cprs_fwd_theta_vector: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s →
68                             ∀V,T,U. L ⊢ ⒶV1s. ⓓV. T ➡* U →
69                             ⒶV1s. ⓓV. T ≃ U ∨ L ⊢ ⓓV. ⒶV2s. T ➡* U.