X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Funfold%2Ftpss.ma;h=93916208b649a6d301c53ab4ef1e56f9d6417871;hb=9c09a0b1f8801e40612eef429b82fc6dbae01b85;hp=26dd58ee6b04cb30349b86fb79c1ce6719a867de;hpb=38a7664fd355705596cb63cac87779688790fcb1;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma index 26dd58ee6..93916208b 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/unfold/tpss.ma @@ -170,3 +170,13 @@ lemma tpss_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → #{T1} ≤ #{T2}. lapply (tps_fwd_tw … HT2) -HT2 #HT2 @(transitive_le … IHT1) // qed-. + +lemma tpss_fwd_shift1: ∀L,L1,T1,T,d,e. L ⊢ L1 @@ T1 ▶*[d, e] T → + ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. +#L #L1 #T1 #T #d #e #H @(tpss_ind … H) -T +[ /2 width=4/ +| #T #X #_ #H0 * #L0 #T0 #HL10 #H destruct + elim (tps_fwd_shift1 … H0) -H0 #L2 #T2 #HL02 #H destruct /2 width=4/ +] +qed-. + \ No newline at end of file