]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_1/clear/fwd.ma
refactoring of \lambda\delta version 1 in matita
[helm.git] / matita / matita / contribs / lambdadelta / basic_1 / clear / 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/clear/defs.ma".
18
19 theorem clear_gen_sort:
20  \forall (x: C).(\forall (n: nat).((clear (CSort n) x) \to (\forall (P: 
21 Prop).P)))
22 \def
23  \lambda (x: C).(\lambda (n: nat).(\lambda (H: (clear (CSort n) x)).(\lambda 
24 (P: Prop).(insert_eq C (CSort n) (\lambda (c: C).(clear c x)) (\lambda (_: 
25 C).P) (\lambda (y: C).(\lambda (H0: (clear y x)).(clear_ind (\lambda (c: 
26 C).(\lambda (_: C).((eq C c (CSort n)) \to P))) (\lambda (b: B).(\lambda (e: 
27 C).(\lambda (u: T).(\lambda (H1: (eq C (CHead e (Bind b) u) (CSort n))).(let 
28 H2 \def (eq_ind C (CHead e (Bind b) u) (\lambda (ee: C).(match ee in C return 
29 (\lambda (_: C).Prop) with [(CSort _) \Rightarrow False | (CHead _ _ _) 
30 \Rightarrow True])) I (CSort n) H1) in (False_ind P H2)))))) (\lambda (e: 
31 C).(\lambda (c: C).(\lambda (_: (clear e c)).(\lambda (_: (((eq C e (CSort 
32 n)) \to P))).(\lambda (f: F).(\lambda (u: T).(\lambda (H3: (eq C (CHead e 
33 (Flat f) u) (CSort n))).(let H4 \def (eq_ind C (CHead e (Flat f) u) (\lambda 
34 (ee: C).(match ee in C return (\lambda (_: C).Prop) with [(CSort _) 
35 \Rightarrow False | (CHead _ _ _) \Rightarrow True])) I (CSort n) H3) in 
36 (False_ind P H4))))))))) y x H0))) H)))).
37 (* COMMENTS
38 Initial nodes: 215
39 END *)
40
41 theorem clear_gen_bind:
42  \forall (b: B).(\forall (e: C).(\forall (x: C).(\forall (u: T).((clear 
43 (CHead e (Bind b) u) x) \to (eq C x (CHead e (Bind b) u))))))
44 \def
45  \lambda (b: B).(\lambda (e: C).(\lambda (x: C).(\lambda (u: T).(\lambda (H: 
46 (clear (CHead e (Bind b) u) x)).(insert_eq C (CHead e (Bind b) u) (\lambda 
47 (c: C).(clear c x)) (\lambda (c: C).(eq C x c)) (\lambda (y: C).(\lambda (H0: 
48 (clear y x)).(clear_ind (\lambda (c: C).(\lambda (c0: C).((eq C c (CHead e 
49 (Bind b) u)) \to (eq C c0 c)))) (\lambda (b0: B).(\lambda (e0: C).(\lambda 
50 (u0: T).(\lambda (H1: (eq C (CHead e0 (Bind b0) u0) (CHead e (Bind b) 
51 u))).(let H2 \def (f_equal C C (\lambda (e1: C).(match e1 in C return 
52 (\lambda (_: C).C) with [(CSort _) \Rightarrow e0 | (CHead c _ _) \Rightarrow 
53 c])) (CHead e0 (Bind b0) u0) (CHead e (Bind b) u) H1) in ((let H3 \def 
54 (f_equal C B (\lambda (e1: C).(match e1 in C return (\lambda (_: C).B) with 
55 [(CSort _) \Rightarrow b0 | (CHead _ k _) \Rightarrow (match k in K return 
56 (\lambda (_: K).B) with [(Bind b1) \Rightarrow b1 | (Flat _) \Rightarrow 
57 b0])])) (CHead e0 (Bind b0) u0) (CHead e (Bind b) u) H1) in ((let H4 \def 
58 (f_equal C T (\lambda (e1: C).(match e1 in C return (\lambda (_: C).T) with 
59 [(CSort _) \Rightarrow u0 | (CHead _ _ t) \Rightarrow t])) (CHead e0 (Bind 
60 b0) u0) (CHead e (Bind b) u) H1) in (\lambda (H5: (eq B b0 b)).(\lambda (H6: 
61 (eq C e0 e)).(eq_ind_r T u (\lambda (t: T).(eq C (CHead e0 (Bind b0) t) 
62 (CHead e0 (Bind b0) t))) (eq_ind_r C e (\lambda (c: C).(eq C (CHead c (Bind 
63 b0) u) (CHead c (Bind b0) u))) (eq_ind_r B b (\lambda (b1: B).(eq C (CHead e 
64 (Bind b1) u) (CHead e (Bind b1) u))) (refl_equal C (CHead e (Bind b) u)) b0 
65 H5) e0 H6) u0 H4)))) H3)) H2)))))) (\lambda (e0: C).(\lambda (c: C).(\lambda 
66 (_: (clear e0 c)).(\lambda (_: (((eq C e0 (CHead e (Bind b) u)) \to (eq C c 
67 e0)))).(\lambda (f: F).(\lambda (u0: T).(\lambda (H3: (eq C (CHead e0 (Flat 
68 f) u0) (CHead e (Bind b) u))).(let H4 \def (eq_ind C (CHead e0 (Flat f) u0) 
69 (\lambda (ee: C).(match ee in C return (\lambda (_: C).Prop) with [(CSort _) 
70 \Rightarrow False | (CHead _ k _) \Rightarrow (match k in K return (\lambda 
71 (_: K).Prop) with [(Bind _) \Rightarrow False | (Flat _) \Rightarrow 
72 True])])) I (CHead e (Bind b) u) H3) in (False_ind (eq C c (CHead e0 (Flat f) 
73 u0)) H4))))))))) y x H0))) H))))).
74 (* COMMENTS
75 Initial nodes: 525
76 END *)
77
78 theorem clear_gen_flat:
79  \forall (f: F).(\forall (e: C).(\forall (x: C).(\forall (u: T).((clear 
80 (CHead e (Flat f) u) x) \to (clear e x)))))
81 \def
82  \lambda (f: F).(\lambda (e: C).(\lambda (x: C).(\lambda (u: T).(\lambda (H: 
83 (clear (CHead e (Flat f) u) x)).(insert_eq C (CHead e (Flat f) u) (\lambda 
84 (c: C).(clear c x)) (\lambda (_: C).(clear e x)) (\lambda (y: C).(\lambda 
85 (H0: (clear y x)).(clear_ind (\lambda (c: C).(\lambda (c0: C).((eq C c (CHead 
86 e (Flat f) u)) \to (clear e c0)))) (\lambda (b: B).(\lambda (e0: C).(\lambda 
87 (u0: T).(\lambda (H1: (eq C (CHead e0 (Bind b) u0) (CHead e (Flat f) 
88 u))).(let H2 \def (eq_ind C (CHead e0 (Bind b) u0) (\lambda (ee: C).(match ee 
89 in C return (\lambda (_: C).Prop) with [(CSort _) \Rightarrow False | (CHead 
90 _ k _) \Rightarrow (match k in K return (\lambda (_: K).Prop) with [(Bind _) 
91 \Rightarrow True | (Flat _) \Rightarrow False])])) I (CHead e (Flat f) u) H1) 
92 in (False_ind (clear e (CHead e0 (Bind b) u0)) H2)))))) (\lambda (e0: 
93 C).(\lambda (c: C).(\lambda (H1: (clear e0 c)).(\lambda (H2: (((eq C e0 
94 (CHead e (Flat f) u)) \to (clear e c)))).(\lambda (f0: F).(\lambda (u0: 
95 T).(\lambda (H3: (eq C (CHead e0 (Flat f0) u0) (CHead e (Flat f) u))).(let H4 
96 \def (f_equal C C (\lambda (e1: C).(match e1 in C return (\lambda (_: C).C) 
97 with [(CSort _) \Rightarrow e0 | (CHead c0 _ _) \Rightarrow c0])) (CHead e0 
98 (Flat f0) u0) (CHead e (Flat f) u) H3) in ((let H5 \def (f_equal C F (\lambda 
99 (e1: C).(match e1 in C return (\lambda (_: C).F) with [(CSort _) \Rightarrow 
100 f0 | (CHead _ k _) \Rightarrow (match k in K return (\lambda (_: K).F) with 
101 [(Bind _) \Rightarrow f0 | (Flat f1) \Rightarrow f1])])) (CHead e0 (Flat f0) 
102 u0) (CHead e (Flat f) u) H3) in ((let H6 \def (f_equal C T (\lambda (e1: 
103 C).(match e1 in C return (\lambda (_: C).T) with [(CSort _) \Rightarrow u0 | 
104 (CHead _ _ t) \Rightarrow t])) (CHead e0 (Flat f0) u0) (CHead e (Flat f) u) 
105 H3) in (\lambda (_: (eq F f0 f)).(\lambda (H8: (eq C e0 e)).(let H9 \def 
106 (eq_ind C e0 (\lambda (c0: C).((eq C c0 (CHead e (Flat f) u)) \to (clear e 
107 c))) H2 e H8) in (let H10 \def (eq_ind C e0 (\lambda (c0: C).(clear c0 c)) H1 
108 e H8) in H10))))) H5)) H4))))))))) y x H0))) H))))).
109 (* COMMENTS
110 Initial nodes: 453
111 END *)
112
113 theorem clear_gen_flat_r:
114  \forall (f: F).(\forall (x: C).(\forall (e: C).(\forall (u: T).((clear x 
115 (CHead e (Flat f) u)) \to (\forall (P: Prop).P)))))
116 \def
117  \lambda (f: F).(\lambda (x: C).(\lambda (e: C).(\lambda (u: T).(\lambda (H: 
118 (clear x (CHead e (Flat f) u))).(\lambda (P: Prop).(insert_eq C (CHead e 
119 (Flat f) u) (\lambda (c: C).(clear x c)) (\lambda (_: C).P) (\lambda (y: 
120 C).(\lambda (H0: (clear x y)).(clear_ind (\lambda (_: C).(\lambda (c0: 
121 C).((eq C c0 (CHead e (Flat f) u)) \to P))) (\lambda (b: B).(\lambda (e0: 
122 C).(\lambda (u0: T).(\lambda (H1: (eq C (CHead e0 (Bind b) u0) (CHead e (Flat 
123 f) u))).(let H2 \def (eq_ind C (CHead e0 (Bind b) u0) (\lambda (ee: C).(match 
124 ee in C return (\lambda (_: C).Prop) with [(CSort _) \Rightarrow False | 
125 (CHead _ k _) \Rightarrow (match k in K return (\lambda (_: K).Prop) with 
126 [(Bind _) \Rightarrow True | (Flat _) \Rightarrow False])])) I (CHead e (Flat 
127 f) u) H1) in (False_ind P H2)))))) (\lambda (e0: C).(\lambda (c: C).(\lambda 
128 (H1: (clear e0 c)).(\lambda (H2: (((eq C c (CHead e (Flat f) u)) \to 
129 P))).(\lambda (_: F).(\lambda (_: T).(\lambda (H3: (eq C c (CHead e (Flat f) 
130 u))).(let H4 \def (eq_ind C c (\lambda (c0: C).((eq C c0 (CHead e (Flat f) 
131 u)) \to P)) H2 (CHead e (Flat f) u) H3) in (let H5 \def (eq_ind C c (\lambda 
132 (c0: C).(clear e0 c0)) H1 (CHead e (Flat f) u) H3) in (H4 (refl_equal C 
133 (CHead e (Flat f) u)))))))))))) x y H0))) H)))))).
134 (* COMMENTS
135 Initial nodes: 303
136 END *)
137
138 theorem clear_gen_all:
139  \forall (c1: C).(\forall (c2: C).((clear c1 c2) \to (ex_3 B C T (\lambda (b: 
140 B).(\lambda (e: C).(\lambda (u: T).(eq C c2 (CHead e (Bind b) u))))))))
141 \def
142  \lambda (c1: C).(\lambda (c2: C).(\lambda (H: (clear c1 c2)).(clear_ind 
143 (\lambda (_: C).(\lambda (c0: C).(ex_3 B C T (\lambda (b: B).(\lambda (e: 
144 C).(\lambda (u: T).(eq C c0 (CHead e (Bind b) u)))))))) (\lambda (b: 
145 B).(\lambda (e: C).(\lambda (u: T).(ex_3_intro B C T (\lambda (b0: 
146 B).(\lambda (e0: C).(\lambda (u0: T).(eq C (CHead e (Bind b) u) (CHead e0 
147 (Bind b0) u0))))) b e u (refl_equal C (CHead e (Bind b) u)))))) (\lambda (e: 
148 C).(\lambda (c: C).(\lambda (H0: (clear e c)).(\lambda (H1: (ex_3 B C T 
149 (\lambda (b: B).(\lambda (e0: C).(\lambda (u: T).(eq C c (CHead e0 (Bind b) 
150 u))))))).(\lambda (_: F).(\lambda (_: T).(let H2 \def H1 in (ex_3_ind B C T 
151 (\lambda (b: B).(\lambda (e0: C).(\lambda (u0: T).(eq C c (CHead e0 (Bind b) 
152 u0))))) (ex_3 B C T (\lambda (b: B).(\lambda (e0: C).(\lambda (u0: T).(eq C c 
153 (CHead e0 (Bind b) u0)))))) (\lambda (x0: B).(\lambda (x1: C).(\lambda (x2: 
154 T).(\lambda (H3: (eq C c (CHead x1 (Bind x0) x2))).(let H4 \def (eq_ind C c 
155 (\lambda (c0: C).(clear e c0)) H0 (CHead x1 (Bind x0) x2) H3) in (eq_ind_r C 
156 (CHead x1 (Bind x0) x2) (\lambda (c0: C).(ex_3 B C T (\lambda (b: B).(\lambda 
157 (e0: C).(\lambda (u0: T).(eq C c0 (CHead e0 (Bind b) u0))))))) (ex_3_intro B 
158 C T (\lambda (b: B).(\lambda (e0: C).(\lambda (u0: T).(eq C (CHead x1 (Bind 
159 x0) x2) (CHead e0 (Bind b) u0))))) x0 x1 x2 (refl_equal C (CHead x1 (Bind x0) 
160 x2))) c H3)))))) H2)))))))) c1 c2 H))).
161 (* COMMENTS
162 Initial nodes: 381
163 END *)
164