]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_feqx.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / csx_feqx.ma
index 7cf01b28dbfa103a451d62825d04db1c00f5caca..5c8c2c00ca1d672a81dcd0349891f41dd714834f 100644 (file)
@@ -19,8 +19,9 @@ include "basic_2/rt_computation/csx_reqx.ma".
 
 (* Properties with sort-irrelevant equivalence for closures *****************)
 
-lemma csx_feqx_conf: ∀h,G1,L1,T1. ❪G1,L1❫ ⊢ ⬈*[h] 𝐒❪T1❫ →
-                     ∀G2,L2,T2. ❪G1,L1,T1❫ ≛ ❪G2,L2,T2❫ → ❪G2,L2❫ ⊢ ⬈*[h] 𝐒❪T2❫.
+lemma csx_feqx_conf (h):
+      ∀G1,L1,T1. ❪G1,L1❫ ⊢ ⬈*𝐒[h] T1 →
+      ∀G2,L2,T2. ❪G1,L1,T1❫ ≛ ❪G2,L2,T2❫ → ❪G2,L2❫ ⊢ ⬈*𝐒[h] T2.
 #h #G1 #L1 #T1 #HT1 #G2 #L2 #T2 * -G2 -L2 -T2
 /3 width=3 by csx_reqx_conf, csx_teqx_trans/
 qed-.