]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tdeq.ma
advances on cpxs and cnx (cnxa removed) ,,,
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpxs_tdeq.ma
index c3e2e20658368c7d0185d4dbbd60f891b2b7b3dd..4960595ffb7c758a144e2d07f75a43cd0fbeb692 100644 (file)
 (**************************************************************************)
 
 include "basic_2/syntax/tdeq_tdeq.ma".
-include "basic_2/rt_computation/cpxs.ma".
-include "basic_2/rt_transition/cpx_lfdeq.ma".
 include "basic_2/rt_transition/lfpx_fqup.ma".
+include "basic_2/rt_transition/lfpx_lfdeq.ma".
+include "basic_2/rt_computation/cpxs.ma".
 
 (* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS ************)
 
+(* Properties with degree-based equivalence for terms ***********************)
+
 lemma tdeq_cpxs_trans: ∀h,o,U1,T1. U1 ≡[h, o] T1 → ∀G,L,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 → 
                        ∃∃U2.  ⦃G, L⦄ ⊢ U1 ⬈*[h] U2 & U2 ≡[h, o] T2.
 #h #o #U1 #T1 #HUT1 #G #L #T2 #HT12 @(cpxs_ind … HT12) -T2 /2 width=3 by ex2_intro/
@@ -40,5 +42,3 @@ lemma cpxs_tdneq_inv_step_sn: ∀h,o,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ⬈*[h] T2 →
   ]
 ]
 qed-.
-
-(* Basic_2A1: removed theorems 1: cpxs_neq_inv_step_sn *)