]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpxs.ma
- ldrop is now drop as in basic_1
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / cpxs_cpxs.ma
index b6431c294bcee4fd18c9014c5527cb3fac1ce98e..a19c92545b1444d26c54dd6479a4c43a264b2fe7 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/reduction/lpx_ldrop.ma".
+include "basic_2/reduction/lpx_drop.ma".
 include "basic_2/computation/cpxs_lift.ma".
 
 (* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************)
@@ -98,7 +98,7 @@ lemma lpx_cpx_trans: ∀h,g,G. s_r_transitive … (cpx h g G) (λ_.lpx h g G).
 [ /2 width=3 by/
 | /3 width=2 by cpx_cpxs, cpx_st/
 | #I #G #L2 #K2 #V0 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV02 #L1 #HL12
-  elim (lpx_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
+  elim (lpx_drop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H
   elim (lpx_inv_pair2 … H) -H #K1 #V1 #HK12 #HV10 #H destruct
   /4 width=7 by cpxs_delta, cpxs_strap2/
 |4,9: /4 width=1 by cpxs_beta, cpxs_bind, lpx_pair/
@@ -132,7 +132,7 @@ lemma fqu_cpxs_trans_neq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2,
 #h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
 [ #I #G #L #V1 #V2 #HV12 #_ elim (lift_total V2 0 1)
   #U2 #HVU2 @(ex3_intro … U2)
-  [1,3: /3 width=7 by fqu_drop, cpxs_delta, ldrop_pair, ldrop_drop/
+  [1,3: /3 width=7 by fqu_drop, cpxs_delta, drop_pair, drop_drop/
   | #H destruct /2 width=7 by lift_inv_lref2_be/
   ]
 | #I #G #L #V1 #T #V2 #HV12 #H @(ex3_intro … (②{I}V2.T))