(**************************************************************************) (* ___ *) (* ||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/cpys_cpys.ma". (* SN EXTENDED MULTIPLE SUBSTITUTION FOR LOCAL ENVIRONMENTS *****************) (* Main properties **********************************************************) theorem lpys_trans: ∀G. Transitive … (lpys G). /3 width=5 by lpx_sn_trans, cpys_trans_lpys/ qed-. (* Advanced forward lemmas **************************************************) lemma cpys_fwd_shift1_ext: ∀G,L1,L,T1,T. ⦃G, L⦄ ⊢ L1 @@ T1 ▶*× T → ∃∃L2,T2. ⦃G, L @@ L1⦄ ⊢ ▶*× L @@ L2 & ⦃G, L @@ L1⦄ ⊢ T1 ▶*× T2 & T = L2 @@ T2. #G #L1 @(lenv_ind_dx … L1) -L1 [ #L #T1 #T #HT1 @ex3_2_intro [3: // |4,5: // |1,2: skip ] (**) (* auto does not work *) | #I #L1 #V1 #IH #L #T1 #T >shift_append_assoc #H