]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/unfold/thin_delift.ma
- more properties on lifting, slicing, delifting and thinning
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / unfold / thin_delift.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/unfold/delift_tpss.ma".
16 include "basic_2/unfold/delift_ltpss.ma".
17 include "basic_2/unfold/thin.ma".
18
19 (* DELIFT ON LOCAL ENVIRONMENTS *********************************************)
20
21 (* Properties on inverse term relocation ************************************)
22
23 lemma thin_delift1: ∀L1,L2,d,e. L1 [d, e] ≡ L2 → ∀V1,V2. L1 ⊢ V1 [d, e] ≡ V2 →
24                     ∀I. L1.ⓑ{I}V1 [d + 1, e] ≡ L2.ⓑ{I}V2.
25 #L1 #L2 #d #e * #L #HL1 #HL2 #V1 #V2 * #V #HV1 #HV2 #I
26 elim (ltpss_tpss_conf … HV1 … HL1) -HV1 #V0 #HV10 #HV0
27 elim (tpss_inv_lift1_be … HV0 … HL2 … HV2 ? ?) -HV0 // <minus_n_n #X #H1 #H2
28 lapply (tpss_inv_refl_O2 … H1) -H1 #H destruct
29 lapply (lift_mono … H2 … HV2) -H2 #H destruct /3 width=5/
30 qed.
31
32 lemma thin_delift_tpss_conf_be: ∀L,U1,U2,d,e. L ⊢ U1 [d, e] ▶* U2 →
33                                 ∀T1,dd,ee. L ⊢ U1 [dd, ee] ≡ T1 →
34                                 ∀K. L [dd, ee] ≡ K → d ≤ dd → dd + ee ≤ d + e →
35                                 ∃∃T2. K ⊢ T1 [d, e - ee] ▶* T2 &
36                                       L ⊢ U2 [dd, ee] ≡ T2.
37 #L #U1 #U2 #d #e #HU12 #T1 #dd #ee #HUT1 #K * #Y #HLY #HYK #Hdd #Hddee
38 lapply (delift_ltpss_conf_eq … HUT1 … HLY) -HUT1 #HUT1
39 elim (ltpss_tpss_conf … HU12 … HLY) -HU12 #U #HU1 #HU2
40 elim (delift_tpss_conf_be … HU1 … HUT1 … HYK ? ?) -HU1 -HUT1 // -Hdd -Hddee #T #HT1 #HUT
41 lapply (tpss_delift_trans_eq … HU2 … HUT) -U #HU2T
42 lapply (ltpss_delift_trans_eq … HLY … HU2T) -Y /2 width=3/
43 qed.
44
45 lemma thin_delift_tps_conf_be: ∀L,U1,U2,d,e. L ⊢ U1 [d, e] ▶ U2 →
46                                ∀T1,dd,ee. L ⊢ U1 [dd, ee] ≡ T1 →
47                                ∀K. L [dd, ee] ≡ K → d ≤ dd → dd + ee ≤ d + e →
48                                ∃∃T2. K ⊢ T1 [d, e - ee] ▶* T2 &
49                                      L ⊢ U2 [dd, ee] ≡ T2.
50 /3 width=3/ qed.