X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fcprs_tstc_vector.ma;h=5a4c2bc3421c5da738b15ec6e51a91379dd72471;hb=fde3b3d2e6cc48f6c9880136b1a0d565e2c78c1f;hp=feaedbf76bfd6be07b00496a2b065845f9cffe6b;hpb=7cf0b6c720e4d7fa05dd25ec0ad0478c0802ba67;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc_vector.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc_vector.ma index feaedbf76..5a4c2bc34 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc_vector.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/cprs_tstc_vector.ma @@ -21,7 +21,7 @@ include "basic_2/computation/cprs_tstc.ma". (* Vector form of forward lemmas involving same top term constructor ********) (* 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. +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 * @@ -134,8 +134,8 @@ elim (cprs_inv_appl1 … H) -H * qed-. (* Basic_1: was: pr3_iso_appls_cast *) -lemma cprs_fwd_tau_vector: ∀L,Vs,W,T,U. L ⊢ ⒶVs. ⓣW. T ➡* U → - ⒶVs. ⓣW. T ≃ U ∨ L ⊢ ⒶVs. T ➡* U. +lemma cprs_fwd_tau_vector: ∀L,Vs,W,T,U. L ⊢ ⒶVs. ⓝW. T ➡* U → + ⒶVs. ⓝW. 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 *