]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc_vector.ma
- relation between native type and atomic arity proced
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / cprs_tstc_vector.ma
index 2f8486cc26d70f2b557428f87b0927010e2867ac..5a4c2bc3421c5da738b15ec6e51a91379dd72471 100644 (file)
@@ -20,7 +20,8 @@ include "basic_2/computation/cprs_tstc.ma".
 
 (* Vector form of forward lemmas involving same top term constructor ********)
 
-lemma cprs_fwd_cnf_vector: ∀L,T.  𝐒[T] → L ⊢ 𝐍[T] → ∀Vs,U. L ⊢ ⒶVs.T ➡* U → ⒶVs.T ≃ U.
+(* Basic_1: was just: nf2_iso_appls_lref *)
+lemma cprs_fwd_cnf_vector: ∀L,T.  𝐒⦃T⦄ → L ⊢ 𝐍⦃T⦄ → ∀Vs,U. L ⊢ ⒶVs.T ➡* U → ⒶVs.T ≃ U.
 #L #T #H1T #H2T #Vs elim Vs -Vs [ @(cprs_fwd_cnf … H2T) ] (**) (* /2 width=3 by cprs_fwd_cnf/ does not work *)
 #V #Vs #IHVs #U #H
 elim (cprs_inv_appl1 … H) -H *
@@ -133,8 +134,8 @@ elim (cprs_inv_appl1 … H) -H *
 qed-.
 
 (* Basic_1: was: pr3_iso_appls_cast *)
-lemma cprs_fwd_tau_vector: â\88\80L,Vs,W,T,U. L â\8a¢ â\92¶Vs. â\93£W. T ➡* U →
-                           â\92¶Vs. â\93£W. T ≃ U ∨ L ⊢ ⒶVs. T ➡* U.
+lemma cprs_fwd_tau_vector: â\88\80L,Vs,W,T,U. L â\8a¢ â\92¶Vs. â\93\9dW. T ➡* U →
+                           â\92¶Vs. â\93\9dW. T ≃ U ∨ L ⊢ ⒶVs. T ➡* U.
 #L #Vs elim Vs -Vs /2 width=1 by cprs_fwd_tau/
 #V #Vs #IHVs #W #T #U #H
 elim (cprs_inv_appl1 … H) -H *