X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fcsn_lcpr.ma;h=169073516d2f5730b6a2ba0406244049d555fec2;hb=2ba7ef901a6b72210692792f2396c08bc0cff52c;hp=92d30c7fd663da23494af959ca462f8ce04d8f1a;hpb=eaaea3c18083de3e442e939768ff450d3b093911;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr.ma index 92d30c7fd..169073516 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn_lcpr.ma @@ -16,20 +16,20 @@ include "basic_2/grammar/tstc_tstc.ma". include "basic_2/computation/cprs_cprs.ma". include "basic_2/computation/csn_lift.ma". include "basic_2/computation/csn_cpr.ma". -include "basic_2/computation/csn_cprs.ma". +include "basic_2/computation/csn_alt.ma". (* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) (* Advanced properties ******************************************************) lemma csn_lcpr_conf: ∀L1,L2. L1 ⊢ ➡ L2 → ∀T. L1 ⊢ ⬇* T → L2 ⊢ ⬇* T. -#L1 #L2 #HL12 #T #H @(csn_ind_cprs … H) -T #T #_ #IHT +#L1 #L2 #HL12 #T #H @(csn_ind_alt … H) -T #T #_ #IHT @csn_intro #T0 #HLT0 #HT0 @IHT /2 width=2/ -IHT -HT0 /2 width=3/ qed. lemma csn_abbr: ∀L,V. L ⊢ ⬇* V → ∀T. L. ⓓV ⊢ ⬇* T → L ⊢ ⬇* ⓓV. T. -#L #V #HV elim HV -V #V #_ #IHV #T #HT @(csn_ind_cprs … HT) -T #T #HT #IHT +#L #V #HV elim HV -V #V #_ #IHV #T #HT @(csn_ind_alt … HT) -T #T #HT #IHT @csn_intro #X #H1 #H2 elim (cpr_inv_abbr1 … H1) -H1 * [ #V0 #V1 #T1 #HLV0 #HLV01 #HLT1 #H destruct @@ -47,7 +47,7 @@ qed. fact csn_appl_beta_aux: ∀L,W. L ⊢ ⬇* W → ∀U. L ⊢ ⬇* U → ∀V,T. U = ⓓV. T → L ⊢ ⬇* ⓐV. ⓛW. T. -#L #W #H elim H -W #W #_ #IHW #X #H @(csn_ind_cprs … H) -X #X #HVT #IHVT #V #T #H destruct +#L #W #H elim H -W #W #_ #IHW #X #H @(csn_ind_alt … H) -X #X #HVT #IHVT #V #T #H destruct lapply (csn_fwd_pair_sn … HVT) #HV lapply (csn_fwd_bind_dx … HVT) #HT -HVT @csn_intro #X #H #H2 @@ -76,7 +76,7 @@ lemma csn_appl_beta: ∀L,W. L ⊢ ⬇* W → ∀V,T. L ⊢ ⬇* ⓓV. T → fact csn_appl_theta_aux: ∀L,U. L ⊢ ⬇* U → ∀V1,V2. ⇧[0, 1] V1 ≡ V2 → ∀V,T. U = ⓓV. ⓐV2. T → L ⊢ ⬇* ⓐV1. ⓓV. T. -#L #X #H @(csn_ind_cprs … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct +#L #X #H @(csn_ind_alt … H) -X #X #HVT #IHVT #V1 #V2 #HV12 #V #T #H destruct lapply (csn_fwd_pair_sn … HVT) #HV lapply (csn_fwd_bind_dx … HVT) -HVT #HVT @csn_intro #X #HL #H