]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_lsuba.ma
- we introduced the pointer_step rc in the perspective of proving
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / lsubc_lsuba.ma
index e50551086c32a92f094251cddb424e1d3e757350..aad454f626941fb4829208c76542f9ec4633bb52 100644 (file)
@@ -20,7 +20,7 @@ include "basic_2/computation/acp_aaa.ma".
 (* properties concerning lenv refinement for atomic arity assignment ********)
 
 lemma lsubc_lsuba: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) →
-                   ∀L1,L2. L1 ÷⊑ L2 → L1 [RP] ⊑ L2.
+                   ∀L1,L2. L1 ⁝⊑ L2 → L1 ⊑[RP] L2.
 #RR #RS #RP #H1RP #H2RP #L1 #L2 #H elim H -L1 -L2
 // /2 width=1/ /3 width=4/
 qed.