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=d12a9f498811ef033f202c59c46dd3460a946aae;hpb=bbac36729dab046d61019081c1523af06d876103;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 d12a9f498..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 @@ -20,6 +20,21 @@ 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. +#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 * +[ -IHVs #V0 #T0 #_ #_ #H destruct /2 width=1/ +| #V0 #W0 #T0 #HV0 #HT0 #HU + lapply (IHVs … HT0) -IHVs -HT0 #HT0 + elim (tstc_inv_bind_appls_simple … HT0 ?) // +| #V1 #V2 #V0 #T0 #HV1 #HV12 #HT0 #HU + lapply (IHVs … HT0) -IHVs -HT0 #HT0 + elim (tstc_inv_bind_appls_simple … HT0 ?) // +] +qed-. + (* Basic_1: was: pr3_iso_appls_beta *) lemma cprs_fwd_beta_vector: ∀L,Vs,V,W,T,U. L ⊢ ⒶVs. ⓐV. ⓛW. T ➡* U → ⒶVs. ⓐV. ⓛW. T ≃ U ∨ L ⊢ ⒶVs. ⓓV. T ➡* U. @@ -119,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 *