]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_1/drop1/fwd.ma
flavour and source information exported for the objects of lambdadelta version 1
[helm.git] / matita / matita / contribs / lambdadelta / basic_1 / drop1 / fwd.ma
index 1aacc171f50f37f761372c4fee855372e4ca2269..3270fe008c9905a42e76293741fe16ec503390e6 100644 (file)
@@ -16,8 +16,8 @@
 
 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: 
+implied 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 
@@ -25,7 +25,7 @@ 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))))))))