X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcsx_csx_vector.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcsx_csx_vector.ma;h=068f4dcb4120a4c9878b01bdf0e0b0ba3e45052a;hb=8ec019202bff90959cf1a7158b309e7f83fa222e;hp=c5d14bbe2273e6a3e619775570991d57e7a9ec79;hpb=33d0a7a9029859be79b25b5a495e0f30dab11f37;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx_vector.ma index c5d14bbe2..068f4dcb4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx_vector.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_csx_vector.ma @@ -24,8 +24,8 @@ include "basic_2/rt_computation/csx_vector.ma". (* Basic_1: was just: sn3_appls_beta *) lemma csx_applv_beta (G) (L): - ∀p,Vs,V,W,T. ❪G,L❫ ⊢ ⬈*𝐒 ⒶVs.ⓓ[p]ⓝW.V.T → - ❪G,L❫ ⊢ ⬈*𝐒 ⒶVs.ⓐV.ⓛ[p]W.T. + ∀p,Vs,V,W,T. ❨G,L❩ ⊢ ⬈*𝐒 ⒶVs.ⓓ[p]ⓝW.V.T → + ❨G,L❩ ⊢ ⬈*𝐒 ⒶVs.ⓐV.ⓛ[p]W.T. #G #L #p #Vs elim Vs -Vs /2 width=1 by csx_appl_beta/ #V0 #Vs #IHV #V #W #T #H1T lapply (csx_fwd_pair_sn … H1T) #HV0 @@ -41,7 +41,7 @@ qed. lemma csx_applv_delta_drops (G) (L): ∀I,K,V1,i. ⇩[i] L ≘ K.ⓑ[I]V1 → ∀V2. ⇧[↑i] V1 ≘ V2 → - ∀Vs. ❪G,L❫ ⊢ ⬈*𝐒 ⒶVs.V2 → ❪G,L❫ ⊢ ⬈*𝐒 ⒶVs.#i. + ∀Vs. ❨G,L❩ ⊢ ⬈*𝐒 ⒶVs.V2 → ❨G,L❩ ⊢ ⬈*𝐒 ⒶVs.#i. #G #L #I #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs [ /4 width=11 by csx_inv_lifts, csx_lref_pair_drops, drops_isuni_fwd_drop2/ | #V #Vs #IHV #H1T @@ -59,7 +59,7 @@ qed. (* Basic_1: was just: sn3_appls_abbr *) lemma csx_applv_theta (G) (L): ∀p,V1b,V2b. ⇧[1] V1b ≘ V2b → - ∀V,T. ❪G,L❫ ⊢ ⬈*𝐒 ⓓ[p]V.ⒶV2b.T → ❪G,L❫ ⊢ ⬈*𝐒 ⒶV1b.ⓓ[p]V.T. + ∀V,T. ❨G,L❩ ⊢ ⬈*𝐒 ⓓ[p]V.ⒶV2b.T → ❨G,L❩ ⊢ ⬈*𝐒 ⒶV1b.ⓓ[p]V.T. #G #L #p #V1b #V2b * -V1b -V2b /2 width=1 by/ #V1b #V2b #V1 #V2 #HV12 #H generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1 @@ -77,8 +77,8 @@ qed. (* Basic_1: was just: sn3_appls_cast *) lemma csx_applv_cast (G) (L): - ∀Vs,U. ❪G,L❫ ⊢ ⬈*𝐒 ⒶVs.U → - ∀T. ❪G,L❫ ⊢ ⬈*𝐒 ⒶVs.T → ❪G,L❫ ⊢ ⬈*𝐒 ⒶVs.ⓝU.T. + ∀Vs,U. ❨G,L❩ ⊢ ⬈*𝐒 ⒶVs.U → + ∀T. ❨G,L❩ ⊢ ⬈*𝐒 ⒶVs.T → ❨G,L❩ ⊢ ⬈*𝐒 ⒶVs.ⓝU.T. #G #L #Vs elim Vs -Vs /2 width=1 by csx_cast/ #V #Vs #IHV #U #H1U #T #H1T lapply (csx_fwd_pair_sn … H1U) #HV