X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Freducibility%2Ftnf.ma;h=6f6d99096942fbef746c84665aea4204504cee06;hb=a386e0eb6b20909ae78c825203d77647b8fde30c;hp=5a7205cc01a7c61fd4c098f8772a74f99ecc3508;hpb=eb918fc784eacd2094e3986ba321ef47690d9983;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf.ma index 5a7205cc0..6f6d99096 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/tnf.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "Basic_2/reducibility/tpr.ma". +include "basic_2/reducibility/tpr.ma". (* CONTEXT-FREE NORMAL TERMS ************************************************) @@ -42,7 +42,7 @@ lemma tnf_inv_appl: ∀V,T. 𝐍[ⓐV.T] → ∧∧ 𝐍[V] & 𝐍[T] & 𝐒[T]. ] qed-. -lemma tnf_inv_abbr: ∀V,T. 𝐍[ⓓV.T] → False. +lemma tnf_inv_abbr: ∀V,T. 𝐍[ⓓV.T] → ⊥. #V #T #H elim (is_lift_dec T 0 1) [ * #U #HTU lapply (H U ?) -H /2 width=3/ #H destruct @@ -53,7 +53,7 @@ lemma tnf_inv_abbr: ∀V,T. 𝐍[ⓓV.T] → False. ] qed. -lemma tnf_inv_cast: ∀V,T. 𝐍[ⓣV.T] → False. +lemma tnf_inv_cast: ∀V,T. 𝐍[ⓣV.T] → ⊥. #V #T #H lapply (H T ?) -H /2 width=1/ #H -@(discr_tpair_xy_y … H) +@discr_tpair_xy_y // qed-.