X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fsubstitution%2Ftps.ma;h=7e35dca7ca5e31ae427ae32bbe2ab50c19fcec83;hb=f6464ba2cffc9936b4d8285604786cd91531e0d0;hp=db116f77e63b526eb2dc408ec50dfb645a0b9327;hpb=abed266f42c25bf77c5c4bfbff0450abe9680e7a;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma index db116f77e..7e35dca7c 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/tps.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "basic_2/grammar/lenv_top.ma". include "basic_2/substitution/ldrop.ma". (* PARALLEL SUBSTITUTION ON TERMS *******************************************) @@ -250,6 +251,17 @@ lemma tps_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → #{T1} ≤ #{T2}. /3 by monotonic_le_plus_l, le_plus/ (**) (* just /3 width=1/ is too slow *) qed-. +lemma tps_fwd_shift1: ∀L1,L,T1,T,d,e. L ⊢ L1 @@ T1 ▶ [d, e] T → + ∃∃L2,T2. L1 𝟙 L2 & T = L2 @@ T2. +#L1 @(lenv_ind_dx … L1) -L1 +[ #L #T1 #T #d #e #_ @ex2_2_intro [3: // |4: // |1,2: skip ] (**) (* /2 width=4/ does not work *) +| #I #L1 #V1 #IH #L #T1 #T #d #e >shift_append_assoc #H + elim (tps_inv_bind1 … H) -H #V2 #T2 #_ #HT12 #H destruct + elim (IH … HT12) -IH -L -T1 -d -e #L2 #T #HL12 #H destruct + @(ex2_2_intro … (⋆.ⓑ{I}V2@@L2)) /2 width=4/ /3 width=2/ +] +qed-. + (* Basic_1: removed theorems 25: subst0_gen_sort subst0_gen_lref subst0_gen_head subst0_gen_lift_lt subst0_gen_lift_false subst0_gen_lift_ge subst0_refl subst0_trans