X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fcsn_tstc_vector.ma;h=e1a64150599580ba687f254c81f3e33ce4ebe60c;hb=636c25914e83819c2f529edc891a7eb899499a97;hp=00de73f528cc89cf85672740d583628754ac7439;hpb=4a5254d45ba455e195b7ae2afca2212446e65ca3;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 00de73f52..e1a641505 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 @@ -28,11 +28,9 @@ lemma csn_applv_cnf: ∀L,T. 𝐒[T] → L ⊢ 𝐍[T] → #V #Vs #IHV #H elim (csnv_inv_cons … H) -H #HV #HVs @csn_appl_simple_tstc // -HV /2 width=1/ -IHV -HVs -[ #X #H #H0 - lapply (cprs_fwd_cnf_vector … H) -H // -H1T -H2T #H - elim (H0 ?) -H0 // -| -L -V elim Vs -Vs // -] +#X #H #H0 +lapply (cprs_fwd_cnf_vector … H) -H // -H1T -H2T #H +elim (H0 ?) -H0 // qed. (* Basic_1: was: sn3_appls_beta *) @@ -44,12 +42,10 @@ lemma csn_applv_beta: ∀L,W. L ⊢ ⬇* W → lapply (csn_fwd_pair_sn … H1T) #HV0 lapply (csn_fwd_flat_dx … H1T) #H2T @csn_appl_simple_tstc // -HV0 /2 width=1/ -IHV -H2T -[ #X #H #H0 - elim (cprs_fwd_beta_vector … H) -H #H - [ -H1T elim (H0 ?) -H0 // - | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/ - ] -| -H1T elim Vs -Vs // +#X #H #H0 +elim (cprs_fwd_beta_vector … H) -H #H +[ -H1T elim (H0 ?) -H0 // +| -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/ ] qed. @@ -64,12 +60,10 @@ lemma csn_applv_delta: ∀L,K,V1,i. ⇩[0, i] L ≡ K. ⓓV1 → lapply (csn_fwd_pair_sn … H1T) #HV lapply (csn_fwd_flat_dx … H1T) #H2T @csn_appl_simple_tstc // -HV /2 width=1/ -IHV -H2T - [ #X #H #H0 - elim (cprs_fwd_delta_vector … HLK … HV12 … H) -HLK -HV12 -H #H - [ -H1T elim (H0 ?) -H0 // - | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/ - ] - | -L -K -V -V1 -V2 elim Vs -Vs // + #X #H #H0 + elim (cprs_fwd_delta_vector … HLK … HV12 … H) -HLK -HV12 -H #H + [ -H1T elim (H0 ?) -H0 // + | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/ ] ] qed. @@ -87,7 +81,7 @@ 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/ ] @@ -102,12 +96,10 @@ lemma csn_applv_tau: ∀L,W. L ⊢ ⬇* W → lapply (csn_fwd_pair_sn … H1T) #HV lapply (csn_fwd_flat_dx … H1T) #H2T @csn_appl_simple_tstc // -HV /2 width=1/ -IHV -H2T -[ #X #H #H0 - elim (cprs_fwd_tau_vector … H) -H #H - [ -H1T elim (H0 ?) -H0 // - | -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/ - ] -| -H1T elim Vs -Vs // +#X #H #H0 +elim (cprs_fwd_tau_vector … H) -H #H +[ -H1T elim (H0 ?) -H0 // +| -H0 @(csn_cprs_trans … H1T) -H1T /2 width=1/ ] qed.