X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fcsn_tstc_vector.ma;h=91e3653f3decd3da0e2d22c3d02f25357da096dc;hb=439b6ec33d749ba4e6ae0938e973a85bc23e306e;hp=3aa7170f50da7ca220c06e2c529fa290cf4636b6;hpb=7cf0b6c720e4d7fa05dd25ec0ad0478c0802ba67;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_tstc_vector.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_tstc_vector.ma index 3aa7170f5..91e3653f3 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_tstc_vector.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_tstc_vector.ma @@ -22,8 +22,8 @@ include "basic_2/computation/csn_vector.ma". (* Advanced properties ******************************************************) (* Basic_1: was only: sn3_appls_lref *) -lemma csn_applv_cnf: ∀L,T. 𝐒[T] → L ⊢ 𝐍[T] → - ∀Vs. L ⊢ ⬇* Vs → L ⊢ ⬇* ⒶVs.T. +lemma csn_applv_cnf: ∀L,T. 𝐒⦃T⦄ → L ⊢ 𝐍⦃T⦄ → + ∀Vs. L ⊢ ⬊* Vs → L ⊢ ⬊* ⒶVs.T. #L #T #H1T #H2T #Vs elim Vs -Vs [ #_ @(csn_cnf … H2T) ] (**) (* /2 width=1/ does not work *) #V #Vs #IHV #H elim (csnv_inv_cons … H) -H #HV #HVs @@ -34,10 +34,10 @@ elim (H0 ?) -H0 // qed. (* Basic_1: was: sn3_appls_beta *) -lemma csn_applv_beta: ∀L,W. L ⊢ ⬇* W → - ∀Vs,V,T. L ⊢ ⬇* ⒶVs.ⓓV.T → - L ⊢ ⬇* ⒶVs. ⓐV.ⓛW. T. -#L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW +lemma csn_applv_beta: ∀a,L,W. L ⊢ ⬊* W → + ∀Vs,V,T. L ⊢ ⬊* ⒶVs.ⓓ{a}V.T → + L ⊢ ⬊* ⒶVs. ⓐV.ⓛ{a}W. T. +#a #L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW #V0 #Vs #IHV #V #T #H1T lapply (csn_fwd_pair_sn … H1T) #HV0 lapply (csn_fwd_flat_dx … H1T) #H2T @@ -51,7 +51,7 @@ qed. lemma csn_applv_delta: ∀L,K,V1,i. ⇩[0, i] L ≡ K. ⓓV1 → ∀V2. ⇧[0, i + 1] V1 ≡ V2 → - ∀Vs.L ⊢ ⬇* (ⒶVs. V2) → L ⊢ ⬇* (ⒶVs. #i). + ∀Vs.L ⊢ ⬊* (ⒶVs. V2) → L ⊢ ⬊* (ⒶVs. #i). #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs [ #H lapply (ldrop_fwd_ldrop2 … HLK) #HLK0 @@ -69,10 +69,10 @@ lemma csn_applv_delta: ∀L,K,V1,i. ⇩[0, i] L ≡ K. ⓓV1 → qed. (* Basic_1: was: sn3_appls_abbr *) -lemma csn_applv_theta: ∀L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → - ∀V,T. L ⊢ ⬇* ⓓV. ⒶV2s. T → L ⊢ ⬇* V → - L ⊢ ⬇* ⒶV1s. ⓓV. T. -#L #V1s #V2s * -V1s -V2s /2 width=1/ +lemma csn_applv_theta: ∀a,L,V1s,V2s. ⇧[0, 1] V1s ≡ V2s → + ∀V,T. L ⊢ ⬊* ⓓ{a}V. ⒶV2s. T → L ⊢ ⬊* V → + L ⊢ ⬊* ⒶV1s. ⓓ{a}V. T. +#a #L #V1s #V2s * -V1s -V2s /2 width=1/ #V1s #V2s #V1 #V2 #HV12 #H generalize in match HV12; -HV12 generalize in match V2; -V2 generalize in match V1; -V1 elim H -V1s -V2s /2 width=3/ @@ -81,16 +81,16 @@ lapply (csn_appl_theta … HW12 … H) -H -HW12 #H lapply (csn_fwd_pair_sn … H) #HW1 lapply (csn_fwd_flat_dx … H) #H1 @csn_appl_simple_tstc // -HW1 /2 width=3/ -IHV12s -HV -H1 #X #H1 #H2 -elim (cprs_fwd_theta_vector … (V2::V2s) … H1) -H1 /2 width=1/ -HV12s -HV12 +elim (cprs_fwd_theta_vector … (V2@V2s) … H1) -H1 /2 width=1/ -HV12s -HV12 [ -H #H elim (H2 ?) -H2 // | -H2 #H1 @(csn_cprs_trans … H) -H /2 width=1/ ] qed. (* Basic_1: was: sn3_appls_cast *) -lemma csn_applv_tau: ∀L,W. L ⊢ ⬇* W → - ∀Vs,T. L ⊢ ⬇* ⒶVs. T → - L ⊢ ⬇* ⒶVs. ⓣW. T. +lemma csn_applv_tau: ∀L,W. L ⊢ ⬊* W → + ∀Vs,T. L ⊢ ⬊* ⒶVs. T → + L ⊢ ⬊* ⒶVs. ⓝW. T. #L #W #HW #Vs elim Vs -Vs /2 width=1/ -HW #V #Vs #IHV #T #H1T lapply (csn_fwd_pair_sn … H1T) #HV @@ -103,12 +103,12 @@ elim (cprs_fwd_tau_vector … H) -H #H ] qed. -theorem csn_acr: acr cpr (eq …) (csn …) (λL,T. L ⊢ ⬇* T). +theorem csn_acr: acr cpr (eq …) (csn …) (λL,T. L ⊢ ⬊* T). @mk_acr // [ /3 width=1/ | /2 width=1/ | /2 width=6/ -| #L #V1 #V2 #HV12 #V #T #H #HVT +| #L #V1 #V2 #HV12 #a #V #T #H #HVT @(csn_applv_theta … HV12) -HV12 // @(csn_abbr) // | /2 width=1/