X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fcprs_delift.ma;fp=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fcprs_delift.ma;h=65488a7347a09d4c68aacdd6adc1d97cee17af94;hb=636c25914e83819c2f529edc891a7eb899499a97;hp=0000000000000000000000000000000000000000;hpb=4400d6a1bf09682b3c679edf0e1775117fe059c2;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_delift.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_delift.ma new file mode 100644 index 000000000..65488a734 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_delift.ma @@ -0,0 +1,30 @@ +(**************************************************************************) +(* ___ *) +(* ||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/cpr_delift.ma". +include "basic_2/computation/cprs.ma". + +(* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************) + +(* Properties on inverse basic term relocation ******************************) + +(* Basic_1: was only: pr3_gen_cabbr *) +lemma thin_cprs_delift_conf: ∀L,U1,U2. L ⊢ U1 ➡* U2 → + ∀K,d,e. L [d, e] ≡ K → ∀T1. L ⊢ U1 [d, e] ≡ T1 → + ∃∃T2. K ⊢ T1 ➡* T2 & L ⊢ U2 [d, e] ≡ T2. +#L #U1 #U2 #H @(cprs_ind … H) -U2 /2 width=3/ +#U #U2 #_ #HU2 #IHU1 #K #d #e #HLK #T1 #HTU1 +elim (IHU1 … HLK … HTU1) -U1 #T #HT1 #HUT +elim (thin_cpr_delift_conf … HU2 … HLK … HUT) -U -HLK /3 width=3/ +qed.