X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_1%2Fcsubst0%2Fprops.ma;h=bb427a677eb88e1841d59891604160156c6e0638;hb=57ae1762497a5f3ea75740e2908e04adb8642cc2;hp=28b75ec37270c7b2869e9b110745284126404cad;hpb=88a68a9c334646bc17314d5327cd3b790202acd6;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_1/csubst0/props.ma b/matita/matita/contribs/lambdadelta/basic_1/csubst0/props.ma index 28b75ec37..bb427a677 100644 --- a/matita/matita/contribs/lambdadelta/basic_1/csubst0/props.ma +++ b/matita/matita/contribs/lambdadelta/basic_1/csubst0/props.ma @@ -14,9 +14,9 @@ (* This file was automatically generated: do not edit *********************) -include "Basic-1/csubst0/defs.ma". +include "basic_1/csubst0/defs.ma". -theorem csubst0_snd_bind: +lemma csubst0_snd_bind: \forall (b: B).(\forall (i: nat).(\forall (v: T).(\forall (u1: T).(\forall (u2: T).((subst0 i v u1 u2) \to (\forall (c: C).(csubst0 (S i) v (CHead c (Bind b) u1) (CHead c (Bind b) u2)))))))) @@ -26,11 +26,8 @@ theorem csubst0_snd_bind: b) i) (\lambda (n: nat).(csubst0 n v (CHead c (Bind b) u1) (CHead c (Bind b) u2))) (csubst0_snd (Bind b) i v u1 u2 H c) (S i) (refl_equal nat (S i))))))))). -(* COMMENTS -Initial nodes: 91 -END *) -theorem csubst0_fst_bind: +lemma csubst0_fst_bind: \forall (b: B).(\forall (i: nat).(\forall (c1: C).(\forall (c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to (\forall (u: T).(csubst0 (S i) v (CHead c1 (Bind b) u) (CHead c2 (Bind b) u)))))))) @@ -39,9 +36,6 @@ theorem csubst0_fst_bind: (v: T).(\lambda (H: (csubst0 i v c1 c2)).(\lambda (u: T).(eq_ind nat (s (Bind b) i) (\lambda (n: nat).(csubst0 n v (CHead c1 (Bind b) u) (CHead c2 (Bind b) u))) (csubst0_fst (Bind b) i c1 c2 v H u) (S i) (refl_equal nat (S i))))))))). -(* COMMENTS -Initial nodes: 91 -END *) theorem csubst0_both_bind: \forall (b: B).(\forall (i: nat).(\forall (v: T).(\forall (u1: T).(\forall @@ -55,7 +49,4 @@ C).(\lambda (H0: (csubst0 i v c1 c2)).(eq_ind nat (s (Bind b) i) (\lambda (n: nat).(csubst0 n v (CHead c1 (Bind b) u1) (CHead c2 (Bind b) u2))) (csubst0_both (Bind b) i v u1 u2 H c1 c2 H0) (S i) (refl_equal nat (S i))))))))))). -(* COMMENTS -Initial nodes: 107 -END *)