X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fetc%2Fssta2%2Fssta_lift.etc;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fetc%2Fssta2%2Fssta_lift.etc;h=5398364e0c5c94f80a7374662a80e9f9ba76f891;hb=ff7754f834f937bfe2384c7703cf63f552885395;hp=0000000000000000000000000000000000000000;hpb=4720368dcf18593959c6d21484f62fb5b61f3d26;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/ssta2/ssta_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/ssta2/ssta_lift.etc new file mode 100644 index 000000000..5398364e0 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/ssta2/ssta_lift.etc @@ -0,0 +1,119 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/ldrop_ldrop.ma". +include "basic_2/static/ssta.ma". + +(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************) + +(* Properties on relocation *************************************************) + +lemma ssta_lift: ∀h,g,G,L1,T1,U1,l. ⦃G, L1⦄ ⊢ T1 •[h, g] ⦃l, U1⦄ → + ∀L2,d,e. ⇩[d, e] L2 ≡ L1 → ∀T2. ⇧[d, e] T1 ≡ T2 → + ∀U2. ⇧[d, e] U1 ≡ U2 → ⦃G, L2⦄ ⊢ T2 •[h, g] ⦃l, U2⦄. +#h #g #G #L1 #T1 #U1 #l #H elim H -G -L1 -T1 -U1 -l +[ #G #L1 #k #l #Hkl #L2 #d #e #HL21 #X1 #H1 #X2 #H2 + >(lift_inv_sort1 … H1) -X1 + >(lift_inv_sort1 … H2) -X2 /2 width=1/ +| #G #L1 #K1 #V1 #W1 #W #i #l #HLK1 #_ #HW1 #IHVW1 #L2 #d #e #HL21 #X #H #U2 #HWU2 + elim (lift_inv_lref1 … H) * #Hid #H destruct + [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // #W2 #HW12 #HWU2 + elim (ldrop_trans_le … HL21 … HLK1 ?) -L1 /2 width=2/ #X #HLK2 #H + elim (ldrop_inv_skip2 … H ?) -H /2 width=1/ -Hid #K2 #V2 #HK21 #HV12 #H destruct + /3 width=8/ + | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W // /2 width=1/ #HW1U2 + lapply (ldrop_trans_ge … HL21 … HLK1 ?) -L1 // -Hid /3 width=8/ + ] +| #G #L1 #K1 #W1 #V1 #W #i #l #HLK1 #_ #HW1 #IHWV1 #L2 #d #e #HL21 #X #H #U2 #HWU2 + elim (lift_inv_lref1 … H) * #Hid #H destruct + [ elim (lift_trans_ge … HW1 … HWU2 ?) -W // (lift_inv_sort2 … H) -X /3 width=3/ +| #G #L2 #K2 #V2 #W2 #W #i #l #HLK2 #HVW2 #HW2 #IHVW2 #L1 #d #e #HL21 #X #H + elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ] + [ elim (ldrop_conf_lt … HL21 … HLK2 ?) -L2 // #K1 #V1 #HLK1 #HK21 #HV12 + elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HVW1 #HW12 + elim (lift_trans_le … HW12 … HW2 ?) -W2 // >minus_plus minus_minus_m_m /2 width=1/ /3 width=6/ + | minus_plus minus_minus_m_m /2 width=1/ /3 width=6/ + |