]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx.ma
syntactic components detached from basic_2 become static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / csx.ma
index 5e255d297e7ee9e3f5521e5b0e0326383d85346e..4c8371a0b74f9b8c748e9e9a35140b56ed215879 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "basic_2/notation/relations/predtystrong_5.ma".
-include "basic_2/syntax/tdeq.ma".
+include "static_2/syntax/tdeq.ma".
 include "basic_2/rt_transition/cpx.ma".
 
 (* STRONGLY NORMALIZING TERMS FOR UNBOUND PARALLEL RT-TRANSITION ************)
@@ -27,13 +27,13 @@ interpretation
 
 (* Basic eliminators ********************************************************)
 
-lemma csx_ind: ∀h,o,G,L. ∀R:predicate term.
+lemma csx_ind: ∀h,o,G,L. ∀Q:predicate term.
                (∀T1. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T1⦄ →
-                     (∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → (T1 ≛[h, o] T2 → ⊥) → R T2) →
-                     R T1
+                     (∀T2. ⦃G, L⦄ ⊢ T1 ⬈[h] T2 → (T1 ≛[h, o] T2 → ⊥) → Q T2) →
+                     Q T1
                ) →
-               ∀T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → R T.
-#h #o #G #L #R #H0 #T1 #H elim H -T1
+               ∀T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ →  Q T.
+#h #o #G #L #Q #H0 #T1 #H elim H -T1
 /5 width=1 by SN_intro/
 qed-.