(**************************************************************************) (* ___ *) (* ||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/reducibility/tpr_tpr.ma". include "basic_2/reducibility/cpr.ma". (* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) (* Advanced properties ******************************************************) lemma cpr_bind_sn: ∀a,I,L,V1,V2,T1,T2. L ⊢ V1 ➡ V2 → T1 ➡ T2 → L ⊢ ⓑ{a,I} V1. T1 ➡ ⓑ{a,I} V2. T2. #a #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 #HT12 @ex2_intro [2: @(tpr_delta … HV1 HT12) | skip ] /2 width=3/ (* /3 width=5/ is too slow *) qed. (* Basic_1: was only: pr2_gen_cbind *) lemma cpr_bind_dx: ∀a,I,L,V1,V2,T1,T2. V1 ➡ V2 → L. ⓑ{I} V2 ⊢ T1 ➡ T2 → L ⊢ ⓑ{a,I} V1. T1 ➡ ⓑ{a,I} V2. T2. #a #I #L #V1 #V2 #T1 #T2 #HV12 * #T #HT1 normalize #HT2 elim (tpss_split_up … HT2 1 ? ?) -HT2 // #T0