]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_lpx.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpxs_lpx.ma
index c4fcab18634f23e39e60c11be422e9e9558eb07a..5e17096f6f2b72260425c6ef322a6796bc9848d0 100644 (file)
@@ -16,9 +16,9 @@ include "basic_2/rt_transition/lpx.ma".
 include "basic_2/rt_computation/cpxs_drops.ma".
 include "basic_2/rt_computation/cpxs_cpxs.ma".
 
-(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS ************)
+(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************)
 
-(* Properties with uncounted parallel rt-transition for local environments **)
+(* Properties with unbound parallel rt-transition for local environments ****)
 
 lemma lpx_cpx_trans: ∀h,G. s_r_transitive … (cpx h G) (λ_.lpx h G).
 #h #G #L2 #T1 #T2 #H @(cpx_ind … H) -G -L2 -T1 -T2
@@ -38,5 +38,5 @@ lemma lpx_cpx_trans: ∀h,G. s_r_transitive … (cpx h G) (λ_.lpx h G).
 qed-.
 
 lemma lpx_cpxs_trans: ∀h,G. s_rs_transitive … (cpx h G) (λ_.lpx h G).
-#h #G @s_r_trans_LTC1 /2 width=3 by lpx_cpx_trans/ (**) (* full auto fails *)
+#h #G @s_r_trans_CTC1 /2 width=3 by lpx_cpx_trans/ (**) (* full auto fails *)
 qed-.