]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_1/drop1/fwd.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_1 / drop1 / fwd.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 (* This file was automatically generated: do not edit *********************)
16
17 include "basic_1/drop1/defs.ma".
18
19 implied rec lemma drop1_ind (P: (PList \to (C \to (C \to Prop)))) (f: 
20 (\forall (c: C).(P PNil c c))) (f0: (\forall (c1: C).(\forall (c2: 
21 C).(\forall (h: nat).(\forall (d: nat).((drop h d c1 c2) \to (\forall (c3: 
22 C).(\forall (hds: PList).((drop1 hds c2 c3) \to ((P hds c2 c3) \to (P (PCons 
23 h d hds) c1 c3))))))))))) (p: PList) (c: C) (c0: C) (d: drop1 p c c0) on d: P 
24 p c c0 \def match d with [(drop1_nil c1) \Rightarrow (f c1) | (drop1_cons c1 
25 c2 h d0 d1 c3 hds d2) \Rightarrow (f0 c1 c2 h d0 d1 c3 hds d2 ((drop1_ind P f 
26 f0) hds c2 c3 d2))].
27
28 lemma drop1_gen_pnil:
29  \forall (c1: C).(\forall (c2: C).((drop1 PNil c1 c2) \to (eq C c1 c2)))
30 \def
31  \lambda (c1: C).(\lambda (c2: C).(\lambda (H: (drop1 PNil c1 c2)).(insert_eq 
32 PList PNil (\lambda (p: PList).(drop1 p c1 c2)) (\lambda (_: PList).(eq C c1 
33 c2)) (\lambda (y: PList).(\lambda (H0: (drop1 y c1 c2)).(drop1_ind (\lambda 
34 (p: PList).(\lambda (c: C).(\lambda (c0: C).((eq PList p PNil) \to (eq C c 
35 c0))))) (\lambda (c: C).(\lambda (_: (eq PList PNil PNil)).(refl_equal C c))) 
36 (\lambda (c3: C).(\lambda (c4: C).(\lambda (h: nat).(\lambda (d: 
37 nat).(\lambda (_: (drop h d c3 c4)).(\lambda (c5: C).(\lambda (hds: 
38 PList).(\lambda (_: (drop1 hds c4 c5)).(\lambda (_: (((eq PList hds PNil) \to 
39 (eq C c4 c5)))).(\lambda (H4: (eq PList (PCons h d hds) PNil)).(let H5 \def 
40 (eq_ind PList (PCons h d hds) (\lambda (ee: PList).(match ee with [PNil 
41 \Rightarrow False | (PCons _ _ _) \Rightarrow True])) I PNil H4) in 
42 (False_ind (eq C c3 c5) H5)))))))))))) y c1 c2 H0))) H))).
43
44 lemma drop1_gen_pcons:
45  \forall (c1: C).(\forall (c3: C).(\forall (hds: PList).(\forall (h: 
46 nat).(\forall (d: nat).((drop1 (PCons h d hds) c1 c3) \to (ex2 C (\lambda 
47 (c2: C).(drop h d c1 c2)) (\lambda (c2: C).(drop1 hds c2 c3))))))))
48 \def
49  \lambda (c1: C).(\lambda (c3: C).(\lambda (hds: PList).(\lambda (h: 
50 nat).(\lambda (d: nat).(\lambda (H: (drop1 (PCons h d hds) c1 c3)).(insert_eq 
51 PList (PCons h d hds) (\lambda (p: PList).(drop1 p c1 c3)) (\lambda (_: 
52 PList).(ex2 C (\lambda (c2: C).(drop h d c1 c2)) (\lambda (c2: C).(drop1 hds 
53 c2 c3)))) (\lambda (y: PList).(\lambda (H0: (drop1 y c1 c3)).(drop1_ind 
54 (\lambda (p: PList).(\lambda (c: C).(\lambda (c0: C).((eq PList p (PCons h d 
55 hds)) \to (ex2 C (\lambda (c2: C).(drop h d c c2)) (\lambda (c2: C).(drop1 
56 hds c2 c0))))))) (\lambda (c: C).(\lambda (H1: (eq PList PNil (PCons h d 
57 hds))).(let H2 \def (eq_ind PList PNil (\lambda (ee: PList).(match ee with 
58 [PNil \Rightarrow True | (PCons _ _ _) \Rightarrow False])) I (PCons h d hds) 
59 H1) in (False_ind (ex2 C (\lambda (c2: C).(drop h d c c2)) (\lambda (c2: 
60 C).(drop1 hds c2 c))) H2)))) (\lambda (c2: C).(\lambda (c4: C).(\lambda (h0: 
61 nat).(\lambda (d0: nat).(\lambda (H1: (drop h0 d0 c2 c4)).(\lambda (c5: 
62 C).(\lambda (hds0: PList).(\lambda (H2: (drop1 hds0 c4 c5)).(\lambda (H3: 
63 (((eq PList hds0 (PCons h d hds)) \to (ex2 C (\lambda (c6: C).(drop h d c4 
64 c6)) (\lambda (c6: C).(drop1 hds c6 c5)))))).(\lambda (H4: (eq PList (PCons 
65 h0 d0 hds0) (PCons h d hds))).(let H5 \def (f_equal PList nat (\lambda (e: 
66 PList).(match e with [PNil \Rightarrow h0 | (PCons n _ _) \Rightarrow n])) 
67 (PCons h0 d0 hds0) (PCons h d hds) H4) in ((let H6 \def (f_equal PList nat 
68 (\lambda (e: PList).(match e with [PNil \Rightarrow d0 | (PCons _ n _) 
69 \Rightarrow n])) (PCons h0 d0 hds0) (PCons h d hds) H4) in ((let H7 \def 
70 (f_equal PList PList (\lambda (e: PList).(match e with [PNil \Rightarrow hds0 
71 | (PCons _ _ p) \Rightarrow p])) (PCons h0 d0 hds0) (PCons h d hds) H4) in 
72 (\lambda (H8: (eq nat d0 d)).(\lambda (H9: (eq nat h0 h)).(let H10 \def 
73 (eq_ind PList hds0 (\lambda (p: PList).((eq PList p (PCons h d hds)) \to (ex2 
74 C (\lambda (c6: C).(drop h d c4 c6)) (\lambda (c6: C).(drop1 hds c6 c5))))) 
75 H3 hds H7) in (let H11 \def (eq_ind PList hds0 (\lambda (p: PList).(drop1 p 
76 c4 c5)) H2 hds H7) in (let H12 \def (eq_ind nat d0 (\lambda (n: nat).(drop h0 
77 n c2 c4)) H1 d H8) in (let H13 \def (eq_ind nat h0 (\lambda (n: nat).(drop n 
78 d c2 c4)) H12 h H9) in (ex_intro2 C (\lambda (c6: C).(drop h d c2 c6)) 
79 (\lambda (c6: C).(drop1 hds c6 c5)) c4 H13 H11)))))))) H6)) H5)))))))))))) y 
80 c1 c3 H0))) H)))))).
81