]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/etc/cpys2/lpx_cpys.etc
milestone update in ground_2 and basic_2A
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / etc / cpys2 / lpx_cpys.etc
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/relocation/cpy_lift.ma".
16 include "basic_2/substitution/cpys.ma".
17 include "basic_2/reduction/lpx_ldrop.ma".
18
19 (* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************)
20
21 (* Properties on context-sensitive extended substitution for terms **********)
22
23 lemma cpx_cpy_trans_lpx: ∀h,g,G,L1,T1,T. ⦃G, L1⦄ ⊢ T1 ➡[h, g] T →
24                          ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 →
25                          ∀T2,d,e. ⦃G, L2⦄ ⊢ T ▶[d, e] T2 → ⦃G, L1⦄ ⊢ T1 ➡[h, g] T2.
26 #h #g #G #L1 #T1 #T #H elim H -G -L1 -T1 -T
27 [ #J #G #L1 #L2 #HL12 #T2 #d #e #H elim (cpy_inv_atom1 … H) -H //
28   * #I #K2 #V2 #i #_ #_ #HLK2 #HVT2 #H destruct
29   elim (lpx_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
30   elim (lpx_inv_pair2 … H) -H #K1 #V1 #_ #HV12 #H destruct
31   /2 width=7 by cpx_delta/
32 | #G #L1 #k #l #Hkl #L2 #_ #X #d #e #H >(cpy_inv_sort1 … H) -X /2 width=2 by cpx_sort/
33 | #I #G #L1 #K1 #V1 #V #T #i #HLK1 #_ #HVT #IHV1 #L2 #HL12 #T2 #d #e #HT2
34   elim (lpx_ldrop_conf … HLK1 … HL12) -HL12 #X #H #HLK2
35   elim (lpx_inv_pair1 … H) -H #K2 #V0 #HK12 #_ #H destruct
36   lapply (ldrop_fwd_drop2 … HLK2) -V0 #HLK2
37   lapply (cpy_weak … HT2 0 (∞) ? ?) -HT2 // #HT2
38   elim (cpy_inv_lift1_be … HT2 … HLK2 … HVT) -HT2 -HLK2 -HVT
39   /3 width=7 by cpx_delta/
40 | #a #I #G #L1 #V1 #V #T1 #T #HV1 #_ #IHV1 #IHT1 #L2 #HL12 #X #d #e #H elim (cpy_inv_bind1 … H) -H
41   #V2 #T2 #HV2 #HT2 #H destruct /4 width=5 by lpx_pair, cpx_bind/
42 | #I #G #L1 #V1 #V #T1 #T #_ #_ #IHV1 #IHT1 #L2 #HL12 #X #d #e #H elim (cpy_inv_flat1 … H) -H
43   #V2 #T2 #HV2 #HT2 #H destruct /3 width=5 by cpx_flat/
44 | #G #L1 #V1 #U1 #U #T #_ #HTU #IHU1 #L2 #HL12 #T2 #d #e #HT2
45   lapply (cpy_weak … HT2 0 (∞) ? ?) -HT2 // #HT2
46   elim (lift_total T2 0 1) #U2 #HTU2
47   lapply (cpy_lift_be … HT2 (L2.ⓓV1) … (Ⓕ) … HTU … HTU2 ? ?) -T
48   /4 width=5 by cpx_zeta, lpx_pair, ldrop_drop/
49 | /3 width=5 by cpx_tau/
50 | /3 width=5 by cpx_ti/
51 | #a #G #L1 #V1 #V #W1 #W #T1 #T #HV1 #HW1 #_ #IHV1 #IHW1 #IHT1 #L2 #HL12 #X #d #e #HX
52   elim (cpy_inv_bind1 … HX) -HX #Y #T2 #HY #HT2 #H destruct
53   elim (cpy_inv_flat1 … HY) -HY #W2 #V2 #HW2 #HV2 #H destruct
54   /5 width=11 by lpx_pair, cpx_beta, lsuby_cpy_trans, lsuby_succ/
55 | #a #G #L1 #V1 #V #U #W1 #W #T1 #T #_ #HVU #HW1 #_ #IHV1 #IHW1 #IHT1 #L2 #HL12 #X #d #e #HX
56   elim (cpy_inv_bind1 … HX) -HX #W2 #Y #HW2 #HY #H destruct
57   elim (cpy_inv_flat1 … HY) -HY #U2 #T2 #HU2 #HT2 #H destruct
58   lapply (cpy_weak … HU2 0 (∞) ? ?) -HU2 // #HU2
59   elim (cpy_inv_lift1_be … HU2 L2 … HVU) -U
60   /4 width=7 by lpx_pair, cpx_theta, ldrop_drop/
61 ]
62 qed-.
63
64 lemma cpx_cpys_trans_lpx: ∀h,g,G,L1,T1,T. ⦃G, L1⦄ ⊢ T1 ➡[h, g] T →
65                           ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 →
66                           ∀T2,d,e. ⦃G, L2⦄ ⊢ T ▶*[d, e] T2 → ⦃G, L1⦄ ⊢ T1 ➡[h, g] T2.
67 #h #g #G #L1 #T1 #T #HT1 #L2 #HL12 #T2 #d #e #H @(cpys_ind … H) -T2
68 /2 width=7 by cpx_cpy_trans_lpx/
69 qed-.