]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_cnx.ma
additions and corrections for the article on λδ-2B
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpxs_cnx.ma
index 6dd8020012ed4fcf26379ff10596ba1763ea656b..453c91e452825612c20b170cc3d322cd6f40c1e7 100644 (file)
@@ -17,10 +17,16 @@ include "basic_2/rt_computation/cpxs.ma".
 
 (* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************)
 
+(* Properties with normal forms *********************************************)
+
+lemma cpxs_cnx (h) (G) (L) (T1):
+      (∀T2. ⦃G,L⦄ ⊢ T1 ⬈*[h] T2 → T1 ≛ T2) → ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃T1⦄.
+/3 width=1 by cpx_cpxs/ qed.
+
 (* Inversion lemmas with normal terms ***************************************)
 
-lemma cpxs_inv_cnx1: ∀h,G,L,T1,T2. ⦃G,L⦄ ⊢ T1 ⬈*[h] T2 → ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃T1⦄ →
-                     T1 ≛ T2.
+lemma cpxs_inv_cnx1 (h) (G) (L):
+      ∀T1,T2. ⦃G,L⦄ ⊢ T1 ⬈*[h] T2 → ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃T1⦄ → T1 ≛ T2.
 #h #G #L #T1 #T2 #H @(cpxs_ind_dx … H) -T1
 /5 width=9 by cnx_tdeq_trans, tdeq_trans/
 qed-.