X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_1%2Fdrop1%2Ffwd.ma;h=3195918959af365def7b86753c8411be4f649020;hp=1aacc171f50f37f761372c4fee855372e4ca2269;hb=57ae1762497a5f3ea75740e2908e04adb8642cc2;hpb=639e798161afea770f41d78673c0fe3be4125beb diff --git a/matita/matita/contribs/lambdadelta/basic_1/drop1/fwd.ma b/matita/matita/contribs/lambdadelta/basic_1/drop1/fwd.ma index 1aacc171f..319591895 100644 --- a/matita/matita/contribs/lambdadelta/basic_1/drop1/fwd.ma +++ b/matita/matita/contribs/lambdadelta/basic_1/drop1/fwd.ma @@ -16,16 +16,16 @@ include "basic_1/drop1/defs.ma". -let rec drop1_ind (P: (PList \to (C \to (C \to Prop)))) (f: (\forall (c: -C).(P PNil c c))) (f0: (\forall (c1: C).(\forall (c2: C).(\forall (h: -nat).(\forall (d: nat).((drop h d c1 c2) \to (\forall (c3: C).(\forall (hds: -PList).((drop1 hds c2 c3) \to ((P hds c2 c3) \to (P (PCons h d hds) c1 -c3))))))))))) (p: PList) (c: C) (c0: C) (d: drop1 p c c0) on d: P p c c0 \def -match d with [(drop1_nil c1) \Rightarrow (f c1) | (drop1_cons c1 c2 h d0 d1 -c3 hds d2) \Rightarrow (f0 c1 c2 h d0 d1 c3 hds d2 ((drop1_ind P f f0) hds c2 -c3 d2))]. +implied rec lemma drop1_ind (P: (PList \to (C \to (C \to Prop)))) (f: +(\forall (c: C).(P PNil c c))) (f0: (\forall (c1: C).(\forall (c2: +C).(\forall (h: nat).(\forall (d: nat).((drop h d c1 c2) \to (\forall (c3: +C).(\forall (hds: PList).((drop1 hds c2 c3) \to ((P hds c2 c3) \to (P (PCons +h d hds) c1 c3))))))))))) (p: PList) (c: C) (c0: C) (d: drop1 p c c0) on d: P +p c c0 \def match d with [(drop1_nil c1) \Rightarrow (f c1) | (drop1_cons c1 +c2 h d0 d1 c3 hds d2) \Rightarrow (f0 c1 c2 h d0 d1 c3 hds d2 ((drop1_ind P f +f0) hds c2 c3 d2))]. -theorem drop1_gen_pnil: +lemma drop1_gen_pnil: \forall (c1: C).(\forall (c2: C).((drop1 PNil c1 c2) \to (eq C c1 c2))) \def \lambda (c1: C).(\lambda (c2: C).(\lambda (H: (drop1 PNil c1 c2)).(insert_eq @@ -41,7 +41,7 @@ PList).(\lambda (_: (drop1 hds c4 c5)).(\lambda (_: (((eq PList hds PNil) \to \Rightarrow False | (PCons _ _ _) \Rightarrow True])) I PNil H4) in (False_ind (eq C c3 c5) H5)))))))))))) y c1 c2 H0))) H))). -theorem drop1_gen_pcons: +lemma drop1_gen_pcons: \forall (c1: C).(\forall (c3: C).(\forall (hds: PList).(\forall (h: nat).(\forall (d: nat).((drop1 (PCons h d hds) c1 c3) \to (ex2 C (\lambda (c2: C).(drop h d c1 c2)) (\lambda (c2: C).(drop1 hds c2 c3))))))))