X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fcsn.ma;h=6b58893271fe1f020ff2061e402c68e1d8f56b1a;hb=fde3b3d2e6cc48f6c9880136b1a0d565e2c78c1f;hp=903c62ea2eca5351aaad0be7c48a56b7b920612d;hpb=78d4844bcccb3deb58a3179151c3045298782b18;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/csn.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/csn.ma index 903c62ea2..6b5889327 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/computation/csn.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/csn.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -include "basic_2/reducibility/cpr.ma". include "basic_2/reducibility/cnf.ma". (* CONTEXT-SENSITIVE STRONGLY NORMALIZING TERMS *****************************) @@ -43,7 +42,7 @@ lemma csn_intro: ∀L,T1. /4 width=1/ qed. (* Basic_1: was: sn3_nf2 *) -lemma csn_cnf: ∀L,T. L ⊢ 𝐍[T] → L ⊢ ⬇* T. +lemma csn_cnf: ∀L,T. L ⊢ 𝐍⦃T⦄ → L ⊢ ⬇* T. /2 width=1/ qed. lemma csn_cpr_trans: ∀L,T1. L ⊢ ⬇* T1 → ∀T2. L ⊢ T1 ➡ T2 → L ⊢ ⬇* T2. @@ -55,7 +54,7 @@ elim (term_eq_dec T1 T2) #HT12 qed. (* Basic_1: was: sn3_cast *) -lemma csn_cast: ∀L,W. L ⊢ ⬇* W → ∀T. L ⊢ ⬇* T → L ⊢ ⬇* ⓣW. T. +lemma csn_cast: ∀L,W. L ⊢ ⬇* W → ∀T. L ⊢ ⬇* T → L ⊢ ⬇* ⓝW. T. #L #W #HW elim HW -W #W #_ #IHW #T #HT @(csn_ind … HT) -T #T #HT #IHT @csn_intro #X #H1 #H2 elim (cpr_inv_cast1 … H1) -H1