]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/unfold/thin_delift.ma
- lambda_delta: bug fix in static type assignment
[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/tpss_alt.ma".
16 include "basic_2/unfold/ltpss_ltpss.ma".
17 include "basic_2/unfold/delift_alt.ma".
18 include "basic_2/unfold/thin.ma".
19
20 (* DELIFT ON LOCAL ENVIRONMENTS *********************************************)
21
22 (* Properties on deflift on terms *******************************************)
23
24 lemma thin_delift1: ∀L1,L2,d,e. L1 [d, e] ≡ L2 → ∀V1,V2. L1 ⊢ V1 [d, e] ≡ V2 →
25                     ∀I. L1.ⓑ{I}V1 [d + 1, e] ≡ L2.ⓑ{I}V2.
26 #L1 #L2 #d #e * #L #HL1 #HL2 #V1 #V2 * #V #HV1 #HV2 #I
27 elim (ltpss_tpss_conf … HV1 … HL1) -HV1 #V0 #HV10 #HV0
28 elim (tpss_inv_lift1_be … HV0 … HL2 … HV2 ? ?) -HV0 // <minus_n_n #X #H1 #H2
29 lapply (tpss_inv_refl_O2 … H1) -H1 #H destruct
30 lapply (lift_mono … H2 … HV2) -H2 #H destruct /3 width=5/
31 qed.
32
33 axiom delift_tpss_conf_be: ∀L,U1,U2,d,e. L ⊢ U1 [d, e] ▶* U2 →
34                            ∀T1,dd,ee. L ⊢ U1 [dd, ee] ≡ T1 → ∀K. L [dd, ee] ≡ K →
35                            d ≤ dd → dd + ee ≤ d + e →
36                            ∃∃T2. K ⊢ T1 [dd - d, e - ee] ▶* T2 & L ⊢ U2 [dd, ee] ≡ T2.
37 (*
38 #L #U1 #U2 #d #e #H @(tpss_ind_alt … H) -L -U1 -U2 -d -e
39 [ #L * #i #d #e #X #dd #ee #H
40   [ >(delift_inv_sort1 … H) -X /2 width=3/
41   | elim (delift_inv_lref1 … H) -H * [1,3: /3 width=3/ | /3 width=6/ ]
42   | >(delift_inv_gref1 … H) -X /2 width=3/
43   ]
44 | #L #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #_ #HVW2 #IHV12 #T1 #dd #ee #H #K2 #HLK2 #Hdd #Hddee
45   lapply 
46     
47     @(ex2_1_intro … X) // /2 width=6/
48 *)