X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Fcprs_lift.ma;h=c36bf3cc7d26701833d19962b23172aaf3077527;hb=43282d3750af8831c8100c60d75c56fdfb7ff3c9;hp=0a7775096790672ed220251224f3f604c953888a;hpb=6c985e4e2e7846a2b9abd0c84569f21c24e9ce2f;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma index 0a7775096..c36bf3cc7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cprs_lift.ma @@ -21,8 +21,8 @@ include "basic_2/computation/cprs.ma". (* Note: apparently this was missing in basic_1 *) lemma cprs_delta: ∀G,L,K,V,V2,i. - ⇩[i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V ➡* V2 → - ∀W2. ⇧[0, i + 1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡* W2. + ⬇[i] L ≡ K.ⓓV → ⦃G, K⦄ ⊢ V ➡* V2 → + ∀W2. ⬆[0, i + 1] V2 ≡ W2 → ⦃G, L⦄ ⊢ #i ➡* W2. #G #L #K #V #V2 #i #HLK #H elim H -V2 [ /3 width=6 by cpr_cprs, cpr_delta/ ] #V1 #V2 #_ #HV12 #IHV1 #W2 #HVW2 lapply (drop_fwd_drop2 … HLK) -HLK #HLK @@ -34,8 +34,8 @@ qed. (* Basic_1: was: pr3_gen_lref *) lemma cprs_inv_lref1: ∀G,L,T2,i. ⦃G, L⦄ ⊢ #i ➡* T2 → T2 = #i ∨ - ∃∃K,V1,T1. ⇩[i] L ≡ K.ⓓV1 & ⦃G, K⦄ ⊢ V1 ➡* T1 & - ⇧[0, i + 1] T1 ≡ T2. + ∃∃K,V1,T1. ⬇[i] L ≡ K.ⓓV1 & ⦃G, K⦄ ⊢ V1 ➡* T1 & + ⬆[0, i + 1] T1 ≡ T2. #G #L #T2 #i #H @(cprs_ind … H) -T2 /2 width=1 by or_introl/ #T #T2 #_ #HT2 * [ #H destruct