]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_vector.ma
syntactic components detached from basic_2 become static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / csx_vector.ma
index 6fdda2bbfb0ba80334d3f068698209206b1509b4..32afe79a4a87cf73d6d69943f70b3261fd8d655e 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/syntax/term_vector.ma".
+include "static_2/syntax/term_vector.ma".
 include "basic_2/rt_computation/csx.ma".
 
 (* STRONGLY NORMALIZING TERMS VECTORS FOR UNBOUND PARALLEL RT-TRANSITION ****)
@@ -26,7 +26,7 @@ interpretation
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma csxv_inv_cons: ∀h,o,G,L,T,Ts. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T@Ts⦄ →
+lemma csxv_inv_cons: ∀h,o,G,L,T,Ts. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃TTs⦄ →
                      ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ ∧ ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃Ts⦄.
 normalize // qed-.