X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Fdrops.ma;h=57bf1bfe510eabe95e007e54f0d34a45b5fee7d0;hb=268e7f336d036f77ffc9663358e9afda92b97730;hp=d469f0ac178cb7c9c66d767fd22d392a46de8c08;hpb=c879284b576409cec07e96c1f08510d9d9ac14f3;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma index d469f0ac1..57bf1bfe5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma @@ -92,7 +92,7 @@ definition co_dropable_dx: predicate (rtmap → relation lenv) ≝ definition co_dedropable_sn: predicate (rtmap → relation lenv) ≝ λR. ∀b,f,L1,K1. ⬇*[b, f] L1 ≡ K1 → ∀f1,K2. R f1 K1 K2 → ∀f2. f ~⊚ f1 ≡ f2 → - ∃∃L2. R f2 L1 L2 & ⬇*[b, f] L2 ≡ K2 & L1 ≡[f] L2. + ∃∃L2. R f2 L1 L2 & ⬇*[b, f] L2 ≡ K2 & L1 ≐[f] L2. (* Basic properties *********************************************************) @@ -432,10 +432,12 @@ qed-. (* Properties with context-sensitive equivalence for terms ******************) -lemma ceq_lift_sn: d_liftable2_sn … liftsb ceq. +lemma ceq_lift_sn: d_liftable2_sn … liftsb ceq_ext. +#K #I1 #I2 #H <(ceq_ext_inv_eq … H) -I2 /2 width=3 by ex2_intro/ qed-. -lemma ceq_inv_lift_sn: d_deliftable2_sn … liftsb ceq. +lemma ceq_inv_lift_sn: d_deliftable2_sn … liftsb ceq_ext. +#L #J1 #J2 #H <(ceq_ext_inv_eq … H) -J2 /2 width=3 by ex2_intro/ qed-. (* Note: d_deliftable2_sn cfull does not hold *)