]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/Basic_2/unfold/tpss_ltps.ma
- the relocation properties of cpr are closed!
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / unfold / tpss_ltps.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/substitution/ltps_tps.ma".
16 include "Basic_2/unfold/tpss_tpss.ma".
17
18 (* PARTIAL UNFOLD ON TERMS **************************************************)
19
20 (* Properties concerning parallel substitution on local environments ********)
21
22 lemma ltps_tpss_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2.
23                          d1 + e1 ≤ d2 → L0 [d1, e1] ≫ L1 →
24                          L0 ⊢ T2 [d2, e2] ≫* U2 → L1 ⊢ T2 [d2, e2] ≫* U2.
25 #L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HL01 #H @(tpss_ind … H) -U2 //
26 #U #U2 #_ #HU2 #IHU
27 lapply (ltps_tps_conf_ge … HU2 … HL01 ?) -HU2 HL01 /2/
28 qed.
29
30 lemma ltps_tpss_conf: ∀L0,L1,T2,U2,d1,e1,d2,e2.
31                       L0 [d1, e1] ≫ L1 → L0 ⊢ T2 [d2, e2] ≫* U2 →
32                       ∃∃T. L1 ⊢ T2 [d2, e2] ≫* T & L1 ⊢ U2 [d1, e1] ≫* T.
33 #L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #HL01 #H @(tpss_ind … H) -U2
34 [ /3/
35 | #U #U2 #_ #HU2 * #T #HT2 #HUT
36   elim (ltps_tps_conf … HU2 … HL01) -HU2 HL01 #W #HUW #HU2W
37   elim (tpss_strip_eq … HUT … HUW) -U
38   /3 width=5 by ex2_1_intro, step, tpss_strap/ (**) (* just /3 width=5/ is too slow *)
39 ]
40 qed.
41
42 lemma ltps_tpss_trans_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2.
43                           d1 + e1 ≤ d2 → L1 [d1, e1] ≫ L0 →
44                           L0 ⊢ T2 [d2, e2] ≫* U2 → L1 ⊢ T2 [d2, e2] ≫* U2.
45 #L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HL10 #H @(tpss_ind … H) -U2 //
46 #U #U2 #_ #HU2 #IHU
47 lapply (ltps_tps_trans_ge … HU2 … HL10 ?) -HU2 HL10 /2/
48 qed.
49
50 lemma ltps_tpss_trans_down: ∀L0,L1,T2,U2,d1,e1,d2,e2. d2 + e2 ≤ d1 →
51                             L1 [d1, e1] ≫ L0 → L0 ⊢ T2 [d2, e2] ≫* U2 →
52                             ∃∃T. L1 ⊢ T2 [d2, e2] ≫* T & L0 ⊢ T [d1, e1] ≫* U2.
53 #L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde2d1 #HL10 #H @(tpss_ind … H) -U2
54 [ /3/
55 | #U #U2 #_ #HU2 * #T #HT2 #HTU
56   elim (tpss_strap1_down … HTU … HU2 ?) -U // #U #HTU #HU2
57   elim (ltps_tps_trans … HTU … HL10) -HTU HL10 #W #HTW #HWU
58   @(ex2_1_intro … W) /2/ (**) (* /3 width=5/ does not work as in ltps_tpss_conf *)
59 ]
60 qed.
61
62 fact ltps_tps_trans_eq_aux: ∀Y1,X2,L1,T2,U2,d,e.
63                             L1 ⊢ T2 [d, e] ≫ U2 → ∀L0. L0 [d, e] ≫ L1 →
64                             Y1 = L1 → X2 = T2 → L0 ⊢ T2 [d, e] ≫* U2.
65 #Y1 #X2 @(cw_wf_ind … Y1 X2) -Y1 X2 #Y1 #X2 #IH
66 #L1 #T2 #U2 #d #e * -L1 T2 U2 d e
67 [ //
68 | #L1 #K1 #V1 #W1 #i #d #e #Hdi #Hide #HLK1 #HVW1 #L0 #HL10 #H1 #H2 destruct -Y1 X2;
69   lapply (ldrop_fwd_lw … HLK1) normalize #H1
70   elim (ltps_ldrop_trans_be … HL10 … HLK1 ? ?) -HL10 HLK1 [2,3: /2/ ] #X #H #HLK0
71   elim (ltps_inv_tps22 … H ?) -H [2: /2/ ] #K0 #V0 #HK01 #HV01 #H destruct -X;
72   lapply (tps_fwd_tw … HV01) #H2
73   lapply (transitive_le (#[K1] + #[V0]) … H1) -H1 [ /2/ ] -H2 #H
74   lapply (IH … HV01 … HK01 ? ?) -IH HV01 HK01 [1,3: // |2,4: skip | normalize /2/ | /3 width=6/ ]
75 | #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #L0 #HL0 #H1 #H2 destruct -Y1 X2;
76   lapply (tps_lsubs_conf … HT12 (L. 𝕓{I} V1) ?) -HT12 /2/ #HT12
77   lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3,5: normalize // |2,4: skip ] #HV12
78   lapply (IH … HT12 (L0. 𝕓{I} V1) ? ? ?) -IH HT12 [1,3,5: /2/ |2,4: skip | normalize // ] -HL0 #HT12
79   lapply (tpss_lsubs_conf … HT12 (L0. 𝕓{I} V2) ?) -HT12 /2/
80 | #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #L0 #HL0 #H1 #H2 destruct -Y1 X2;
81   lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3,5: normalize // |2,4: skip ]
82   lapply (IH … HT12 … HL0 ? ?) -IH HT12 [1,3,5: normalize // |2,4: skip ] -HL0 /2/
83 ]
84 qed.
85
86 lemma ltps_tps_trans_eq: ∀L1,T2,U2,d,e. L1 ⊢ T2 [d, e] ≫ U2 →
87                          ∀L0. L0 [d, e] ≫ L1 → L0 ⊢ T2 [d, e] ≫* U2.
88 /2 width=5/ qed.
89
90 lemma ltps_tpss_trans_eq: ∀L0,L1,T2,U2,d,e. L0 [d, e] ≫ L1 →
91                           L1 ⊢ T2 [d, e] ≫* U2 → L0 ⊢ T2 [d, e] ≫* U2.
92 #L0 #L1 #T2 #U2 #d #e #HL01 #H @(tpss_ind … H) -U2 //
93 #U #U2 #_ #HU2 #IHU @(tpss_trans_eq … IHU) /2/
94 qed.