X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fmultiple%2Flifts_lift.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fmultiple%2Flifts_lift.ma;h=efa5038d27e580b5510ca78586bbe9ae4e8ad225;hb=598a5c56535a8339f6533227ab580aff64e2d41c;hp=0000000000000000000000000000000000000000;hpb=c671743a83bbc7fff114e3e3694f628c0ec6685b;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/lifts_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/lifts_lift.ma new file mode 100644 index 000000000..efa5038d2 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/lifts_lift.ma @@ -0,0 +1,59 @@ +(**************************************************************************) +(* ___ *) +(* ||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/substitution/lift_lift.ma". +include "basic_2/multiple/gr2_minus.ma". +include "basic_2/multiple/lifts.ma". + +(* GENERIC TERM RELOCATION **************************************************) + +(* Properties concerning basic term relocation ******************************) + +(* Basic_1: was: lift1_xhg (right to left) *) +lemma lifts_lift_trans_le: ∀T1,T,des. ⇧*[des] T1 ≡ T → ∀T2. ⇧[0, 1] T ≡ T2 → + ∃∃T0. ⇧[0, 1] T1 ≡ T0 & ⇧*[des + 1] T0 ≡ T2. +#T1 #T #des #H elim H -T1 -T -des +[ /2 width=3/ +| #T1 #T3 #T #des #d #e #HT13 #_ #IHT13 #T2 #HT2 + elim (IHT13 … HT2) -T #T #HT3 #HT2 + elim (lift_trans_le … HT13 … HT3) -T3 // /3 width=5/ +] +qed-. + +(* Basic_1: was: lift1_free (right to left) *) +lemma lifts_lift_trans: ∀des,i,i0. @⦃i, des⦄ ≡ i0 → + ∀des0. des + 1 ▭ i + 1 ≡ des0 + 1 → + ∀T1,T0. ⇧*[des0] T1 ≡ T0 → + ∀T2. ⇧[O, i0 + 1] T0 ≡ T2 → + ∃∃T. ⇧[0, i + 1] T1 ≡ T & ⇧*[des] T ≡ T2. +#des elim des -des normalize +[ #i #x #H1 #des0 #H2 #T1 #T0 #HT10 #T2 + <(at_inv_nil … H1) -x #HT02 + lapply (minuss_inv_nil1 … H2) -H2 #H + >(pluss_inv_nil2 … H) in HT10; -des0 #H + >(lifts_inv_nil … H) -T1 /2 width=3/ +| #d #e #des #IHdes #i #i0 #H1 #des0 #H2 #T1 #T0 #HT10 #T2 #HT02 + elim (at_inv_cons … H1) -H1 * #Hid #Hi0 + [ elim (minuss_inv_cons1_lt … H2) -H2 [2: /2 width=1/ ] #des1 #Hdes1 minus_plus #HT1 #HT0 + elim (IHdes … Hi0 … Hdes1 … HT0 … HT02) -IHdes -Hi0 -Hdes1 -T0 #T0 #HT0 #HT02 + elim (lift_trans_le … HT1 … HT0) -T /2 width=1/ #T #HT1 commutative_plus in Hi0; #Hi0 + lapply (minuss_inv_cons1_ge … H2 ?) -H2 [ /2 width=1/ ]