include "basic_1/A/defs.ma".
-let rec A_rect (P: (A \to Type[0])) (f: (\forall (n: nat).(\forall (n0:
-nat).(P (ASort n n0))))) (f0: (\forall (a: A).((P a) \to (\forall (a0: A).((P
-a0) \to (P (AHead a a0))))))) (a: A) on a: P a \def match a with [(ASort n
-n0) \Rightarrow (f n n0) | (AHead a0 a1) \Rightarrow (f0 a0 ((A_rect P f f0)
-a0) a1 ((A_rect P f f0) a1))].
+implied let rec A_rect (P: (A \to Type[0])) (f: (\forall (n: nat).(\forall
+(n0: nat).(P (ASort n n0))))) (f0: (\forall (a: A).((P a) \to (\forall (a0:
+A).((P a0) \to (P (AHead a a0))))))) (a: A) on a: P a \def match a with
+[(ASort n n0) \Rightarrow (f n n0) | (AHead a0 a1) \Rightarrow (f0 a0
+((A_rect P f f0) a0) a1 ((A_rect P f f0) a1))].
-theorem A_ind:
+implied lemma A_ind:
\forall (P: ((A \to Prop))).(((\forall (n: nat).(\forall (n0: nat).(P (ASort
n n0))))) \to (((\forall (a: A).((P a) \to (\forall (a0: A).((P a0) \to (P
(AHead a a0))))))) \to (\forall (a: A).(P a))))
include "basic_1/C/defs.ma".
-let rec C_rect (P: (C \to Type[0])) (f: (\forall (n: nat).(P (CSort n))))
-(f0: (\forall (c: C).((P c) \to (\forall (k: K).(\forall (t: T).(P (CHead c k
-t))))))) (c: C) on c: P c \def match c with [(CSort n) \Rightarrow (f n) |
-(CHead c0 k t) \Rightarrow (f0 c0 ((C_rect P f f0) c0) k t)].
+implied let rec C_rect (P: (C \to Type[0])) (f: (\forall (n: nat).(P (CSort
+n)))) (f0: (\forall (c: C).((P c) \to (\forall (k: K).(\forall (t: T).(P
+(CHead c k t))))))) (c: C) on c: P c \def match c with [(CSort n) \Rightarrow
+(f n) | (CHead c0 k t) \Rightarrow (f0 c0 ((C_rect P f f0) c0) k t)].
-theorem C_ind:
+implied lemma C_ind:
\forall (P: ((C \to Prop))).(((\forall (n: nat).(P (CSort n)))) \to
(((\forall (c: C).((P c) \to (\forall (k: K).(\forall (t: T).(P (CHead c k
t))))))) \to (\forall (c: C).(P c))))
\def
\lambda (P: ((C \to Prop))).(C_rect P).
-theorem clt_wf__q_ind:
+fact clt_wf__q_ind:
\forall (P: ((C \to Prop))).(((\forall (n: nat).((\lambda (P0: ((C \to
Prop))).(\lambda (n0: nat).(\forall (c: C).((eq nat (cweight c) n0) \to (P0
c))))) P n))) \to (\forall (c: C).(P c)))
n) \to (P c)))))).(\lambda (c: C).(H (cweight c) c (refl_equal nat (cweight
c)))))).
-theorem clt_wf_ind:
+lemma clt_wf_ind:
\forall (P: ((C \to Prop))).(((\forall (c: C).(((\forall (d: C).((clt d c)
\to (P d)))) \to (P c)))) \to (\forall (c: C).(P c)))
\def
include "basic_1/T/props.ma".
-theorem cle_r:
+lemma cle_r:
\forall (c: C).(cle c c)
\def
\lambda (c: C).(C_ind (\lambda (c0: C).(le (cweight c0) (cweight c0)))
(cweight c0))).(\lambda (_: K).(\lambda (t: T).(le_n (plus (cweight c0)
(tweight t))))))) c).
-theorem cle_head:
+lemma cle_head:
\forall (c1: C).(\forall (c2: C).((cle c1 c2) \to (\forall (u1: T).(\forall
(u2: T).((tle u1 u2) \to (\forall (k: K).(cle (CHead c1 k u1) (CHead c2 k
u2))))))))
(tweight u2))).(\lambda (_: K).(le_plus_plus (cweight c1) (cweight c2)
(tweight u1) (tweight u2) H H0))))))).
-theorem cle_trans_head:
+lemma cle_trans_head:
\forall (c1: C).(\forall (c2: C).((cle c1 c2) \to (\forall (k: K).(\forall
(u: T).(cle c1 (CHead c2 k u))))))
\def
c2))).(\lambda (_: K).(\lambda (u: T).(le_plus_trans (cweight c1) (cweight
c2) (tweight u) H))))).
-theorem clt_cong:
+lemma clt_cong:
\forall (c: C).(\forall (d: C).((clt c d) \to (\forall (k: K).(\forall (t:
T).(clt (CHead c k t) (CHead d k t))))))
\def
d))).(\lambda (_: K).(\lambda (t: T).(lt_reg_r (cweight c) (cweight d)
(tweight t) H))))).
-theorem clt_head:
+lemma clt_head:
\forall (k: K).(\forall (c: C).(\forall (u: T).(clt c (CHead c k u))))
\def
\lambda (_: K).(\lambda (c: C).(\lambda (u: T).(eq_ind_r nat (plus (cweight
c) O) (\lambda (n: nat).(lt n (plus (cweight c) (tweight u)))) (lt_reg_l O
(tweight u) (cweight c) (tweight_lt u)) (cweight c) (plus_n_O (cweight c))))).
-theorem chead_ctail:
+lemma chead_ctail:
\forall (c: C).(\forall (t: T).(\forall (k: K).(ex_3 K C T (\lambda (h:
K).(\lambda (d: C).(\lambda (u: T).(eq C (CHead c k t) (CTail h u d))))))))
\def
(CHead x1 k0 t0) x2 (refl_equal C (CHead (CTail x0 x2 x1) k0 t0))) (CHead c0
k t) H1))))) H0))))))))) c).
-theorem clt_thead:
+lemma clt_thead:
\forall (k: K).(\forall (u: T).(\forall (c: C).(clt c (CTail k u c))))
\def
\lambda (k: K).(\lambda (u: T).(\lambda (c: C).(C_ind (\lambda (c0: C).(clt
C).(\lambda (H: (clt c0 (CTail k u c0))).(\lambda (k0: K).(\lambda (t:
T).(clt_cong c0 (CTail k u c0) H k0 t))))) c))).
-theorem c_tail_ind:
+lemma c_tail_ind:
\forall (P: ((C \to Prop))).(((\forall (n: nat).(P (CSort n)))) \to
(((\forall (c: C).((P c) \to (\forall (k: K).(\forall (t: T).(P (CTail k t
c))))))) \to (\forall (c: C).(P c))))
include "basic_1/T/fwd.ma".
-theorem terms_props__bind_dec:
+fact terms_props__bind_dec:
\forall (b1: B).(\forall (b2: B).(or (eq B b1 b2) ((eq B b1 b2) \to (\forall
(P: Prop).P))))
\def
Void Void) ((eq B Void Void) \to (\forall (P: Prop).P)) (refl_equal B Void))
b2)) b1).
-theorem bind_dec_not:
+lemma bind_dec_not:
\forall (b1: B).(\forall (b2: B).(or (eq B b1 b2) (not (eq B b1 b2))))
\def
\lambda (b1: B).(\lambda (b2: B).(let H_x \def (terms_props__bind_dec b1 b2)
(((eq B b1 b2) \to (\forall (P: Prop).P)))).(or_intror (eq B b1 b2) ((eq B b1
b2) \to False) (\lambda (H1: (eq B b1 b2)).(H0 H1 False)))) H)))).
-theorem terms_props__flat_dec:
+fact terms_props__flat_dec:
\forall (f1: F).(\forall (f2: F).(or (eq F f1 f2) ((eq F f1 f2) \to (\forall
(P: Prop).P))))
\def
Cast) ((eq F Cast Cast) \to (\forall (P: Prop).P)) (refl_equal F Cast)) f2))
f1).
-theorem terms_props__kind_dec:
+fact terms_props__kind_dec:
\forall (k1: K).(\forall (k2: K).(or (eq K k1 k2) ((eq K k1 k2) \to (\forall
(P: Prop).P))))
\def
\to (\forall (P0: Prop).P0))) H0 f H2) in (H3 (refl_equal F f) P))))))) H))))
k2))) k1).
-theorem term_dec:
+lemma term_dec:
\forall (t1: T).(\forall (t2: T).(or (eq T t1 t2) ((eq T t1 t2) \to (\forall
(P: Prop).P))))
\def
(THead k t t0) t5) \to (\forall (P0: Prop).P0)))) H1 t H9) in (H12
(refl_equal T t) P))))))) H7)) H6)))))) H3)))))))) t2))))))) t1).
-theorem binder_dec:
+lemma binder_dec:
\forall (t: T).(or (ex_3 B T T (\lambda (b: B).(\lambda (w: T).(\lambda (u:
T).(eq T t (THead (Bind b) w u)))))) (\forall (b: B).(\forall (w: T).(\forall
(u: T).((eq T t (THead (Bind b) w u)) \to (\forall (P: Prop).P))))))
\Rightarrow False | (Flat _) \Rightarrow True])])) I (THead (Bind b) w u) H1)
in (False_ind P H2))))))))))))) k)) t).
-theorem abst_dec:
+lemma abst_dec:
\forall (u: T).(\forall (v: T).(or (ex T (\lambda (t: T).(eq T u (THead
(Bind Abst) v t)))) (\forall (t: T).((eq T u (THead (Bind Abst) v t)) \to
(\forall (P: Prop).P)))))
include "basic_1/T/defs.ma".
-let rec T_rect (P: (T \to Type[0])) (f: (\forall (n: nat).(P (TSort n))))
-(f0: (\forall (n: nat).(P (TLRef n)))) (f1: (\forall (k: K).(\forall (t:
-T).((P t) \to (\forall (t0: T).((P t0) \to (P (THead k t t0)))))))) (t: T) on
-t: P t \def match t with [(TSort n) \Rightarrow (f n) | (TLRef n) \Rightarrow
-(f0 n) | (THead k t0 t1) \Rightarrow (f1 k t0 ((T_rect P f f0 f1) t0) t1
-((T_rect P f f0 f1) t1))].
+implied let rec T_rect (P: (T \to Type[0])) (f: (\forall (n: nat).(P (TSort
+n)))) (f0: (\forall (n: nat).(P (TLRef n)))) (f1: (\forall (k: K).(\forall
+(t: T).((P t) \to (\forall (t0: T).((P t0) \to (P (THead k t t0)))))))) (t:
+T) on t: P t \def match t with [(TSort n) \Rightarrow (f n) | (TLRef n)
+\Rightarrow (f0 n) | (THead k t0 t1) \Rightarrow (f1 k t0 ((T_rect P f f0 f1)
+t0) t1 ((T_rect P f f0 f1) t1))].
-theorem T_ind:
+implied lemma T_ind:
\forall (P: ((T \to Prop))).(((\forall (n: nat).(P (TSort n)))) \to
(((\forall (n: nat).(P (TLRef n)))) \to (((\forall (k: K).(\forall (t: T).((P
t) \to (\forall (t0: T).((P t0) \to (P (THead k t t0)))))))) \to (\forall (t:
\def
\lambda (P: ((T \to Prop))).(T_rect P).
-theorem thead_x_y_y:
+lemma thead_x_y_y:
\forall (k: K).(\forall (v: T).(\forall (t: T).((eq T (THead k v t) t) \to
(\forall (P: Prop).P))))
\def
include "basic_1/T/fwd.ma".
-theorem not_abbr_abst:
+lemma not_abbr_abst:
not (eq B Abbr Abst)
\def
\lambda (H: (eq B Abbr Abst)).(let H0 \def (eq_ind B Abbr (\lambda (ee:
B).(match ee with [Abbr \Rightarrow True | Abst \Rightarrow False | Void
\Rightarrow False])) I Abst H) in (False_ind False H0)).
-theorem not_void_abst:
+lemma not_void_abst:
not (eq B Void Abst)
\def
\lambda (H: (eq B Void Abst)).(let H0 \def (eq_ind B Void (\lambda (ee:
B).(match ee with [Abbr \Rightarrow False | Abst \Rightarrow False | Void
\Rightarrow True])) I Abst H) in (False_ind False H0)).
-theorem not_abbr_void:
+lemma not_abbr_void:
not (eq B Abbr Void)
\def
\lambda (H: (eq B Abbr Void)).(let H0 \def (eq_ind B Abbr (\lambda (ee:
B).(match ee with [Abbr \Rightarrow True | Abst \Rightarrow False | Void
\Rightarrow False])) I Void H) in (False_ind False H0)).
-theorem not_abst_void:
+lemma not_abst_void:
not (eq B Abst Void)
\def
\lambda (H: (eq B Abst Void)).(let H0 \def (eq_ind B Abst (\lambda (ee:
B).(match ee with [Abbr \Rightarrow False | Abst \Rightarrow True | Void
\Rightarrow False])) I Void H) in (False_ind False H0)).
-theorem tweight_lt:
+lemma tweight_lt:
\forall (t: T).(lt O (tweight t))
\def
\lambda (t: T).(T_ind (\lambda (t0: T).(lt O (tweight t0))) (\lambda (_:
(tweight t1))).(le_S (S O) (plus (tweight t0) (tweight t1)) (le_plus_trans (S
O) (tweight t0) (tweight t1) H))))))) t).
-theorem tle_r:
+lemma tle_r:
\forall (t: T).(tle t t)
\def
\lambda (t: T).(T_ind (\lambda (t0: T).(le (tweight t0) (tweight t0)))
include "basic_1/next_plus/props.ma".
-theorem aplus_reg_r:
+lemma aplus_reg_r:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).(\forall (h1: nat).(\forall
(h2: nat).((eq A (aplus g a1 h1) (aplus g a2 h2)) \to (\forall (h: nat).(eq A
(aplus g a1 (plus h h1)) (aplus g a2 (plus h h2)))))))))
h1)) (aplus g a2 (plus n h2)))).(f_equal2 G A A asucc g g (aplus g a1 (plus n
h1)) (aplus g a2 (plus n h2)) (refl_equal G g) H0))) h))))))).
-theorem aplus_assoc:
+lemma aplus_assoc:
\forall (g: G).(\forall (a: A).(\forall (h1: nat).(\forall (h2: nat).(eq A
(aplus g (aplus g a h1) h2) (aplus g a (plus h1 h2))))))
\def
n0) (asucc g (aplus g a (plus n n0))) (refl_equal G g) H0) (plus n (S n0))
(plus_n_Sm n n0)))) h2)))) h1))).
-theorem aplus_asucc:
+lemma aplus_asucc:
\forall (g: G).(\forall (h: nat).(\forall (a: A).(eq A (aplus g (asucc g a)
h) (asucc g (aplus g a h)))))
\def
(refl_equal A (asucc g (aplus g a h))) (aplus g (aplus g a (S O)) h)
(aplus_assoc g a (S O) h)))).
-theorem aplus_sort_O_S_simpl:
+lemma aplus_sort_O_S_simpl:
\forall (g: G).(\forall (n: nat).(\forall (k: nat).(eq A (aplus g (ASort O
n) (S k)) (aplus g (ASort O (next g n)) k))))
\def
(refl_equal A (aplus g (ASort O (next g n)) k)) (asucc g (aplus g (ASort O n)
k)) (aplus_asucc g k (ASort O n))))).
-theorem aplus_sort_S_S_simpl:
+lemma aplus_sort_S_S_simpl:
\forall (g: G).(\forall (n: nat).(\forall (h: nat).(\forall (k: nat).(eq A
(aplus g (ASort (S h) n) (S k)) (aplus g (ASort h n) k)))))
\def
(ASort h n) k))) (refl_equal A (aplus g (ASort h n) k)) (asucc g (aplus g
(ASort (S h) n) k)) (aplus_asucc g k (ASort (S h) n)))))).
-theorem aplus_asort_O_simpl:
+lemma aplus_asort_O_simpl:
\forall (g: G).(\forall (h: nat).(\forall (n: nat).(eq A (aplus g (ASort O
n) h) (ASort O (next_plus g n h)))))
\def
(next_plus_next g n0 n)) (asucc g (aplus g (ASort O n0) n)) (aplus_asucc g n
(ASort O n0)))))) h)).
-theorem aplus_asort_le_simpl:
+lemma aplus_asort_le_simpl:
\forall (g: G).(\forall (h: nat).(\forall (k: nat).(\forall (n: nat).((le h
k) \to (eq A (aplus g (ASort k n) h) (ASort (minus k h) n))))))
\def
g (aplus g (ASort (S n) n0) h0)) (aplus_asucc g h0 (ASort (S n) n0)))))))
k)))) h)).
-theorem aplus_asort_simpl:
+lemma aplus_asort_simpl:
\forall (g: G).(\forall (h: nat).(\forall (k: nat).(\forall (n: nat).(eq A
(aplus g (ASort k n) h) (ASort (minus k h) (next_plus g n (minus h k)))))))
\def
(next_plus g n O))) (minus h k) (O_minus h k H)) (aplus g (ASort k n) h)
(aplus_asort_le_simpl g h k n H))))))).
-theorem aplus_ahead_simpl:
+lemma aplus_ahead_simpl:
\forall (g: G).(\forall (h: nat).(\forall (a1: A).(\forall (a2: A).(eq A
(aplus g (AHead a1 a2) h) (AHead a1 (aplus g a2 h))))))
\def
a2)) (asucc g (aplus g (AHead a1 a2) n)) (aplus_asucc g n (AHead a1 a2)))))))
h)).
-theorem aplus_asucc_false:
+lemma aplus_asucc_false:
\forall (g: G).(\forall (a: A).(\forall (h: nat).((eq A (aplus g (asucc g a)
h) a) \to (\forall (P: Prop).P))))
\def
(asucc g a1) h) | (AHead _ a2) \Rightarrow a2])) (AHead a0 (aplus g (asucc g
a1) h)) (AHead a0 a1) H2) in (H0 h H3 P)))))))))) a)).
-theorem aplus_inj:
+lemma aplus_inj:
\forall (g: G).(\forall (h1: nat).(\forall (h2: nat).(\forall (a: A).((eq A
(aplus g a h1) (aplus g a h2)) \to (eq nat h1 h2)))))
\def
include "basic_1/aprem/defs.ma".
-let rec aprem_ind (P: (nat \to (A \to (A \to Prop)))) (f: (\forall (a1:
-A).(\forall (a2: A).(P O (AHead a1 a2) a1)))) (f0: (\forall (a2: A).(\forall
-(a: A).(\forall (i: nat).((aprem i a2 a) \to ((P i a2 a) \to (\forall (a1:
-A).(P (S i) (AHead a1 a2) a)))))))) (n: nat) (a: A) (a0: A) (a1: aprem n a
-a0) on a1: P n a a0 \def match a1 with [(aprem_zero a2 a3) \Rightarrow (f a2
-a3) | (aprem_succ a2 a3 i a4 a5) \Rightarrow (f0 a2 a3 i a4 ((aprem_ind P f
-f0) i a2 a3 a4) a5)].
+implied let rec aprem_ind (P: (nat \to (A \to (A \to Prop)))) (f: (\forall
+(a1: A).(\forall (a2: A).(P O (AHead a1 a2) a1)))) (f0: (\forall (a2:
+A).(\forall (a: A).(\forall (i: nat).((aprem i a2 a) \to ((P i a2 a) \to
+(\forall (a1: A).(P (S i) (AHead a1 a2) a)))))))) (n: nat) (a: A) (a0: A)
+(a1: aprem n a a0) on a1: P n a a0 \def match a1 with [(aprem_zero a2 a3)
+\Rightarrow (f a2 a3) | (aprem_succ a2 a3 i a4 a5) \Rightarrow (f0 a2 a3 i a4
+((aprem_ind P f f0) i a2 a3 a4) a5)].
-theorem aprem_gen_sort:
+lemma aprem_gen_sort:
\forall (x: A).(\forall (i: nat).(\forall (h: nat).(\forall (n: nat).((aprem
i (ASort h n) x) \to False))))
\def
False | (AHead _ _) \Rightarrow True])) I (ASort h n) H3) in (False_ind False
H4))))))))) i y x H0))) H))))).
-theorem aprem_gen_head_O:
+lemma aprem_gen_head_O:
\forall (a1: A).(\forall (a2: A).(\forall (x: A).((aprem O (AHead a1 a2) x)
\to (eq A x a1))))
\def
with [O \Rightarrow False | (S _) \Rightarrow True])) I O H4) in (False_ind
(eq A a a1) H11)))))) H6)))))))))) y0 y x H1))) H0))) H)))).
-theorem aprem_gen_head_S:
+lemma aprem_gen_head_S:
\forall (a1: A).(\forall (a2: A).(\forall (x: A).(\forall (i: nat).((aprem
(S i) (AHead a1 a2) x) \to (aprem i a2 x)))))
\def
include "basic_1/leq/fwd.ma".
-theorem aprem_repl:
+lemma aprem_repl:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).((leq g a1 a2) \to (\forall
(i: nat).(\forall (b2: A).((aprem i a2 b2) \to (ex2 A (\lambda (b1: A).(leq g
b1 b2)) (\lambda (b1: A).(aprem i a1 b1)))))))))
a4) b1)) x H7 (aprem_succ a4 x i0 H8 a0))))) H6))))))) i H4)))))))))))) a1 a2
H)))).
-theorem aprem_asucc:
+lemma aprem_asucc:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).(\forall (i: nat).((aprem i
a1 a2) \to (aprem i (asucc g a1) a2)))))
\def
include "basic_1/aprem/props.ma".
-theorem arity_aprem:
+lemma arity_aprem:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (a: A).((arity g c t
a) \to (\forall (i: nat).(\forall (b: A).((aprem i a b) \to (ex2_3 C T nat
(\lambda (d: C).(\lambda (_: T).(\lambda (j: nat).(drop (plus i j) O d c))))
include "basic_1/cimp/props.ma".
-theorem arity_cimp_conf:
+lemma arity_cimp_conf:
\forall (g: G).(\forall (c1: C).(\forall (t: T).(\forall (a: A).((arity g c1
t a) \to (\forall (c2: C).((cimp c1 c2) \to (arity g c2 t a)))))))
\def
include "basic_1/getl/drop.ma".
-let rec arity_ind (g: G) (P: (C \to (T \to (A \to Prop)))) (f: (\forall (c:
-C).(\forall (n: nat).(P c (TSort n) (ASort O n))))) (f0: (\forall (c:
-C).(\forall (d: C).(\forall (u: T).(\forall (i: nat).((getl i c (CHead d
-(Bind Abbr) u)) \to (\forall (a: A).((arity g d u a) \to ((P d u a) \to (P c
-(TLRef i) a)))))))))) (f1: (\forall (c: C).(\forall (d: C).(\forall (u:
-T).(\forall (i: nat).((getl i c (CHead d (Bind Abst) u)) \to (\forall (a:
+implied let rec arity_ind (g: G) (P: (C \to (T \to (A \to Prop)))) (f:
+(\forall (c: C).(\forall (n: nat).(P c (TSort n) (ASort O n))))) (f0:
+(\forall (c: C).(\forall (d: C).(\forall (u: T).(\forall (i: nat).((getl i c
+(CHead d (Bind Abbr) u)) \to (\forall (a: A).((arity g d u a) \to ((P d u a)
+\to (P c (TLRef i) a)))))))))) (f1: (\forall (c: C).(\forall (d: C).(\forall
+(u: T).(\forall (i: nat).((getl i c (CHead d (Bind Abst) u)) \to (\forall (a:
A).((arity g d u (asucc g a)) \to ((P d u (asucc g a)) \to (P c (TLRef i)
a)))))))))) (f2: (\forall (b: B).((not (eq B b Abst)) \to (\forall (c:
C).(\forall (u: T).(\forall (a1: A).((arity g c u a1) \to ((P c u a1) \to
a2 a3 l) \Rightarrow (f6 c0 t0 a1 a2 ((arity_ind g P f f0 f1 f2 f3 f4 f5 f6)
c0 t0 a1 a2) a3 l)].
-theorem arity_gen_sort:
+lemma arity_gen_sort:
\forall (g: G).(\forall (c: C).(\forall (n: nat).(\forall (a: A).((arity g c
(TSort n) a) \to (leq g a (ASort O n))))))
\def
(leq_trans g a2 a1 (leq_sym g a1 a2 H3) (ASort O n) (H6 (refl_equal T (TSort
n))))))))))))))) c y a H0))) H))))).
-theorem arity_gen_lref:
+lemma arity_gen_lref:
\forall (g: G).(\forall (c: C).(\forall (i: nat).(\forall (a: A).((arity g c
(TLRef i) a) \to (or (ex2_2 C T (\lambda (d: C).(\lambda (u: T).(getl i c
(CHead d (Bind Abbr) u)))) (\lambda (d: C).(\lambda (u: T).(arity g d u a))))
(asucc g a2)))) x0 x1 H10 (arity_repl g x0 x1 (asucc g a1) H11 (asucc g a2)
(asucc_repl g a1 a2 H3)))))))) H9)) H8))))))))))))) c y a H0))) H))))).
-theorem arity_gen_bind:
+lemma arity_gen_bind:
\forall (b: B).((not (eq B b Abst)) \to (\forall (g: G).(\forall (c:
C).(\forall (u: T).(\forall (t: T).(\forall (a2: A).((arity g c (THead (Bind
b) u t) a2) \to (ex2 A (\lambda (a1: A).(arity g c u a1)) (\lambda (_:
(arity_repl g (CHead c0 (Bind b) u) t a1 H11 a0 H4))))) H9))))))))))))) c y
a2 H1))) H0)))))))).
-theorem arity_gen_abst:
+lemma arity_gen_abst:
\forall (g: G).(\forall (c: C).(\forall (u: T).(\forall (t: T).(\forall (a:
A).((arity g c (THead (Bind Abst) u t) a) \to (ex3_2 A A (\lambda (a1:
A).(\lambda (a2: A).(eq A a (AHead a1 a2)))) (\lambda (a1: A).(\lambda (_:
(arity_repl g (CHead c0 (Bind Abst) u) t x1 H11 x3 H16)) a2 H18)))))))
H14)))))))))) H8))))))))))))) c y a H0))) H)))))).
-theorem arity_gen_appl:
+lemma arity_gen_appl:
\forall (g: G).(\forall (c: C).(\forall (u: T).(\forall (t: T).(\forall (a2:
A).((arity g c (THead (Flat Appl) u t) a2) \to (ex2 A (\lambda (a1: A).(arity
g c u a1)) (\lambda (a1: A).(arity g c t (AHead a1 a2)))))))))
(AHead x a0) (leq_head g x x (leq_refl g x) a1 a0 H3)))))) H8))))))))))))) c
y a2 H0))) H)))))).
-theorem arity_gen_cast:
+lemma arity_gen_cast:
\forall (g: G).(\forall (c: C).(\forall (u: T).(\forall (t: T).(\forall (a:
A).((arity g c (THead (Flat Cast) u t) a) \to (land (arity g c u (asucc g a))
(arity g c t a)))))))
g a1 a2 H3)) (arity_repl g c0 t a1 H10 a2 H3)))) H8))))))))))))) c y a H0)))
H)))))).
-theorem arity_gen_appls:
+lemma arity_gen_appls:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (vs: TList).(\forall
(a2: A).((arity g c (THeads (Flat Appl) vs t) a2) \to (ex A (\lambda (a:
A).(arity g c t a))))))))
(\lambda (x0: A).(\lambda (H5: (arity g c t x0)).(ex_intro A (\lambda (a:
A).(arity g c t a)) x0 H5))) H4)))))) H1))))))) vs)))).
-theorem arity_gen_lift:
+lemma arity_gen_lift:
\forall (g: G).(\forall (c1: C).(\forall (t: T).(\forall (a: A).(\forall (h:
nat).(\forall (d: nat).((arity g c1 (lift h d t) a) \to (\forall (c2:
C).((drop h d c1 c2) \to (arity g c2 t a)))))))))
include "basic_1/drop1/fwd.ma".
-theorem arity_lift1:
+lemma arity_lift1:
\forall (g: G).(\forall (a: A).(\forall (c2: C).(\forall (hds:
PList).(\forall (c1: C).(\forall (t: T).((drop1 hds c1 c2) \to ((arity g c2 t
a) \to (arity g c1 (lift1 hds t) a))))))))
include "basic_1/arity/subst0.ma".
-theorem arity_sred_wcpr0_pr0:
+lemma arity_sred_wcpr0_pr0:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (a: A).((arity g
c1 t1 a) \to (\forall (c2: C).((wcpr0 c1 c2) \to (\forall (t2: T).((pr0 t1
t2) \to (arity g c2 t2 a)))))))))
(H4: (pr0 t t2)).(arity_repl g c2 t2 a1 (H1 c2 H3 t2 H4) a2 H2)))))))))))) c1
t1 a H))))).
-theorem arity_sred_wcpr0_pr1:
+lemma arity_sred_wcpr0_pr1:
\forall (t1: T).(\forall (t2: T).((pr1 t1 t2) \to (\forall (g: G).(\forall
(c1: C).(\forall (a: A).((arity g c1 t1 a) \to (\forall (c2: C).((wcpr0 c1
c2) \to (arity g c2 t2 a)))))))))
(arity_sred_wcpr0_pr0 g c1 t4 a H3 c2 H4 t3 H0) c2 (wcpr0_refl
c2)))))))))))))) t1 t2 H))).
-theorem arity_sred_pr2:
+lemma arity_sred_pr2:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall
(g: G).(\forall (a: A).((arity g c t1 a) \to (arity g c t2 a)))))))
\def
(arity_sred_wcpr0_pr0 g c0 t3 a H3 c0 (wcpr0_refl c0) t4 H1) d u i H0 t
H2)))))))))))))) c t1 t2 H)))).
-theorem arity_sred_pr3:
+lemma arity_sred_pr3:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr3 c t1 t2) \to (\forall
(g: G).(\forall (a: A).((arity g c t1 a) \to (arity g c t2 a)))))))
\def
include "basic_1/arity/fwd.ma".
-theorem node_inh:
+lemma node_inh:
\forall (g: G).(\forall (n: nat).(\forall (k: nat).(ex_2 C T (\lambda (c:
C).(\lambda (t: T).(arity g c t (ASort k n)))))))
\def
(TLRef O) (arity_abst g (CHead x0 (Bind Abst) x1) x0 x1 O (getl_refl Abst x0
x1) (ASort (S n0) n) H1))))) H0)))) k))).
-theorem arity_lift:
+lemma arity_lift:
\forall (g: G).(\forall (c2: C).(\forall (t: T).(\forall (a: A).((arity g c2
t a) \to (\forall (c1: C).(\forall (h: nat).(\forall (d: nat).((drop h d c1
c2) \to (arity g c1 (lift h d t) a)))))))))
c)).(arity_repl g c1 (lift h d t0) a1 (H1 c1 h d H3) a2 H2)))))))))))) c2 t a
H))))).
-theorem arity_repellent:
+lemma arity_repellent:
\forall (g: G).(\forall (c: C).(\forall (w: T).(\forall (t: T).(\forall (a1:
A).((arity g (CHead c (Bind Abst) w) t a1) \to (\forall (a2: A).((arity g c
(THead (Bind Abst) w t) a2) \to ((leq g a1 a2) \to (\forall (P:
g a) (asucc g a) (leq_refl g (asucc g a)))) (asucc g (AHead x a)) (leq_refl g
(asucc g (AHead x a)))) H4))))) H5))))) H2)))))))) vs))))).
-theorem arity_appls_abbr:
+lemma arity_appls_abbr:
\forall (g: G).(\forall (c: C).(\forall (d: C).(\forall (v: T).(\forall (i:
nat).((getl i c (CHead d (Bind Abbr) v)) \to (\forall (vs: TList).(\forall
(a: A).((arity g c (THeads (Flat Appl) vs (lift (S i) O v)) a) \to (arity g c
include "basic_1/getl/getl.ma".
-theorem arity_gen_cvoid_subst0:
+lemma arity_gen_cvoid_subst0:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (a: A).((arity g c t
a) \to (\forall (d: C).(\forall (u: T).(\forall (i: nat).((getl i c (CHead d
(Bind Void) u)) \to (\forall (w: T).(\forall (v: T).((subst0 i w t v) \to
T).(\lambda (v: T).(\lambda (H4: (subst0 i w t0 v)).(\lambda (P: Prop).(H1 d
u i H3 w v H4 P)))))))))))))))) c t a H))))).
-theorem arity_gen_cvoid:
+lemma arity_gen_cvoid:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (a: A).((arity g c t
a) \to (\forall (d: C).(\forall (u: T).(\forall (i: nat).((getl i c (CHead d
(Bind Void) u)) \to (ex T (\lambda (v: T).(eq T t (lift (S O) i v))))))))))))
(ex_intro T (\lambda (v: T).(eq T (lift (S O) i x) (lift (S O) i v))) x
(refl_equal T (lift (S O) i x))) t H3))) H2))) H1))))))))))).
-theorem arity_fsubst0:
+lemma arity_fsubst0:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (a: A).((arity g
c1 t1 a) \to (\forall (d1: C).(\forall (u: T).(\forall (i: nat).((getl i c1
(CHead d1 (Bind Abbr) u)) \to (\forall (c2: C).(\forall (t2: T).((fsubst0 i u
c2)).(arity_repl g c2 t2 a1 (H1 d1 u i H3 c2 t2 (fsubst0_both i u c t t2 H7
c2 H8)) a2 H2))) H6)) H5))))))))))))))))) c1 t1 a H))))).
-theorem arity_subst0:
+lemma arity_subst0:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (a: A).((arity g c
t1 a) \to (\forall (d: C).(\forall (u: T).(\forall (i: nat).((getl i c (CHead
d (Bind Abbr) u)) \to (\forall (t2: T).((subst0 i u t1 t2) \to (arity g c t2
include "basic_1/A/fwd.ma".
-theorem asucc_gen_sort:
+lemma asucc_gen_sort:
\forall (g: G).(\forall (h: nat).(\forall (n: nat).(\forall (a: A).((eq A
(ASort h n) (asucc g a)) \to (ex_2 nat nat (\lambda (h0: nat).(\lambda (n0:
nat).(eq A a (ASort h0 n0)))))))))
H1) in (False_ind (ex_2 nat nat (\lambda (h0: nat).(\lambda (n0: nat).(eq A
(AHead a0 a1) (ASort h0 n0))))) H2))))))) a)))).
-theorem asucc_gen_head:
+lemma asucc_gen_head:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).(\forall (a: A).((eq A
(AHead a1 a2) (asucc g a)) \to (ex2 A (\lambda (a0: A).(eq A a (AHead a1
a0))) (\lambda (a0: A).(eq A a2 (asucc g a0))))))))
include "basic_1/getl/getl.ma".
-theorem cimp_flat_sx:
+lemma cimp_flat_sx:
\forall (f: F).(\forall (c: C).(\forall (v: T).(cimp (CHead c (Flat f) v)
c)))
\def
b) w))).(ex_intro C (\lambda (d2: C).(getl (S h0) c (CHead d2 (Bind b) w)))
d1 (getl_gen_S (Flat f) c (CHead d1 (Bind b) w) v h0 H0))))) h H)))))))).
-theorem cimp_flat_dx:
+lemma cimp_flat_dx:
\forall (f: F).(\forall (c: C).(\forall (v: T).(cimp c (CHead c (Flat f)
v))))
\def
b) w))).(ex_intro C (\lambda (d2: C).(getl h (CHead c (Flat f) v) (CHead d2
(Bind b) w))) d1 (getl_flat c (CHead d1 (Bind b) w) h H f v))))))))).
-theorem cimp_bind:
+lemma cimp_bind:
\forall (c1: C).(\forall (c2: C).((cimp c1 c2) \to (\forall (b: B).(\forall
(v: T).(cimp (CHead c1 (Bind b) v) (CHead c2 (Bind b) v))))))
\def
(CHead c2 (Bind b) v) (CHead d2 (Bind b0) w))) x (getl_head (Bind b) h0 c2
(CHead x (Bind b0) w) H3 v)))) H2)))))) h H0)))))))))).
-theorem cimp_getl_conf:
+lemma cimp_getl_conf:
\forall (c1: C).(\forall (c2: C).((cimp c1 c2) \to (\forall (b: B).(\forall
(d1: C).(\forall (w: T).(\forall (i: nat).((getl i c1 (CHead d1 (Bind b) w))
\to (ex2 C (\lambda (d2: C).(cimp d1 d2)) (\lambda (d2: C).(getl i c2 (CHead
include "basic_1/drop/fwd.ma".
-theorem drop_clear:
+lemma drop_clear:
\forall (c1: C).(\forall (c2: C).(\forall (i: nat).((drop (S i) O c1 c2) \to
(ex2_3 B C T (\lambda (b: B).(\lambda (e: C).(\lambda (v: T).(clear c1 (CHead
e (Bind b) v))))) (\lambda (_: B).(\lambda (e: C).(\lambda (_: T).(drop i O e
(CHead x1 (Bind x0) x2) H3 f t) H4)))))) H2)))) k (drop_gen_drop k c c2 t i
H0))))))))) c1).
-theorem drop_clear_O:
+lemma drop_clear_O:
\forall (b: B).(\forall (c: C).(\forall (e1: C).(\forall (u: T).((clear c
(CHead e1 (Bind b) u)) \to (\forall (e2: C).(\forall (i: nat).((drop i O e1
e2) \to (drop (S i) O c e2))))))))
(Bind b) u))).(drop_drop (Flat f) i c0 e2 (H e1 u (clear_gen_flat f c0 (CHead
e1 (Bind b) u) t H2) e2 i H1) t))) k H0))))))))))) c)).
-theorem drop_clear_S:
+lemma drop_clear_S:
\forall (x2: C).(\forall (x1: C).(\forall (h: nat).(\forall (d: nat).((drop
h (S d) x1 x2) \to (\forall (b: B).(\forall (c2: C).(\forall (u: T).((clear
x2 (CHead c2 (Bind b) u)) \to (ex2 C (\lambda (c1: C).(clear x1 (CHead c1
include "basic_1/C/fwd.ma".
-let rec clear_ind (P: (C \to (C \to Prop))) (f: (\forall (b: B).(\forall (e:
-C).(\forall (u: T).(P (CHead e (Bind b) u) (CHead e (Bind b) u)))))) (f0:
-(\forall (e: C).(\forall (c: C).((clear e c) \to ((P e c) \to (\forall (f0:
-F).(\forall (u: T).(P (CHead e (Flat f0) u) c)))))))) (c: C) (c0: C) (c1:
-clear c c0) on c1: P c c0 \def match c1 with [(clear_bind b e u) \Rightarrow
-(f b e u) | (clear_flat e c2 c3 f1 u) \Rightarrow (f0 e c2 c3 ((clear_ind P f
-f0) e c2 c3) f1 u)].
+implied let rec clear_ind (P: (C \to (C \to Prop))) (f: (\forall (b:
+B).(\forall (e: C).(\forall (u: T).(P (CHead e (Bind b) u) (CHead e (Bind b)
+u)))))) (f0: (\forall (e: C).(\forall (c: C).((clear e c) \to ((P e c) \to
+(\forall (f0: F).(\forall (u: T).(P (CHead e (Flat f0) u) c)))))))) (c: C)
+(c0: C) (c1: clear c c0) on c1: P c c0 \def match c1 with [(clear_bind b e u)
+\Rightarrow (f b e u) | (clear_flat e c2 c3 f1 u) \Rightarrow (f0 e c2 c3
+((clear_ind P f f0) e c2 c3) f1 u)].
-theorem clear_gen_sort:
+lemma clear_gen_sort:
\forall (x: C).(\forall (n: nat).((clear (CSort n) x) \to (\forall (P:
Prop).P)))
\def
[(CSort _) \Rightarrow False | (CHead _ _ _) \Rightarrow True])) I (CSort n)
H3) in (False_ind P H4))))))))) y x H0))) H)))).
-theorem clear_gen_bind:
+lemma clear_gen_bind:
\forall (b: B).(\forall (e: C).(\forall (x: C).(\forall (u: T).((clear
(CHead e (Bind b) u) x) \to (eq C x (CHead e (Bind b) u))))))
\def
b) u) H3) in (False_ind (eq C c (CHead e0 (Flat f) u0)) H4))))))))) y x H0)))
H))))).
-theorem clear_gen_flat:
+lemma clear_gen_flat:
\forall (f: F).(\forall (e: C).(\forall (x: C).(\forall (u: T).((clear
(CHead e (Flat f) u) x) \to (clear e x)))))
\def
(\lambda (c0: C).(clear c0 c)) H1 e H8) in H10))))) H5)) H4))))))))) y x
H0))) H))))).
-theorem clear_gen_flat_r:
+lemma clear_gen_flat_r:
\forall (f: F).(\forall (x: C).(\forall (e: C).(\forall (u: T).((clear x
(CHead e (Flat f) u)) \to (\forall (P: Prop).P)))))
\def
(c0: C).(clear e0 c0)) H1 (CHead e (Flat f) u) H3) in (H4 (refl_equal C
(CHead e (Flat f) u)))))))))))) x y H0))) H)))))).
-theorem clear_gen_all:
+lemma clear_gen_all:
\forall (c1: C).(\forall (c2: C).((clear c1 c2) \to (ex_3 B C T (\lambda (b:
B).(\lambda (e: C).(\lambda (u: T).(eq C c2 (CHead e (Bind b) u))))))))
\def
c1)).(\lambda (H3: (clear (CHead c0 (Flat f) t) c2)).(H c1 (clear_gen_flat f
c0 c1 t H2) c2 (clear_gen_flat f c0 c2 t H3))))) k H0 H1))))))))) c).
-theorem clear_cle:
+lemma clear_cle:
\forall (c1: C).(\forall (c2: C).((clear c1 c2) \to (cle c2 c1)))
\def
\lambda (c1: C).(C_ind (\lambda (c: C).(\forall (c2: C).((clear c c2) \to
include "basic_1/clear/fwd.ma".
-theorem clear_clear:
+lemma clear_clear:
\forall (c1: C).(\forall (c2: C).((clear c1 c2) \to (clear c2 c2)))
\def
\lambda (c1: C).(C_ind (\lambda (c: C).(\forall (c2: C).((clear c c2) \to
(\lambda (f: F).(\lambda (H2: (clear (CHead c (Flat f) t) c0)).(clear_flat c
c2 (H c0 (clear_gen_flat f c c0 t H2) c2 H1) f t))) k H0))))))))) c1).
-theorem clear_ctail:
+lemma clear_ctail:
\forall (b: B).(\forall (c1: C).(\forall (c2: C).(\forall (u2: T).((clear c1
(CHead c2 (Bind b) u2)) \to (\forall (k: K).(\forall (u1: T).(clear (CTail k
u1 c1) (CHead (CTail k u1 c2) (Bind b) u2))))))))
include "basic_1/getl/props.ma".
-theorem getl_ctail_clen:
+lemma getl_ctail_clen:
\forall (b: B).(\forall (t: T).(\forall (c: C).(ex nat (\lambda (n:
nat).(getl (clen c) (CTail (Bind b) t c) (CHead (CSort n) (Bind b) t))))))
\def
c0) (Flat f) t0) (CHead (CSort n) (Bind b) t))) x (getl_flat (CTail (Bind b)
t c0) (CHead (CSort x) (Bind b) t) (clen c0) H1 f t0))) k))) H0)))))) c))).
-theorem getl_gen_tail:
+lemma getl_gen_tail:
\forall (k: K).(\forall (b: B).(\forall (u1: T).(\forall (u2: T).(\forall
(c2: C).(\forall (c1: C).(\forall (i: nat).((getl i (CTail k u1 c1) (CHead c2
(Bind b) u2)) \to (or (ex2 C (\lambda (e: C).(eq C c2 (CTail k u1 e)))
include "basic_1/cnt/defs.ma".
-let rec cnt_ind (P: (T \to Prop)) (f: (\forall (n: nat).(P (TSort n)))) (f0:
-(\forall (t: T).((cnt t) \to ((P t) \to (\forall (k: K).(\forall (v: T).(P
-(THead k v t)))))))) (t: T) (c: cnt t) on c: P t \def match c with [(cnt_sort
-n) \Rightarrow (f n) | (cnt_head t0 c0 k v) \Rightarrow (f0 t0 c0 ((cnt_ind P
-f f0) t0 c0) k v)].
+implied let rec cnt_ind (P: (T \to Prop)) (f: (\forall (n: nat).(P (TSort
+n)))) (f0: (\forall (t: T).((cnt t) \to ((P t) \to (\forall (k: K).(\forall
+(v: T).(P (THead k v t)))))))) (t: T) (c: cnt t) on c: P t \def match c with
+[(cnt_sort n) \Rightarrow (f n) | (cnt_head t0 c0 k v) \Rightarrow (f0 t0 c0
+((cnt_ind P f f0) t0 c0) k v)].
include "basic_1/lift/props.ma".
-theorem cnt_lift:
+lemma cnt_lift:
\forall (t: T).((cnt t) \to (\forall (i: nat).(\forall (d: nat).(cnt (lift i
d t)))))
\def
include "basic_1/csubv/getl.ma".
-theorem csuba_arity:
+lemma csuba_arity:
\forall (g: G).(\forall (c1: C).(\forall (t: T).(\forall (a: A).((arity g c1
t a) \to (\forall (c2: C).((csuba g c1 c2) \to (arity g c2 t a)))))))
\def
C).(\lambda (H3: (csuba g c c2)).(arity_repl g c2 t0 a1 (H1 c2 H3) a2
H2)))))))))) c1 t a H))))).
-theorem csuba_arity_rev:
+lemma csuba_arity_rev:
\forall (g: G).(\forall (c1: C).(\forall (t: T).(\forall (a: A).((arity g c1
t a) \to (\forall (c2: C).((csuba g c2 c1) \to ((csubv c2 c1) \to (arity g c2
t a))))))))
include "basic_1/clear/fwd.ma".
-theorem csuba_clear_conf:
+lemma csuba_clear_conf:
\forall (g: G).(\forall (c1: C).(\forall (c2: C).((csuba g c1 c2) \to
(\forall (e1: C).((clear c1 e1) \to (ex2 C (\lambda (e2: C).(csuba g e1 e2))
(\lambda (e2: C).(clear c2 e2))))))))
u) (csuba_abst g c3 c4 H0 t a H2 u H3) (clear_bind Abbr c4 u)) e1
(clear_gen_bind Abst c3 e1 t H4))))))))))))) c1 c2 H)))).
-theorem csuba_clear_trans:
+lemma csuba_clear_trans:
\forall (g: G).(\forall (c1: C).(\forall (c2: C).((csuba g c2 c1) \to
(\forall (e1: C).((clear c1 e1) \to (ex2 C (\lambda (e2: C).(csuba g e2 e1))
(\lambda (e2: C).(clear c2 e2))))))))
include "basic_1/drop/fwd.ma".
-theorem csuba_drop_abbr:
+lemma csuba_drop_abbr:
\forall (i: nat).(\forall (c1: C).(\forall (d1: C).(\forall (u: T).((drop i
O c1 (CHead d1 (Bind Abbr) u)) \to (\forall (g: G).(\forall (c2: C).((csuba g
c1 c2) \to (ex2 C (\lambda (d2: C).(drop i O c2 (CHead d2 (Bind Abbr) u)))
H8)) c2 H6))))) H5)))))) k H2 (drop_gen_drop k c (CHead d1 (Bind Abbr) u) t n
H1)))))))))))) c1)))) i).
-theorem csuba_drop_abst:
+lemma csuba_drop_abst:
\forall (i: nat).(\forall (c1: C).(\forall (d1: C).(\forall (u1: T).((drop i
O c1 (CHead d1 (Bind Abst) u1)) \to (\forall (g: G).(\forall (c2: C).((csuba
g c1 c2) \to (or (ex2 C (\lambda (d2: C).(drop i O c2 (CHead d2 (Bind Abst)
Abbr) x3) H10 x1) H11 H12 H13))))))))) H9)) H8)) c2 H6))))) H5)))))) k H2
(drop_gen_drop k c (CHead d1 (Bind Abst) u1) t n H1)))))))))))) c1)))) i).
-theorem csuba_drop_abst_rev:
+lemma csuba_drop_abst_rev:
\forall (i: nat).(\forall (c1: C).(\forall (d1: C).(\forall (u: T).((drop i
O c1 (CHead d1 (Bind Abst) u)) \to (\forall (g: G).(\forall (c2: C).((csuba g
c2 c1) \to (or (ex2 C (\lambda (d2: C).(drop i O c2 (CHead d2 (Bind Abst)
H8)) c2 H6))))) H5)))))) k H2 (drop_gen_drop k c (CHead d1 (Bind Abst) u) t n
H1)))))))))))) c1)))) i).
-theorem csuba_drop_abbr_rev:
+lemma csuba_drop_abbr_rev:
\forall (i: nat).(\forall (c1: C).(\forall (d1: C).(\forall (u1: T).((drop i
O c1 (CHead d1 (Bind Abbr) u1)) \to (\forall (g: G).(\forall (c2: C).((csuba
g c2 c1) \to (or3 (ex2 C (\lambda (d2: C).(drop i O c2 (CHead d2 (Bind Abbr)
include "basic_1/csuba/defs.ma".
-let rec csuba_ind (g: G) (P: (C \to (C \to Prop))) (f: (\forall (n: nat).(P
-(CSort n) (CSort n)))) (f0: (\forall (c1: C).(\forall (c2: C).((csuba g c1
-c2) \to ((P c1 c2) \to (\forall (k: K).(\forall (u: T).(P (CHead c1 k u)
+implied let rec csuba_ind (g: G) (P: (C \to (C \to Prop))) (f: (\forall (n:
+nat).(P (CSort n) (CSort n)))) (f0: (\forall (c1: C).(\forall (c2: C).((csuba
+g c1 c2) \to ((P c1 c2) \to (\forall (k: K).(\forall (u: T).(P (CHead c1 k u)
(CHead c2 k u))))))))) (f1: (\forall (c1: C).(\forall (c2: C).((csuba g c1
c2) \to ((P c1 c2) \to (\forall (b: B).((not (eq B b Void)) \to (\forall (u1:
T).(\forall (u2: T).(P (CHead c1 (Bind Void) u1) (CHead c2 (Bind b)
f1 f2) c2 c3 c4) b n u1 u2) | (csuba_abst c2 c3 c4 t a a0 u a1) \Rightarrow
(f2 c2 c3 c4 ((csuba_ind g P f f0 f1 f2) c2 c3 c4) t a a0 u a1)].
-theorem csuba_gen_abbr:
+lemma csuba_gen_abbr:
\forall (g: G).(\forall (d1: C).(\forall (c: C).(\forall (u: T).((csuba g
(CHead d1 (Bind Abbr) u) c) \to (ex2 C (\lambda (d2: C).(eq C c (CHead d2
(Bind Abbr) u))) (\lambda (d2: C).(csuba g d1 d2)))))))
(\lambda (d2: C).(eq C (CHead c2 (Bind Abbr) u0) (CHead d2 (Bind Abbr) u)))
(\lambda (d2: C).(csuba g d1 d2))) H6)))))))))))) y c H0))) H))))).
-theorem csuba_gen_void:
+lemma csuba_gen_void:
\forall (g: G).(\forall (d1: C).(\forall (c: C).(\forall (u1: T).((csuba g
(CHead d1 (Bind Void) u1) c) \to (ex2_3 B C T (\lambda (b: B).(\lambda (d2:
C).(\lambda (u2: T).(eq C c (CHead d2 (Bind b) u2))))) (\lambda (_:
(\lambda (_: B).(\lambda (d2: C).(\lambda (_: T).(csuba g d1 d2)))))
H6)))))))))))) y c H0))) H))))).
-theorem csuba_gen_abst:
+lemma csuba_gen_abst:
\forall (g: G).(\forall (d1: C).(\forall (c: C).(\forall (u1: T).((csuba g
(CHead d1 (Bind Abst) u1) c) \to (or (ex2 C (\lambda (d2: C).(eq C c (CHead
d2 (Bind Abst) u1))) (\lambda (d2: C).(csuba g d1 d2))) (ex4_3 C T A (\lambda
A).(arity g d2 u2 a0)))) c2 u a (refl_equal C (CHead c2 (Bind Abbr) u)) H12
H10 H4)))))))) H6)))))))))))) y c H0))) H))))).
-theorem csuba_gen_flat:
+lemma csuba_gen_flat:
\forall (g: G).(\forall (d1: C).(\forall (c: C).(\forall (u1: T).(\forall
(f: F).((csuba g (CHead d1 (Flat f) u1) c) \to (ex2_2 C T (\lambda (d2:
C).(\lambda (u2: T).(eq C c (CHead d2 (Flat f) u2)))) (\lambda (d2:
c2 (Bind Abbr) u) (CHead d2 (Flat f) u2)))) (\lambda (d2: C).(\lambda (_:
T).(csuba g d1 d2)))) H6)))))))))))) y c H0))) H)))))).
-theorem csuba_gen_bind:
+lemma csuba_gen_bind:
\forall (g: G).(\forall (b1: B).(\forall (e1: C).(\forall (c2: C).(\forall
(v1: T).((csuba g (CHead e1 (Bind b1) v1) c2) \to (ex2_3 B C T (\lambda (b2:
B).(\lambda (e2: C).(\lambda (v2: T).(eq C c2 (CHead e2 (Bind b2) v2)))))
Abbr c3 u (refl_equal C (CHead c3 (Bind Abbr) u)) H14))))))))) H7))
H6)))))))))))) y c2 H0))) H)))))).
-theorem csuba_gen_abst_rev:
+lemma csuba_gen_abst_rev:
\forall (g: G).(\forall (d1: C).(\forall (c: C).(\forall (u: T).((csuba g c
(CHead d1 (Bind Abst) u)) \to (or (ex2 C (\lambda (d2: C).(eq C c (CHead d2
(Bind Abst) u))) (\lambda (d2: C).(csuba g d2 d1))) (ex2_2 C T (\lambda (d2:
(u2: T).(eq C (CHead c1 (Bind Abst) t) (CHead d2 (Bind Void) u2)))) (\lambda
(d2: C).(\lambda (_: T).(csuba g d2 d1))))) H6)))))))))))) c y H0))) H))))).
-theorem csuba_gen_void_rev:
+lemma csuba_gen_void_rev:
\forall (g: G).(\forall (d1: C).(\forall (c: C).(\forall (u: T).((csuba g c
(CHead d1 (Bind Void) u)) \to (ex2 C (\lambda (d2: C).(eq C c (CHead d2 (Bind
Void) u))) (\lambda (d2: C).(csuba g d2 d1)))))))
(\lambda (d2: C).(eq C (CHead c1 (Bind Abst) t) (CHead d2 (Bind Void) u)))
(\lambda (d2: C).(csuba g d2 d1))) H6)))))))))))) c y H0))) H))))).
-theorem csuba_gen_abbr_rev:
+lemma csuba_gen_abbr_rev:
\forall (g: G).(\forall (d1: C).(\forall (c: C).(\forall (u1: T).((csuba g c
(CHead d1 (Bind Abbr) u1)) \to (or3 (ex2 C (\lambda (d2: C).(eq C c (CHead d2
(Bind Abbr) u1))) (\lambda (d2: C).(csuba g d2 d1))) (ex4_3 C T A (\lambda
u1 a0)))) c1 t a (refl_equal C (CHead c1 (Bind Abst) t)) H12 H3 H10))))))))
H6)))))))))))) c y H0))) H))))).
-theorem csuba_gen_flat_rev:
+lemma csuba_gen_flat_rev:
\forall (g: G).(\forall (d1: C).(\forall (c: C).(\forall (u1: T).(\forall
(f: F).((csuba g c (CHead d1 (Flat f) u1)) \to (ex2_2 C T (\lambda (d2:
C).(\lambda (u2: T).(eq C c (CHead d2 (Flat f) u2)))) (\lambda (d2:
c1 (Bind Abst) t) (CHead d2 (Flat f) u2)))) (\lambda (d2: C).(\lambda (_:
T).(csuba g d2 d1)))) H6)))))))))))) c y H0))) H)))))).
-theorem csuba_gen_bind_rev:
+lemma csuba_gen_bind_rev:
\forall (g: G).(\forall (b1: B).(\forall (e1: C).(\forall (c2: C).(\forall
(v1: T).((csuba g c2 (CHead e1 (Bind b1) v1)) \to (ex2_3 B C T (\lambda (b2:
B).(\lambda (e2: C).(\lambda (v2: T).(eq C c2 (CHead e2 (Bind b2) v2)))))
include "basic_1/getl/clear.ma".
-theorem csuba_getl_abbr:
+lemma csuba_getl_abbr:
\forall (g: G).(\forall (c1: C).(\forall (d1: C).(\forall (u: T).(\forall
(i: nat).((getl i c1 (CHead d1 (Bind Abbr) u)) \to (\forall (c2: C).((csuba g
c1 c2) \to (ex2 C (\lambda (d2: C).(getl i c2 (CHead d2 (Bind Abbr) u)))
x9 (Bind Abbr) u) n H22) H23)))) H21)))))))) H17)))))) H14))))))) H11))))))))
i) H7))))) k H3 H4))))))) x H1 H2)))) H0))))))).
-theorem csuba_getl_abst:
+lemma csuba_getl_abst:
\forall (g: G).(\forall (c1: C).(\forall (d1: C).(\forall (u1: T).(\forall
(i: nat).((getl i c1 (CHead d1 (Bind Abst) u1)) \to (\forall (c2: C).((csuba
g c1 c2) \to (or (ex2 C (\lambda (d2: C).(getl i c2 (CHead d2 (Bind Abst)
H21)))))))) H17)))))) H14))))))) H11)))))))) i) H7))))) k H3 H4))))))) x H1
H2)))) H0))))))).
-theorem csuba_getl_abst_rev:
+lemma csuba_getl_abst_rev:
\forall (g: G).(\forall (c1: C).(\forall (d1: C).(\forall (u: T).(\forall
(i: nat).((getl i c1 (CHead d1 (Bind Abst) u)) \to (\forall (c2: C).((csuba g
c2 c1) \to (or (ex2 C (\lambda (d2: C).(getl i c2 (CHead d2 (Bind Abst) u)))
Void) x10) n H23) H24)))))) H22)) H21)))))))) H17)))))) H14)))))))
H11)))))))) i) H7))))) k H3 H4))))))) x H1 H2)))) H0))))))).
-theorem csuba_getl_abbr_rev:
+lemma csuba_getl_abbr_rev:
\forall (g: G).(\forall (c1: C).(\forall (d1: C).(\forall (u1: T).(\forall
(i: nat).((getl i c1 (CHead d1 (Bind Abbr) u1)) \to (\forall (c2: C).((csuba
g c2 c1) \to (or3 (ex2 C (\lambda (d2: C).(getl i c2 (CHead d2 (Bind Abbr)
include "basic_1/C/fwd.ma".
-theorem csuba_refl:
+lemma csuba_refl:
\forall (g: G).(\forall (c: C).(csuba g c c))
\def
\lambda (g: G).(\lambda (c: C).(C_ind (\lambda (c0: C).(csuba g c0 c0))
include "basic_1/csubc/csuba.ma".
-theorem csubc_arity_conf:
+lemma csubc_arity_conf:
\forall (g: G).(\forall (c1: C).(\forall (c2: C).((csubc g c1 c2) \to
(\forall (t: T).(\forall (a: A).((arity g c1 t a) \to (arity g c2 t a)))))))
\def
c2)).(\lambda (t: T).(\lambda (a: A).(\lambda (H0: (arity g c1 t
a)).(csuba_arity g c1 t a H0 c2 (csubc_csuba g c1 c2 H)))))))).
-theorem csubc_arity_trans:
+lemma csubc_arity_trans:
\forall (g: G).(\forall (c1: C).(\forall (c2: C).((csubc g c1 c2) \to
((csubv c1 c2) \to (\forall (t: T).(\forall (a: A).((arity g c2 t a) \to
(arity g c1 t a))))))))
include "basic_1/clear/fwd.ma".
-theorem csubc_clear_conf:
+lemma csubc_clear_conf:
\forall (g: G).(\forall (c1: C).(\forall (e1: C).((clear c1 e1) \to (\forall
(c2: C).((csubc g c1 c2) \to (ex2 C (\lambda (e2: C).(clear c2 e2)) (\lambda
(e2: C).(csubc g e1 e2))))))))
include "basic_1/sc3/props.ma".
-theorem csubc_csuba:
+lemma csubc_csuba:
\forall (g: G).(\forall (c1: C).(\forall (c2: C).((csubc g c1 c2) \to (csuba
g c1 c2))))
\def
include "basic_1/sc3/props.ma".
-theorem csubc_drop_conf_O:
+lemma csubc_drop_conf_O:
\forall (g: G).(\forall (c1: C).(\forall (e1: C).(\forall (h: nat).((drop h
O c1 e1) \to (\forall (c2: C).((csubc g c1 c2) \to (ex2 C (\lambda (e2:
C).(drop h O c2 e2)) (\lambda (e2: C).(csubc g e1 e2)))))))))
e2)) (\lambda (e2: C).(csubc g e1 e2)) x (drop_drop (Bind x0) n x1 x H12 x2)
H13)))) H11))))) c2 H5)))))))) H4)) H3)))))))) h))))))) c1)).
-theorem drop_csubc_trans:
+lemma drop_csubc_trans:
\forall (g: G).(\forall (c2: C).(\forall (e2: C).(\forall (d: nat).(\forall
(h: nat).((drop h d c2 e2) \to (\forall (e1: C).((csubc g e2 e1) \to (ex2 C
(\lambda (c1: C).(drop h d c1 e1)) (\lambda (c1: C).(csubc g c2 c1))))))))))
x4)))))) H17))) k H12))) e1 H11)))))))) H10)) H9))) t H4)))))))))
(drop_gen_skip_l c e2 t h n k H1)))))))) d))))))) c2)).
-theorem csubc_drop_conf_rev:
+lemma csubc_drop_conf_rev:
\forall (g: G).(\forall (c2: C).(\forall (e2: C).(\forall (d: nat).(\forall
(h: nat).((drop h d c2 e2) \to (\forall (e1: C).((csubc g e1 e2) \to (ex2 C
(\lambda (c1: C).(drop h d c1 e1)) (\lambda (c1: C).(csubc g c1 c2))))))))))
include "basic_1/csubc/drop.ma".
-theorem drop1_csubc_trans:
+lemma drop1_csubc_trans:
\forall (g: G).(\forall (hds: PList).(\forall (c2: C).(\forall (e2:
C).((drop1 hds c2 e2) \to (\forall (e1: C).((csubc g e2 e1) \to (ex2 C
(\lambda (c1: C).(drop1 hds c1 e1)) (\lambda (c1: C).(csubc g c2 c1)))))))))
(PCons n n0 p) c1 e1)) (\lambda (c1: C).(csubc g c2 c1)) x1 (drop1_cons x1 x0
n n0 H9 e1 p H6) H10)))) H8)))))) H5)))))) H2)))))))))))) hds)).
-theorem csubc_drop1_conf_rev:
+lemma csubc_drop1_conf_rev:
\forall (g: G).(\forall (hds: PList).(\forall (c2: C).(\forall (e2:
C).((drop1 hds c2 e2) \to (\forall (e1: C).((csubc g e1 e2) \to (ex2 C
(\lambda (c1: C).(drop1 hds c1 e1)) (\lambda (c1: C).(csubc g c1 c2)))))))))
include "basic_1/csubc/defs.ma".
-let rec csubc_ind (g: G) (P: (C \to (C \to Prop))) (f: (\forall (n: nat).(P
-(CSort n) (CSort n)))) (f0: (\forall (c1: C).(\forall (c2: C).((csubc g c1
-c2) \to ((P c1 c2) \to (\forall (k: K).(\forall (v: T).(P (CHead c1 k v)
+implied let rec csubc_ind (g: G) (P: (C \to (C \to Prop))) (f: (\forall (n:
+nat).(P (CSort n) (CSort n)))) (f0: (\forall (c1: C).(\forall (c2: C).((csubc
+g c1 c2) \to ((P c1 c2) \to (\forall (k: K).(\forall (v: T).(P (CHead c1 k v)
(CHead c2 k v))))))))) (f1: (\forall (c1: C).(\forall (c2: C).((csubc g c1
c2) \to ((P c1 c2) \to (\forall (b: B).((not (eq B b Void)) \to (\forall (u1:
T).(\forall (u2: T).(P (CHead c1 (Bind Void) u1) (CHead c2 (Bind b)
f1 f2) c2 c3 c4) b n u1 u2) | (csubc_abst c2 c3 c4 v a s0 w s1) \Rightarrow
(f2 c2 c3 c4 ((csubc_ind g P f f0 f1 f2) c2 c3 c4) v a s0 w s1)].
-theorem csubc_gen_sort_l:
+lemma csubc_gen_sort_l:
\forall (g: G).(\forall (x: C).(\forall (n: nat).((csubc g (CSort n) x) \to
(eq C x (CSort n)))))
\def
\Rightarrow True])) I (CSort n) H5) in (False_ind (eq C (CHead c2 (Bind Abbr)
w) (CHead c1 (Bind Abst) v)) H6)))))))))))) y x H0))) H)))).
-theorem csubc_gen_head_l:
+lemma csubc_gen_head_l:
\forall (g: G).(\forall (c1: C).(\forall (x: C).(\forall (v: T).(\forall (k:
K).((csubc g (CHead c1 k v) x) \to (or3 (ex2 C (\lambda (c2: C).(eq C x
(CHead c2 k v))) (\lambda (c2: C).(csubc g c1 c2))) (ex5_3 C T A (\lambda (_:
(Bind Abbr) w)) H14 H12 H4)) k H9))))))))) H7)) H6)))))))))))) y x H0)))
H)))))).
-theorem csubc_gen_sort_r:
+lemma csubc_gen_sort_r:
\forall (g: G).(\forall (x: C).(\forall (n: nat).((csubc g x (CSort n)) \to
(eq C x (CSort n)))))
\def
\Rightarrow True])) I (CSort n) H5) in (False_ind (eq C (CHead c1 (Bind Abst)
v) (CHead c2 (Bind Abbr) w)) H6)))))))))))) x y H0))) H)))).
-theorem csubc_gen_head_r:
+lemma csubc_gen_head_r:
\forall (g: G).(\forall (c2: C).(\forall (x: C).(\forall (w: T).(\forall (k:
K).((csubc g x (CHead c2 k w)) \to (or3 (ex2 C (\lambda (c1: C).(eq C x
(CHead c1 k w))) (\lambda (c1: C).(csubc g c1 c2))) (ex5_3 C T A (\lambda (_:
include "basic_1/csubc/clear.ma".
-theorem csubc_getl_conf:
+lemma csubc_getl_conf:
\forall (g: G).(\forall (c1: C).(\forall (e1: C).(\forall (i: nat).((getl i
c1 e1) \to (\forall (c2: C).((csubc g c1 c2) \to (ex2 C (\lambda (e2:
C).(getl i c2 e2)) (\lambda (e2: C).(csubc g e1 e2)))))))))
include "basic_1/sc3/props.ma".
-theorem csubc_refl:
+lemma csubc_refl:
\forall (g: G).(\forall (c: C).(csubc g c c))
\def
\lambda (g: G).(\lambda (c: C).(C_ind (\lambda (c0: C).(csubc g c0 c0))
include "basic_1/clear/fwd.ma".
-theorem csubst0_clear_O:
+lemma csubst0_clear_O:
\forall (c1: C).(\forall (c2: C).(\forall (v: T).((csubst0 O v c1 c2) \to
(\forall (c: C).((clear c1 c) \to (clear c2 c))))))
\def
H6 O H9) in (clear_flat x1 c0 (H x1 v H10 c0 (clear_gen_flat f c c0 t H8)) f
x0)))))) k H1 H4) c2 H5)))))))) H3)) H2))))))))))) c1).
-theorem csubst0_clear_O_back:
+lemma csubst0_clear_O_back:
\forall (c1: C).(\forall (c2: C).(\forall (v: T).((csubst0 O v c1 c2) \to
(\forall (c: C).((clear c2 c) \to (clear c1 c))))))
\def
n v t x0)) H6 O H9) in (clear_flat c c0 (H x1 v H11 c0 (clear_gen_flat f x1
c0 x0 H10)) f t)))))) k H4 H8))))))))) H3)) H2))))))))))) c1).
-theorem csubst0_clear_S:
+lemma csubst0_clear_S:
\forall (c1: C).(\forall (c2: C).(\forall (v: T).(\forall (i: nat).((csubst0
(S i) v c1 c2) \to (\forall (c: C).((clear c1 c) \to (or4 (clear c2 c) (ex3_4
B C T T (\lambda (b: B).(\lambda (e: C).(\lambda (u1: T).(\lambda (_: T).(eq
(CHead x5 (Bind x3) x7) H16 f x0) H17 H18))))))))))) H14)) H13)))))))) k H1
H4) c2 H5)))))))) H3)) H2)))))))))))) c1).
-theorem csubst0_clear_trans:
+lemma csubst0_clear_trans:
\forall (c1: C).(\forall (c2: C).(\forall (v: T).(\forall (i: nat).((csubst0
i v c1 c2) \to (\forall (e2: C).((clear c2 e2) \to (or (clear c1 e2) (ex2 C
(\lambda (e1: C).(csubst0 i v e1 e2)) (\lambda (e1: C).(clear c1 e1))))))))))
include "basic_1/drop/fwd.ma".
-theorem csubst0_drop_gt:
+lemma csubst0_drop_gt:
\forall (n: nat).(\forall (i: nat).((lt i n) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to (\forall (e: C).((drop n O
c1 e) \to (drop n O c2 e)))))))))
(drop_gen_drop k c e t n0 H3) H10 H11))) c2 H7)))))))) H5)) H4)))))))))))
c1)))))) n).
-theorem csubst0_drop_gt_back:
+lemma csubst0_drop_gt_back:
\forall (n: nat).(\forall (i: nat).((lt i n) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to (\forall (e: C).((drop n O
c2 e) \to (drop n O c1 e)))))))))
t)))) H16)) (lt_gen_xS x2 n0 H14)))))) k H11 H12 (drop_gen_drop k x1 e x0 n0
H10)))))))))))) H5)) H4))))))))))) c1)))))) n).
-theorem csubst0_drop_lt:
+lemma csubst0_drop_lt:
\forall (n: nat).(\forall (i: nat).((lt n i) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to (\forall (e: C).((drop n O
c1 e) \to (or4 (drop n O c2 e) (ex3_4 K C T T (\lambda (k: K).(\lambda (e0:
x7) H17 x0) H18 H19)) e H16)))))))))) H15)) H14)))))) k (drop_gen_drop k c e
t n0 H2) H9 H10) i H5))) c2 H6)))))))) H4)) H3))))))))))) c1)))))) n).
-theorem csubst0_drop_eq:
+lemma csubst0_drop_eq:
\forall (n: nat).(\forall (c1: C).(\forall (c2: C).(\forall (v: T).((csubst0
n v c1 c2) \to (\forall (e: C).((drop n O c1 e) \to (or4 (drop n O c2 e)
(ex3_4 F C T T (\lambda (f: F).(\lambda (e0: C).(\lambda (u: T).(\lambda (_:
(Flat x3) x7) H16 x0) H17 H18)) e H15)))))))))) H14)) H13)))))))) k
(drop_gen_drop k c e t n0 H1) H4) c2 H5)))))))) H3)) H2))))))))))) c1)))) n).
-theorem csubst0_drop_eq_back:
+lemma csubst0_drop_eq_back:
\forall (n: nat).(\forall (c1: C).(\forall (c2: C).(\forall (v: T).((csubst0
n v c1 c2) \to (\forall (e: C).((drop n O c2 e) \to (or4 (drop n O c1 e)
(ex3_4 F C T T (\lambda (f: F).(\lambda (e0: C).(\lambda (_: T).(\lambda (u2:
(Flat x3) x6) H17 t) H18 H19)) e H16))))))))))) H15)) H14)))))))) k H4
(drop_gen_drop k x1 e x0 n0 H8)))))))))) H3)) H2))))))))))) c1)))) n).
-theorem csubst0_drop_lt_back:
+lemma csubst0_drop_lt_back:
\forall (n: nat).(\forall (i: nat).((lt n i) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to (\forall (e2: C).((drop n O
c2 e2) \to (or (drop n O c1 e2) (ex2 C (\lambda (e1: C).(csubst0 (minus i n)
include "basic_1/C/fwd.ma".
-let rec csubst0_ind (P: (nat \to (T \to (C \to (C \to Prop))))) (f: (\forall
-(k: K).(\forall (i: nat).(\forall (v: T).(\forall (u1: T).(\forall (u2:
-T).((subst0 i v u1 u2) \to (\forall (c: C).(P (s k i) v (CHead c k u1) (CHead
-c k u2)))))))))) (f0: (\forall (k: K).(\forall (i: nat).(\forall (c1:
+implied let rec csubst0_ind (P: (nat \to (T \to (C \to (C \to Prop))))) (f:
+(\forall (k: K).(\forall (i: nat).(\forall (v: T).(\forall (u1: T).(\forall
+(u2: T).((subst0 i v u1 u2) \to (\forall (c: C).(P (s k i) v (CHead c k u1)
+(CHead c k u2)))))))))) (f0: (\forall (k: K).(\forall (i: nat).(\forall (c1:
C).(\forall (c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to ((P i v c1 c2)
\to (\forall (u: T).(P (s k i) v (CHead c1 k u) (CHead c2 k u))))))))))) (f1:
(\forall (k: K).(\forall (i: nat).(\forall (v: T).(\forall (u1: T).(\forall
c2 c3 c4) \Rightarrow (f1 k i v u1 u2 s0 c2 c3 c4 ((csubst0_ind P f f0 f1) i
v c2 c3 c4))].
-theorem csubst0_gen_sort:
+lemma csubst0_gen_sort:
\forall (x: C).(\forall (v: T).(\forall (i: nat).(\forall (n: nat).((csubst0
i v (CSort n) x) \to (\forall (P: Prop).P)))))
\def
C).(match ee with [(CSort _) \Rightarrow False | (CHead _ _ _) \Rightarrow
True])) I (CSort n) H4) in (False_ind P H5))))))))))))) i v y x H0))) H)))))).
-theorem csubst0_gen_head:
+lemma csubst0_gen_head:
\forall (k: K).(\forall (c1: C).(\forall (x: C).(\forall (u1: T).(\forall
(v: T).(\forall (i: nat).((csubst0 i v (CHead c1 k u1) x) \to (or3 (ex3_2 T
nat (\lambda (_: T).(\lambda (j: nat).(eq nat i (s k j)))) (\lambda (u2:
i0)) (refl_equal C (CHead c2 k u2)) H12 H11)) k0 H8))))))) H6))
H5))))))))))))) i v y x H0))) H))))))).
-theorem csubst0_gen_S_bind_2:
+lemma csubst0_gen_S_bind_2:
\forall (b: B).(\forall (x: C).(\forall (c2: C).(\forall (v: T).(\forall
(v2: T).(\forall (i: nat).((csubst0 (S i) v x (CHead c2 (Bind b) v2)) \to
(or3 (ex2 T (\lambda (v1: T).(subst0 i v v1 v2)) (\lambda (v1: T).(eq C x
include "basic_1/getl/fwd.ma".
-theorem csubst0_getl_ge:
+lemma csubst0_getl_ge:
\forall (i: nat).(\forall (n: nat).((le i n) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to (\forall (e: C).((getl n c1
e) \to (getl n c2 e)))))))))
x3 H14)) x0 x4)))))))))))) H9)) H8)) n H5)))) (\lambda (H5: (lt n
i)).(le_lt_false i n H H5 (getl n c2 e))))))) H2)))))))))).
-theorem csubst0_getl_lt:
+lemma csubst0_getl_lt:
\forall (i: nat).(\forall (n: nat).((lt n i) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to (\forall (e: C).((getl n c1
e) \to (or4 (getl n c2 e) (ex3_4 B C T T (\lambda (b: B).(\lambda (e0:
(clear_flat x2 (CHead x7 (Bind x5) x9) H20 f x4)) H21 H22)) e H19))))))))))
H18)) H17)))))))) x0 H8 H9 H10 H11))))))))))) H6)) H5))))) H2)))))))))).
-theorem csubst0_getl_ge_back:
+lemma csubst0_getl_ge_back:
\forall (i: nat).(\forall (n: nat).((le i n) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to (\forall (e: C).((getl n c2
e) \to (getl n c1 e)))))))))
x2 e x4 H14)) x0 x3)))))))))))) H9)) H8)) n H5)))) (\lambda (H5: (lt n
i)).(le_lt_false i n H H5 (getl n c1 e))))))) H2)))))))))).
-theorem csubst0_getl_lt_back:
+lemma csubst0_getl_lt_back:
\forall (n: nat).(\forall (i: nat).((lt n i) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst0 i v c1 c2) \to (\forall (e2: C).((getl n c2
e2) \to (or (getl n c1 e2) (ex2 C (\lambda (e1: C).(csubst0 (minus i n) v e1
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))))))))
u2))) (csubst0_snd (Bind b) i v u1 u2 H c) (S i) (refl_equal nat (S
i))))))))).
-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))))))))
include "basic_1/s/fwd.ma".
-theorem csubst1_ind:
+implied lemma csubst1_ind:
\forall (i: nat).(\forall (v: T).(\forall (c1: C).(\forall (P: ((C \to
Prop))).((P c1) \to (((\forall (c2: C).((csubst0 i v c1 c2) \to (P c2)))) \to
(\forall (c: C).((csubst1 i v c1 c) \to (P c))))))))
c0 with [csubst1_refl \Rightarrow f | (csubst1_sing x x0) \Rightarrow (f0 x
x0)])))))))).
-theorem csubst1_gen_head:
+lemma csubst1_gen_head:
\forall (k: K).(\forall (c1: C).(\forall (x: C).(\forall (u1: T).(\forall
(v: T).(\forall (i: nat).((csubst1 (s k i) v (CHead c1 k u1) x) \to (ex3_2 T
C (\lambda (u2: T).(\lambda (c2: C).(eq C x (CHead c2 k u2)))) (\lambda (u2:
include "basic_1/drop/props.ma".
-theorem csubst1_getl_ge:
+lemma csubst1_getl_ge:
\forall (i: nat).(\forall (n: nat).((le i n) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst1 i v c1 c2) \to (\forall (e: C).((getl n c1
e) \to (getl n c2 e)))))))))
(c3: C).(\lambda (H1: (csubst0 i v c1 c3)).(\lambda (e: C).(\lambda (H2:
(getl n c1 e)).(csubst0_getl_ge i n H c1 c3 v H1 e H2))))) c2 H0))))))).
-theorem csubst1_getl_lt:
+lemma csubst1_getl_lt:
\forall (i: nat).(\forall (n: nat).((lt n i) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst1 i v c1 c2) \to (\forall (e1: C).((getl n c1
e1) \to (ex2 C (\lambda (e2: C).(csubst1 (minus i n) v e1 e2)) (\lambda (e2:
x0) x4) (csubst0_both_bind x0 (minus i (S n)) v x3 x4 H7 x1 x2 H8)) H6) e1
H5)))))))))) H4)) H3)) (minus i n) (minus_x_Sy i n H)))))) c2 H0))))))).
-theorem csubst1_getl_ge_back:
+lemma csubst1_getl_ge_back:
\forall (i: nat).(\forall (n: nat).((le i n) \to (\forall (c1: C).(\forall
(c2: C).(\forall (v: T).((csubst1 i v c1 c2) \to (\forall (e: C).((getl n c2
e) \to (getl n c1 e)))))))))
(c3: C).(\lambda (H1: (csubst0 i v c1 c3)).(\lambda (e: C).(\lambda (H2:
(getl n c3 e)).(csubst0_getl_ge_back i n H c1 c3 v H1 e H2))))) c2 H0))))))).
-theorem getl_csubst1:
+lemma getl_csubst1:
\forall (d: nat).(\forall (c: C).(\forall (e: C).(\forall (u: T).((getl d c
(CHead e (Bind Abbr) u)) \to (ex2_2 C C (\lambda (a0: C).(\lambda (_:
C).(csubst1 d u c a0))) (\lambda (a0: C).(\lambda (a: C).(drop (S O) d a0
include "basic_1/clear/fwd.ma".
-theorem csubt_clear_conf:
+lemma csubt_clear_conf:
\forall (g: G).(\forall (c1: C).(\forall (c2: C).((csubt g c1 c2) \to
(\forall (e1: C).((clear c1 e1) \to (ex2 C (\lambda (e2: C).(csubt g e1 e2))
(\lambda (e2: C).(clear c2 e2))))))))
include "basic_1/ty3/arity.ma".
-theorem csubt_csuba:
+lemma csubt_csuba:
\forall (g: G).(\forall (c1: C).(\forall (c2: C).((csubt g c1 c2) \to (csuba
g c1 c2))))
\def
include "basic_1/drop/fwd.ma".
-theorem csubt_drop_flat:
+lemma csubt_drop_flat:
\forall (g: G).(\forall (f: F).(\forall (n: nat).(\forall (c1: C).(\forall
(c2: C).((csubt g c1 c2) \to (\forall (d1: C).(\forall (u: T).((drop n O c1
(CHead d1 (Flat f) u)) \to (ex2 C (\lambda (d2: C).(csubt g d1 d2)) (\lambda
c3 H1 d1 u0 (drop_gen_drop (Bind Abst) c0 (CHead d1 (Flat f) u0) t n0
H5)))))))))))))) c1 c2 H0)))))) n))).
-theorem csubt_drop_abbr:
+lemma csubt_drop_abbr:
\forall (g: G).(\forall (n: nat).(\forall (c1: C).(\forall (c2: C).((csubt g
c1 c2) \to (\forall (d1: C).(\forall (u: T).((drop n O c1 (CHead d1 (Bind
Abbr) u)) \to (ex2 C (\lambda (d2: C).(csubt g d1 d2)) (\lambda (d2: C).(drop
(CHead x (Bind Abbr) u0) H7 u))))) (H c0 c3 H1 d1 u0 (drop_gen_drop (Bind
Abst) c0 (CHead d1 (Bind Abbr) u0) t n0 H5)))))))))))))) c1 c2 H0)))))) n)).
-theorem csubt_drop_abst:
+lemma csubt_drop_abst:
\forall (g: G).(\forall (n: nat).(\forall (c1: C).(\forall (c2: C).((csubt g
c1 c2) \to (\forall (d1: C).(\forall (t: T).((drop n O c1 (CHead d1 (Bind
Abst) t)) \to (or (ex2 C (\lambda (d2: C).(csubt g d1 d2)) (\lambda (d2:
include "basic_1/csubt/defs.ma".
-let rec csubt_ind (g: G) (P: (C \to (C \to Prop))) (f: (\forall (n: nat).(P
-(CSort n) (CSort n)))) (f0: (\forall (c1: C).(\forall (c2: C).((csubt g c1
-c2) \to ((P c1 c2) \to (\forall (k: K).(\forall (u: T).(P (CHead c1 k u)
+implied let rec csubt_ind (g: G) (P: (C \to (C \to Prop))) (f: (\forall (n:
+nat).(P (CSort n) (CSort n)))) (f0: (\forall (c1: C).(\forall (c2: C).((csubt
+g c1 c2) \to ((P c1 c2) \to (\forall (k: K).(\forall (u: T).(P (CHead c1 k u)
(CHead c2 k u))))))))) (f1: (\forall (c1: C).(\forall (c2: C).((csubt g c1
c2) \to ((P c1 c2) \to (\forall (b: B).((not (eq B b Void)) \to (\forall (u1:
T).(\forall (u2: T).(P (CHead c1 (Bind Void) u1) (CHead c2 (Bind b)
(csubt_abst c2 c3 c4 u t t0 t1) \Rightarrow (f2 c2 c3 c4 ((csubt_ind g P f f0
f1 f2) c2 c3 c4) u t t0 t1)].
-theorem csubt_gen_abbr:
+lemma csubt_gen_abbr:
\forall (g: G).(\forall (e1: C).(\forall (c2: C).(\forall (v: T).((csubt g
(CHead e1 (Bind Abbr) v) c2) \to (ex2 C (\lambda (e2: C).(eq C c2 (CHead e2
(Bind Abbr) v))) (\lambda (e2: C).(csubt g e1 e2)))))))
(CHead c3 (Bind Abbr) u) (CHead e2 (Bind Abbr) v))) (\lambda (e2: C).(csubt g
e1 e2))) H6))))))))))) y c2 H0))) H))))).
-theorem csubt_gen_abst:
+lemma csubt_gen_abst:
\forall (g: G).(\forall (e1: C).(\forall (c2: C).(\forall (v1: T).((csubt g
(CHead e1 (Bind Abst) v1) c2) \to (or (ex2 C (\lambda (e2: C).(eq C c2 (CHead
e2 (Bind Abst) v1))) (\lambda (e2: C).(csubt g e1 e2))) (ex4_2 C T (\lambda
g e2 v2 v1))) c3 u (refl_equal C (CHead c3 (Bind Abbr) u)) H13 H11
H9))))))))) H6))))))))))) y c2 H0))) H))))).
-theorem csubt_gen_flat:
+lemma csubt_gen_flat:
\forall (g: G).(\forall (e1: C).(\forall (c2: C).(\forall (v: T).(\forall
(f: F).((csubt g (CHead e1 (Flat f) v) c2) \to (ex2 C (\lambda (e2: C).(eq C
c2 (CHead e2 (Flat f) v))) (\lambda (e2: C).(csubt g e1 e2))))))))
(CHead e2 (Flat f) v))) (\lambda (e2: C).(csubt g e1 e2))) H6))))))))))) y c2
H0))) H)))))).
-theorem csubt_gen_bind:
+lemma csubt_gen_bind:
\forall (g: G).(\forall (b1: B).(\forall (e1: C).(\forall (c2: C).(\forall
(v1: T).((csubt g (CHead e1 (Bind b1) v1) c2) \to (ex2_3 B C T (\lambda (b2:
B).(\lambda (e2: C).(\lambda (v2: T).(eq C c2 (CHead e2 (Bind b2) v2)))))
include "basic_1/getl/clear.ma".
-theorem csubt_getl_abbr:
+lemma csubt_getl_abbr:
\forall (g: G).(\forall (c1: C).(\forall (d1: C).(\forall (u: T).(\forall
(n: nat).((getl n c1 (CHead d1 (Bind Abbr) u)) \to (\forall (c2: C).((csubt g
c1 c2) \to (ex2 C (\lambda (d2: C).(csubt g d1 d2)) (\lambda (d2: C).(getl n
H21)))))))) H17))))) H14))))))) H11)))))))) n) H7))))) k H3 H4))))))) x H1
H2)))) H0))))))).
-theorem csubt_getl_abst:
+lemma csubt_getl_abst:
\forall (g: G).(\forall (c1: C).(\forall (d1: C).(\forall (t: T).(\forall
(n: nat).((getl n c1 (CHead d1 (Bind Abst) t)) \to (\forall (c2: C).((csubt g
c1 c2) \to (or (ex2 C (\lambda (d2: C).(csubt g d1 d2)) (\lambda (d2:
include "basic_1/pc3/left.ma".
-theorem csubt_pr2:
+lemma csubt_pr2:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t2: T).((pr2 c1
t1 t2) \to (\forall (c2: C).((csubt g c1 c2) \to (pr2 c2 t1 t2)))))))
\def
(csubt g d x)).(\lambda (H6: (getl i c2 (CHead x (Bind Abbr) u))).(pr2_delta
c2 x u i H6 t3 t4 H1 t H2)))) H4)))))))))))))) c1 t1 t2 H))))).
-theorem csubt_pc3:
+lemma csubt_pc3:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t2: T).((pc3 c1
t1 t2) \to (\forall (c2: C).((csubt g c1 c2) \to (pc3 c2 t1 t2)))))))
\def
include "basic_1/C/fwd.ma".
-theorem csubt_refl:
+lemma csubt_refl:
\forall (g: G).(\forall (c: C).(csubt g c c))
\def
\lambda (g: G).(\lambda (c: C).(C_ind (\lambda (c0: C).(csubt g c0 c0))
include "basic_1/ty3/fwd.ma".
-theorem csubt_ty3:
+lemma csubt_ty3:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t2: T).((ty3 g c1
t1 t2) \to (\forall (c2: C).((csubt g c1 c2) \to (ty3 g c2 t1 t2)))))))
\def
c2) \to (ty3 g c2 t3 t4))))).(\lambda (c2: C).(\lambda (H4: (csubt g c
c2)).(ty3_cast g c2 t0 t3 (H1 c2 H4) t4 (H3 c2 H4)))))))))))) c1 t1 t2 H))))).
-theorem csubt_ty3_ld:
+lemma csubt_ty3_ld:
\forall (g: G).(\forall (c: C).(\forall (u: T).(\forall (v: T).((ty3 g c u
v) \to (\forall (t1: T).(\forall (t2: T).((ty3 g (CHead c (Bind Abst) v) t1
t2) \to (ty3 g (CHead c (Bind Abbr) u) t1 t2))))))))
include "basic_1/clear/fwd.ma".
-theorem csubv_clear_conf:
+lemma csubv_clear_conf:
\forall (c1: C).(\forall (c2: C).((csubv c1 c2) \to (\forall (b1:
B).(\forall (d1: C).(\forall (v1: T).((clear c1 (CHead d1 (Bind b1) v1)) \to
(ex2_3 B C T (\lambda (_: B).(\lambda (d2: C).(\lambda (_: T).(csubv d1
v2) (CHead d2 (Bind b2) v3))))) x0 x1 x2 H4 (clear_flat c4 (CHead x1 (Bind
x0) x2) H5 f2 v2))))))) H3))))))))))))))) c1 c2 H))).
-theorem csubv_clear_conf_void:
+lemma csubv_clear_conf_void:
\forall (c1: C).(\forall (c2: C).((csubv c1 c2) \to (\forall (d1:
C).(\forall (v1: T).((clear c1 (CHead d1 (Bind Void) v1)) \to (ex2_2 C T
(\lambda (d2: C).(\lambda (_: T).(csubv d1 d2))) (\lambda (d2: C).(\lambda
include "basic_1/drop/fwd.ma".
-theorem csubv_drop_conf:
+lemma csubv_drop_conf:
\forall (c1: C).(\forall (c2: C).((csubv c1 c2) \to (\forall (e1:
C).(\forall (h: nat).((drop h O c1 e1) \to (ex2 C (\lambda (e2: C).(csubv e1
e2)) (\lambda (e2: C).(drop h O c2 e2))))))))
include "basic_1/csubv/defs.ma".
-let rec csubv_ind (P: (C \to (C \to Prop))) (f: (\forall (n: nat).(P (CSort
-n) (CSort n)))) (f0: (\forall (c1: C).(\forall (c2: C).((csubv c1 c2) \to ((P
-c1 c2) \to (\forall (v1: T).(\forall (v2: T).(P (CHead c1 (Bind Void) v1)
-(CHead c2 (Bind Void) v2))))))))) (f1: (\forall (c1: C).(\forall (c2:
+implied let rec csubv_ind (P: (C \to (C \to Prop))) (f: (\forall (n: nat).(P
+(CSort n) (CSort n)))) (f0: (\forall (c1: C).(\forall (c2: C).((csubv c1 c2)
+\to ((P c1 c2) \to (\forall (v1: T).(\forall (v2: T).(P (CHead c1 (Bind Void)
+v1) (CHead c2 (Bind Void) v2))))))))) (f1: (\forall (c1: C).(\forall (c2:
C).((csubv c1 c2) \to ((P c1 c2) \to (\forall (b1: B).((not (eq B b1 Void))
\to (\forall (b2: B).(\forall (v1: T).(\forall (v2: T).(P (CHead c1 (Bind b1)
v1) (CHead c2 (Bind b2) v2)))))))))))) (f2: (\forall (c1: C).(\forall (c2:
include "basic_1/getl/fwd.ma".
-theorem csubv_getl_conf:
+lemma csubv_getl_conf:
\forall (c1: C).(\forall (c2: C).((csubv c1 c2) \to (\forall (b1:
B).(\forall (d1: C).(\forall (v1: T).(\forall (i: nat).((getl i c1 (CHead d1
(Bind b1) v1)) \to (ex2_3 B C T (\lambda (_: B).(\lambda (d2: C).(\lambda (_:
b2) v2))))) x1 x2 x3 H8 (getl_intro i c2 (CHead x2 (Bind x1) x3) x0 H6
H9))))))) H7)))))) H4)))))) H1))))))))).
-theorem csubv_getl_conf_void:
+lemma csubv_getl_conf_void:
\forall (c1: C).(\forall (c2: C).((csubv c1 c2) \to (\forall (d1:
C).(\forall (v1: T).(\forall (i: nat).((getl i c1 (CHead d1 (Bind Void) v1))
\to (ex2_2 C T (\lambda (d2: C).(\lambda (_: T).(csubv d1 d2))) (\lambda (d2:
include "basic_1/T/props.ma".
-theorem csubv_bind_same:
+lemma csubv_bind_same:
\forall (c1: C).(\forall (c2: C).((csubv c1 c2) \to (\forall (b: B).(\forall
(v1: T).(\forall (v2: T).(csubv (CHead c1 (Bind b) v1) (CHead c2 (Bind b)
v2)))))))
T).(\lambda (v2: T).(csubv_bind c1 c2 H Abst not_abst_void Abst v1 v2)))
(\lambda (v1: T).(\lambda (v2: T).(csubv_void c1 c2 H v1 v2))) b)))).
-theorem csubv_refl:
+lemma csubv_refl:
\forall (c: C).(csubv c c)
\def
\lambda (c: C).(C_ind (\lambda (c0: C).(csubv c0 c0)) (\lambda (n:
include "basic_1/C/fwd.ma".
-let rec drop_ind (P: (nat \to (nat \to (C \to (C \to Prop))))) (f: (\forall
-(c: C).(P O O c c))) (f0: (\forall (k: K).(\forall (h: nat).(\forall (c:
-C).(\forall (e: C).((drop (r k h) O c e) \to ((P (r k h) O c e) \to (\forall
-(u: T).(P (S h) O (CHead c k u) e))))))))) (f1: (\forall (k: K).(\forall (h:
-nat).(\forall (d: nat).(\forall (c: C).(\forall (e: C).((drop h (r k d) c e)
-\to ((P h (r k d) c e) \to (\forall (u: T).(P h (S d) (CHead c k (lift h (r k
-d) u)) (CHead e k u))))))))))) (n: nat) (n0: nat) (c: C) (c0: C) (d: drop n
-n0 c c0) on d: P n n0 c c0 \def match d with [(drop_refl c1) \Rightarrow (f
-c1) | (drop_drop k h c1 e d0 u) \Rightarrow (f0 k h c1 e d0 ((drop_ind P f f0
-f1) (r k h) O c1 e d0) u) | (drop_skip k h d0 c1 e d1 u) \Rightarrow (f1 k h
-d0 c1 e d1 ((drop_ind P f f0 f1) h (r k d0) c1 e d1) u)].
+implied let rec drop_ind (P: (nat \to (nat \to (C \to (C \to Prop))))) (f:
+(\forall (c: C).(P O O c c))) (f0: (\forall (k: K).(\forall (h: nat).(\forall
+(c: C).(\forall (e: C).((drop (r k h) O c e) \to ((P (r k h) O c e) \to
+(\forall (u: T).(P (S h) O (CHead c k u) e))))))))) (f1: (\forall (k:
+K).(\forall (h: nat).(\forall (d: nat).(\forall (c: C).(\forall (e: C).((drop
+h (r k d) c e) \to ((P h (r k d) c e) \to (\forall (u: T).(P h (S d) (CHead c
+k (lift h (r k d) u)) (CHead e k u))))))))))) (n: nat) (n0: nat) (c: C) (c0:
+C) (d: drop n n0 c c0) on d: P n n0 c c0 \def match d with [(drop_refl c1)
+\Rightarrow (f c1) | (drop_drop k h c1 e d0 u) \Rightarrow (f0 k h c1 e d0
+((drop_ind P f f0 f1) (r k h) O c1 e d0) u) | (drop_skip k h d0 c1 e d1 u)
+\Rightarrow (f1 k h d0 c1 e d1 ((drop_ind P f f0 f1) h (r k d0) c1 e d1) u)].
-theorem drop_gen_sort:
+lemma drop_gen_sort:
\forall (n: nat).(\forall (h: nat).(\forall (d: nat).(\forall (x: C).((drop
h d (CSort n) x) \to (and3 (eq C x (CSort n)) (eq nat h O) (eq nat d O))))))
\def
k d0) u))) (eq nat h0 O) (eq nat (S d0) O)) H4))))))))))) h d y x H0)))
H))))).
-theorem drop_gen_refl:
+lemma drop_gen_refl:
\forall (x: C).(\forall (e: C).((drop O O x e) \to (eq C x e)))
\def
\lambda (x: C).(\lambda (e: C).(\lambda (H: (drop O O x e)).(insert_eq nat O
in (False_ind (eq C (CHead c k (lift (S d) (r k d) u)) (CHead e0 k u)) H9)) h
H6)))))))))))))) y y0 x e H1))) H0))) H))).
-theorem drop_gen_drop:
+lemma drop_gen_drop:
\forall (k: K).(\forall (c: C).(\forall (x: C).(\forall (u: T).(\forall (h:
nat).((drop (S h) O (CHead c k u) x) \to (drop (r k h) O c x))))))
\def
k h) (S d) c (CHead e k u0)) H22))) k0 H14))))))))) H12)) H11))))))))))))))))
y1 y0 y x H2))) H1))) H0))) H)))))).
-theorem drop_gen_skip_r:
+lemma drop_gen_skip_r:
\forall (c: C).(\forall (x: C).(\forall (u: T).(\forall (h: nat).(\forall
(d: nat).(\forall (k: K).((drop h (S d) x (CHead c k u)) \to (ex2 C (\lambda
(e: C).(eq C x (CHead e k (lift h (r k d) u)))) (\lambda (e: C).(drop h (r k
h0 (r k d) u))) H17) d0 H15)))) k0 H9))))) u0 H8)))) H7)) H6)))))))))))) h y0
x y H1))) H0))) H))))))).
-theorem drop_gen_skip_l:
+lemma drop_gen_skip_l:
\forall (c: C).(\forall (x: C).(\forall (u: T).(\forall (h: nat).(\forall
(d: nat).(\forall (k: K).((drop h (S d) (CHead c k u) x) \to (ex3_2 C T
(\lambda (e: C).(\lambda (v: T).(eq C x (CHead e k v)))) (\lambda (_:
H17)))) u H13)) k0 H9))))))))) H7)) H6)))))))))))) h y0 y x H1))) H0)))
H))))))).
-theorem drop_S:
+lemma drop_S:
\forall (b: B).(\forall (c: C).(\forall (e: C).(\forall (u: T).(\forall (h:
nat).((drop h O c (CHead e (Bind b) u)) \to (drop (S h) O c e))))))
\def
include "basic_1/drop/fwd.ma".
-theorem drop_skip_bind:
+lemma drop_skip_bind:
\forall (h: nat).(\forall (d: nat).(\forall (c: C).(\forall (e: C).((drop h
d c e) \to (\forall (b: B).(\forall (u: T).(drop h (S d) (CHead c (Bind b)
(lift h d u)) (CHead e (Bind b) u))))))))
d) (\lambda (n: nat).(drop h (S d) (CHead c (Bind b) (lift h n u)) (CHead e
(Bind b) u))) (drop_skip (Bind b) h d c e H u) d (refl_equal nat d)))))))).
-theorem drop_skip_flat:
+lemma drop_skip_flat:
\forall (h: nat).(\forall (d: nat).(\forall (c: C).(\forall (e: C).((drop h
(S d) c e) \to (\forall (f: F).(\forall (u: T).(drop h (S d) (CHead c (Flat
f) (lift h (S d) u)) (CHead e (Flat f) u))))))))
e (Flat f) u))) (drop_skip (Flat f) h d c e H u) (S d) (refl_equal nat (S
d))))))))).
-theorem drop_ctail:
+lemma drop_ctail:
\forall (c1: C).(\forall (c2: C).(\forall (d: nat).(\forall (h: nat).((drop
h d c1 c2) \to (\forall (k: K).(\forall (u: T).(drop h d (CTail k u c1)
(CTail k u c2))))))))
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
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
\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))))))))
include "basic_1/getl/drop.ma".
-theorem drop1_getl_trans:
+lemma drop1_getl_trans:
\forall (hds: PList).(\forall (c1: C).(\forall (c2: C).((drop1 hds c2 c1)
\to (\forall (b: B).(\forall (e1: C).(\forall (v: T).(\forall (i: nat).((getl
i c1 (CHead e1 (Bind b) v)) \to (ex2 C (\lambda (e2: C).(drop1 (ptrans hds i)
include "basic_1/getl/defs.ma".
-theorem drop1_skip_bind:
+lemma drop1_skip_bind:
\forall (b: B).(\forall (e: C).(\forall (hds: PList).(\forall (c:
C).(\forall (u: T).((drop1 hds c e) \to (drop1 (Ss hds) (CHead c (Bind b)
(lift1 hds u)) (CHead e (Bind b) u)))))))
(lift1 p u)) n (S n0) (drop_skip_bind n n0 c x H2 b (lift1 p u)) (CHead e
(Bind b) u) (Ss p) (H x u H3))))) H1)))))))))) hds))).
-theorem drop1_cons_tail:
+lemma drop1_cons_tail:
\forall (c2: C).(\forall (c3: C).(\forall (h: nat).(\forall (d: nat).((drop
h d c2 c3) \to (\forall (hds: PList).(\forall (c1: C).((drop1 hds c1 c2) \to
(drop1 (PConsTail hds h d) c1 c3))))))))
include "basic_1/ex0/defs.ma".
-let rec leqz_ind (P: (A \to (A \to Prop))) (f: (\forall (h1: nat).(\forall
-(h2: nat).(\forall (n1: nat).(\forall (n2: nat).((eq nat (plus h1 n2) (plus
-h2 n1)) \to (P (ASort h1 n1) (ASort h2 n2)))))))) (f0: (\forall (a1:
-A).(\forall (a2: A).((leqz a1 a2) \to ((P a1 a2) \to (\forall (a3:
+implied let rec leqz_ind (P: (A \to (A \to Prop))) (f: (\forall (h1:
+nat).(\forall (h2: nat).(\forall (n1: nat).(\forall (n2: nat).((eq nat (plus
+h1 n2) (plus h2 n1)) \to (P (ASort h1 n1) (ASort h2 n2)))))))) (f0: (\forall
+(a1: A).(\forall (a2: A).((leqz a1 a2) \to ((P a1 a2) \to (\forall (a3:
A).(\forall (a4: A).((leqz a3 a4) \to ((P a3 a4) \to (P (AHead a1 a3) (AHead
a2 a4))))))))))) (a: A) (a0: A) (l: leqz a a0) on l: P a a0 \def match l with
[(leqz_sort h1 h2 n1 n2 e) \Rightarrow (f h1 h2 n1 n2 e) | (leqz_head a1 a2
include "basic_1/aplus/props.ma".
-theorem aplus_gz_le:
+lemma aplus_gz_le:
\forall (k: nat).(\forall (h: nat).(\forall (n: nat).((le h k) \to (eq A
(aplus gz (ASort h n) k) (ASort O (plus (minus k h) n))))))
\def
k0)) (asucc gz (aplus gz (ASort (S n) n0) k0)) (aplus_asucc gz k0 (ASort (S
n) n0))) (ASort O (plus (minus k0 n) n0)) (IH n n0 H_y))))))) h)))) k).
-theorem aplus_gz_ge:
+lemma aplus_gz_ge:
\forall (n: nat).(\forall (k: nat).(\forall (h: nat).((le k h) \to (eq A
(aplus gz (ASort h n) k) (ASort (minus h k) n)))))
\def
gz (aplus gz (ASort (S n0) n) k0)) (aplus_asucc gz k0 (ASort (S n0) n)))
(ASort (minus n0 k0) n) (IH n0 H_y)))))) h)))) k)).
-theorem next_plus_gz:
+lemma next_plus_gz:
\forall (n: nat).(\forall (h: nat).(eq nat (next_plus gz n h) (plus h n)))
\def
\lambda (n: nat).(\lambda (h: nat).(nat_ind (\lambda (n0: nat).(eq nat
nat).(\lambda (H: (eq nat (next_plus gz n n0) (plus n0 n))).(f_equal nat nat
S (next_plus gz n n0) (plus n0 n) H))) h)).
-theorem leqz_leq:
+lemma leqz_leq:
\forall (a1: A).(\forall (a2: A).((leq gz a1 a2) \to (leqz a1 a2)))
\def
\lambda (a1: A).(\lambda (a2: A).(\lambda (H: (leq gz a1 a2)).(leq_ind gz
a3)).(\lambda (a4: A).(\lambda (a5: A).(\lambda (_: (leq gz a4 a5)).(\lambda
(H3: (leqz a4 a5)).(leqz_head a0 a3 H1 a4 a5 H3))))))))) a1 a2 H))).
-theorem leq_leqz:
+lemma leq_leqz:
\forall (a1: A).(\forall (a2: A).((leqz a1 a2) \to (leq gz a1 a2)))
\def
\lambda (a1: A).(\lambda (a2: A).(\lambda (H: (leqz a1 a2)).(leqz_ind
include "basic_1/leq/props.ma".
-theorem ex1__leq_sort_SS:
+fact ex1__leq_sort_SS:
\forall (g: G).(\forall (k: nat).(\forall (n: nat).(leq g (ASort k n) (asucc
g (asucc g (ASort (S (S k)) n))))))
\def
\lambda (g: G).(\lambda (k: nat).(\lambda (n: nat).(leq_refl g (asucc g
(asucc g (ASort (S (S k)) n)))))).
-theorem ex1_arity:
+lemma ex1_arity:
\forall (g: G).(arity g ex1_c ex1_t (ASort O O))
\def
\lambda (g: G).(arity_appl g (CHead (CHead (CHead (CSort O) (Bind Abst)
(CHead (CHead (CHead (CSort O) (Bind Abst) (TSort O)) (Bind Abst) (TSort O))
(Bind Abst) (TLRef O)) (Bind Abst) (TLRef (S (S O)))) O))).
-theorem ex1_ty3:
+lemma ex1_ty3:
\forall (g: G).(\forall (u: T).((ty3 g ex1_c ex1_t u) \to (\forall (P:
Prop).P)))
\def
include "basic_1/arity/fwd.ma".
-theorem ex2_nf2:
+lemma ex2_nf2:
nf2 ex2_c ex2_t
\def
\lambda (t2: T).(\lambda (H: (pr2 (CSort O) (THead (Flat Appl) (TSort O)
(TSort O) (TSort O)) (THead (Bind x0) x5 (THead (Flat Appl) (lift (S O) O
(TSort O)) x3))) H9)) t2 H8))))))))))))))) H1)) H0))).
-theorem ex2_arity:
+lemma ex2_arity:
\forall (g: G).(\forall (a: A).((arity g ex2_c ex2_t a) \to (\forall (P:
Prop).P)))
\def
include "basic_1/flt/defs.ma".
-theorem flt_wf__q_ind:
+fact flt_wf__q_ind:
\forall (P: ((C \to (T \to Prop)))).(((\forall (n: nat).((\lambda (P0: ((C
\to (T \to Prop)))).(\lambda (n0: nat).(\forall (c: C).(\forall (t: T).((eq
nat (fweight c t) n0) \to (P0 c t)))))) P n))) \to (\forall (c: C).(\forall
C).(\forall (t: T).((eq nat (fweight c t) n) \to (P c t))))))).(\lambda (c:
C).(\lambda (t: T).(H (fweight c t) c t (refl_equal nat (fweight c t))))))).
-theorem flt_wf_ind:
+lemma flt_wf_ind:
\forall (P: ((C \to (T \to Prop)))).(((\forall (c2: C).(\forall (t2:
T).(((\forall (c1: C).(\forall (t1: T).((flt c1 t1 c2 t2) \to (P c1 t1)))))
\to (P c2 t2))))) \to (\forall (c: C).(\forall (t: T).(P c t))))
include "basic_1/C/props.ma".
-theorem flt_thead_sx:
+lemma flt_thead_sx:
\forall (k: K).(\forall (c: C).(\forall (u: T).(\forall (t: T).(flt c u c
(THead k u t)))))
\def
(tweight u) (S (plus (tweight u) (tweight t))) (cweight c) (le_n_S (tweight
u) (plus (tweight u) (tweight t)) (le_plus_l (tweight u) (tweight t))))))).
-theorem flt_thead_dx:
+lemma flt_thead_dx:
\forall (k: K).(\forall (c: C).(\forall (u: T).(\forall (t: T).(flt c t c
(THead k u t)))))
\def
(tweight t) (S (plus (tweight u) (tweight t))) (cweight c) (le_n_S (tweight
t) (plus (tweight u) (tweight t)) (le_plus_r (tweight u) (tweight t))))))).
-theorem flt_shift:
+lemma flt_shift:
\forall (k: K).(\forall (c: C).(\forall (u: T).(\forall (t: T).(flt (CHead c
k u) t c (THead k u t)))))
\def
(plus (tweight u) (tweight t)))) (plus_n_Sm (cweight c) (plus (tweight u)
(tweight t))))))).
-theorem flt_arith0:
+lemma flt_arith0:
\forall (k: K).(\forall (c: C).(\forall (t: T).(\forall (i: nat).(flt c t
(CHead c k t) (TLRef i)))))
\def
\lambda (_: K).(\lambda (c: C).(\lambda (t: T).(\lambda (_:
nat).(lt_x_plus_x_Sy (plus (cweight c) (tweight t)) O)))).
-theorem flt_arith1:
+lemma flt_arith1:
\forall (k1: K).(\forall (c1: C).(\forall (c2: C).(\forall (t1: T).((cle
(CHead c1 k1 t1) c2) \to (\forall (k2: K).(\forall (t2: T).(\forall (i:
nat).(flt c1 t1 (CHead c2 k2 t2) (TLRef i)))))))))
(tweight t2)) (S O)) (plus_sym (plus (cweight c2) (tweight t2)) (S
O))))))))))).
-theorem flt_arith2:
+lemma flt_arith2:
\forall (c1: C).(\forall (c2: C).(\forall (t1: T).(\forall (i: nat).((flt c1
t1 c2 (TLRef i)) \to (\forall (k2: K).(\forall (t2: T).(\forall (j: nat).(flt
c1 t1 (CHead c2 k2 t2) (TLRef j)))))))))
t2)) (S O)) H (le_plus_plus (cweight c2) (plus (cweight c2) (tweight t2)) (S
O) (S O) (le_plus_l (cweight c2) (tweight t2)) (le_n (S O))))))))))).
-theorem cle_flt_trans:
+lemma cle_flt_trans:
\forall (c1: C).(\forall (c2: C).((cle c1 c2) \to (\forall (c3: C).(\forall
(u2: T).(\forall (u3: T).((flt c2 u2 c3 u3) \to (flt c1 u2 c3 u3)))))))
\def
include "basic_1/fsubst0/defs.ma".
-theorem fsubst0_ind:
+implied lemma fsubst0_ind:
\forall (i: nat).(\forall (v: T).(\forall (c1: C).(\forall (t1: T).(\forall
(P: ((C \to (T \to Prop)))).(((\forall (t2: T).((subst0 i v t1 t2) \to (P c1
t2)))) \to (((\forall (c2: C).((csubst0 i v c1 c2) \to (P c2 t1)))) \to
[(fsubst0_snd x x0) \Rightarrow (f x x0) | (fsubst0_fst x x0) \Rightarrow (f0
x x0) | (fsubst0_both x x0 x1 x2) \Rightarrow (f1 x x0 x1 x2)]))))))))))).
-theorem fsubst0_gen_base:
+lemma fsubst0_gen_base:
\forall (c1: C).(\forall (c2: C).(\forall (t1: T).(\forall (t2: T).(\forall
(v: T).(\forall (i: nat).((fsubst0 i v c1 t1 c2 t2) \to (or3 (land (eq C c1
c2) (subst0 i v t1 t2)) (land (eq T t1 t2) (csubst0 i v c1 c2)) (land (subst0
include "basic_1/clear/drop.ma".
-theorem clear_getl_trans:
+lemma clear_getl_trans:
\forall (i: nat).(\forall (c2: C).(\forall (c3: C).((getl i c2 c3) \to
(\forall (c1: C).((clear c1 c2) \to (getl i c1 c3))))))
\def
t) c3)).(\lambda (H4: (clear c1 (CHead c (Flat f) t))).(clear_gen_flat_r f c1
c t H4 (getl (S n) c1 c3))))) k H1 H2))))))))) c2)))) i).
-theorem getl_clear_trans:
+lemma getl_clear_trans:
\forall (i: nat).(\forall (c1: C).(\forall (c2: C).((getl i c1 c2) \to
(\forall (c3: C).((clear c2 c3) \to (getl i c1 c3))))))
\def
c)) (getl_intro i c1 (CHead x1 (Bind x0) x2) x H2 H6) c3 (clear_gen_bind x0
x1 c3 x2 H7)))))))) H4))))) H1))))))).
-theorem getl_clear_bind:
+lemma getl_clear_bind:
\forall (b: B).(\forall (c: C).(\forall (e1: C).(\forall (v: T).((clear c
(CHead e1 (Bind b) v)) \to (\forall (e2: C).(\forall (n: nat).((getl n e1 e2)
\to (getl (S n) c e2))))))))
(Bind b) v))).(getl_flat c0 e2 (S n) (H e1 v (clear_gen_flat f c0 (CHead e1
(Bind b) v) t H2) e2 n H1) f t))) k H0))))))))))) c)).
-theorem getl_clear_conf:
+lemma getl_clear_conf:
\forall (i: nat).(\forall (c1: C).(\forall (c3: C).((getl i c1 c3) \to
(\forall (c2: C).((clear c1 c2) \to (getl i c2 c3))))))
\def
include "basic_1/getl/props.ma".
-theorem getl_dec:
+lemma getl_dec:
\forall (c: C).(\forall (i: nat).(or (ex_3 C B T (\lambda (e: C).(\lambda
(b: B).(\lambda (v: T).(getl i c (CHead e (Bind b) v)))))) (\forall (d:
C).((getl i c d) \to (\forall (P: Prop).P)))))
include "basic_1/clear/drop.ma".
-theorem getl_drop:
+lemma getl_drop:
\forall (b: B).(\forall (c: C).(\forall (e: C).(\forall (u: T).(\forall (h:
nat).((getl h c (CHead e (Bind b) u)) \to (drop (S h) O c e))))))
\def
(\lambda (n0: nat).(drop n0 O c0 e)) (H e u (r k n) (getl_gen_S k c0 (CHead e
(Bind b) u) t n H1)) (r k (S n)) (r_S k n)) t)))) h)))))))) c)).
-theorem getl_drop_conf_lt:
+lemma getl_drop_conf_lt:
\forall (b: B).(\forall (c: C).(\forall (c0: C).(\forall (u: T).(\forall (i:
nat).((getl i c (CHead c0 (Bind b) u)) \to (\forall (e: C).(\forall (h:
nat).(\forall (d: nat).((drop h (S (plus i d)) c e) \to (ex3_2 T C (\lambda
H21)))))) e H11))))))))) (drop_gen_skip_l c0 e t h (plus (S i0) d) k
H9))))))) i H1 H7 IHx)))) k0 H5 H6))))))) x H3 H4)))) H2)))))))))))))) c)).
-theorem getl_drop_conf_ge:
+lemma getl_drop_conf_ge:
\forall (i: nat).(\forall (a: C).(\forall (c: C).((getl i c a) \to (\forall
(e: C).(\forall (h: nat).(\forall (d: nat).((drop h d c e) \to ((le (plus d
h) i) \to (getl (minus i h) e a)))))))))
x)).(\lambda (H4: (clear x a)).(getl_intro (minus i h) e a x (drop_conf_ge i
x c H3 e h d H0 H1) H4)))) H2)))))))))).
-theorem getl_conf_ge_drop:
+lemma getl_conf_ge_drop:
\forall (b: B).(\forall (c1: C).(\forall (e: C).(\forall (u: T).(\forall (i:
nat).((getl i c1 (CHead e (Bind b) u)) \to (\forall (c2: C).((drop (S O) i c1
c2) \to (drop i O c2 e))))))))
i))) (le_n (S i)) (plus i (S O)) (plus_sym i (S O)))) i (minus_Sx_SO i)) in
H3)))))))).
-theorem getl_drop_conf_rev:
+lemma getl_drop_conf_rev:
\forall (j: nat).(\forall (e1: C).(\forall (e2: C).((drop j O e1 e2) \to
(\forall (b: B).(\forall (c2: C).(\forall (v2: T).(\forall (i: nat).((getl i
c2 (CHead e2 (Bind b) v2)) \to (ex2 C (\lambda (c1: C).(drop j O c1 c2))
nat).(\lambda (H0: (getl i c2 (CHead e2 (Bind b) v2))).(drop_conf_rev j e1 e2
H c2 (S i) (getl_drop b c2 e2 v2 i H0)))))))))).
-theorem drop_getl_trans_lt:
+lemma drop_getl_trans_lt:
\forall (i: nat).(\forall (d: nat).((lt i d) \to (\forall (c1: C).(\forall
(c2: C).(\forall (h: nat).((drop h d c1 c2) \to (\forall (b: B).(\forall (e2:
C).(\forall (v: T).((getl i c2 (CHead e2 (Bind b) v)) \to (ex2 C (\lambda
(le_S_n (S i) (S d) (le_S (S (S i)) (S d) (le_n_S (S i) d H)))) c1 c2 h H0 x
H3))))) H2)))))))))))).
-theorem drop_getl_trans_le:
+lemma drop_getl_trans_le:
\forall (i: nat).(\forall (d: nat).((le i d) \to (\forall (c1: C).(\forall
(c2: C).(\forall (h: nat).((drop h d c1 c2) \to (\forall (e2: C).((getl i c2
e2) \to (ex3_2 C C (\lambda (e0: C).(\lambda (_: C).(drop i O c1 e0)))
(\lambda (_: C).(\lambda (e1: C).(clear e1 e2))) x0 x H6 H7 H4)))) H5)))))
H2)))))))))).
-theorem drop_getl_trans_ge:
+lemma drop_getl_trans_ge:
\forall (i: nat).(\forall (c1: C).(\forall (c2: C).(\forall (d:
nat).(\forall (h: nat).((drop h d c1 c2) \to (\forall (e2: C).((getl i c2 e2)
\to ((le d i) \to (getl (plus i h) c1 e2)))))))))
C).(\lambda (H3: (drop i O c2 x)).(\lambda (H4: (clear x e2)).(getl_intro
(plus i h) c1 e2 x (drop_trans_ge i c1 c2 d h H x H3 H1) H4)))) H2)))))))))).
-theorem getl_drop_trans:
+lemma getl_drop_trans:
\forall (c1: C).(\forall (c2: C).(\forall (h: nat).((getl h c1 c2) \to
(\forall (e2: C).(\forall (i: nat).((drop (S i) O c2 e2) \to (drop (S (plus i
h)) O c1 e2)))))))
include "basic_1/flt/props.ma".
-theorem getl_flt:
+lemma getl_flt:
\forall (b: B).(\forall (c: C).(\forall (e: C).(\forall (u: T).(\forall (i:
nat).((getl i c (CHead e (Bind b) u)) \to (flt e u c (TLRef i)))))))
\def
include "basic_1/clear/fwd.ma".
-theorem getl_ind:
+implied lemma getl_ind:
\forall (h: nat).(\forall (c1: C).(\forall (c2: C).(\forall (P:
Prop).(((\forall (e: C).((drop h O c1 e) \to ((clear e c2) \to P)))) \to
((getl h c1 c2) \to P)))))
P))))).(\lambda (g: (getl h c1 c2)).(match g with [(getl_intro x x0 x1)
\Rightarrow (f x x0 x1)])))))).
-theorem getl_gen_all:
+lemma getl_gen_all:
\forall (c1: C).(\forall (c2: C).(\forall (i: nat).((getl i c1 c2) \to (ex2
C (\lambda (e: C).(drop i O c1 e)) (\lambda (e: C).(clear e c2))))))
\def
(H1: (clear e c2)).(ex_intro2 C (\lambda (e0: C).(drop i O c1 e0)) (\lambda
(e0: C).(clear e0 c2)) e H0 H1)))) H)))).
-theorem getl_gen_sort:
+lemma getl_gen_sort:
\forall (n: nat).(\forall (h: nat).(\forall (x: C).((getl h (CSort n) x) \to
(\forall (P: Prop).P))))
\def
(CSort n) H3) in (clear_gen_sort x n H6 P))))) (drop_gen_sort n h O x0
H1))))) H0)))))).
-theorem getl_gen_O:
+lemma getl_gen_O:
\forall (e: C).(\forall (x: C).((getl O e x) \to (clear e x)))
\def
\lambda (e: C).(\lambda (x: C).(\lambda (H: (getl O e x)).(let H0 \def
(drop O O e x0)).(\lambda (H2: (clear x0 x)).(let H3 \def (eq_ind_r C x0
(\lambda (c: C).(clear c x)) H2 e (drop_gen_refl e x0 H1)) in H3)))) H0)))).
-theorem getl_gen_S:
+lemma getl_gen_S:
\forall (k: K).(\forall (c: C).(\forall (x: C).(\forall (u: T).(\forall (h:
nat).((getl (S h) (CHead c k u) x) \to (getl (r k h) c x))))))
\def
C).(\lambda (H1: (drop (S h) O (CHead c k u) x0)).(\lambda (H2: (clear x0
x)).(getl_intro (r k h) c x x0 (drop_gen_drop k c x0 u h H1) H2)))) H0))))))).
-theorem getl_gen_2:
+lemma getl_gen_2:
\forall (c1: C).(\forall (c2: C).(\forall (i: nat).((getl i c1 c2) \to (ex_3
B C T (\lambda (b: B).(\lambda (c: C).(\lambda (v: T).(eq C c2 (CHead c (Bind
b) v)))))))))
(Bind b) v))))) x0 x1 x2 (refl_equal C (CHead x1 (Bind x0) x2))) c2 H4))))))
H3))))) H0))))).
-theorem getl_gen_flat:
+lemma getl_gen_flat:
\forall (f: F).(\forall (e: C).(\forall (d: C).(\forall (v: T).(\forall (i:
nat).((getl i (CHead e (Flat f) v) d) \to (getl i e d))))))
\def
(getl n e d)))).(\lambda (H0: (getl (S n) (CHead e (Flat f) v)
d)).(getl_gen_S (Flat f) e d v n H0)))) i))))).
-theorem getl_gen_bind:
+lemma getl_gen_bind:
\forall (b: B).(\forall (e: C).(\forall (d: C).(\forall (v: T).(\forall (i:
nat).((getl i (CHead e (Bind b) v) d) \to (or (land (eq nat i O) (eq C d
(CHead e (Bind b) v))) (ex2 nat (\lambda (j: nat).(eq nat i (S j))) (\lambda
include "basic_1/getl/clear.ma".
-theorem getl_conf_le:
+lemma getl_conf_le:
\forall (i: nat).(\forall (a: C).(\forall (c: C).((getl i c a) \to (\forall
(e: C).(\forall (h: nat).((getl h c e) \to ((le h i) \to (getl (minus i h) e
a))))))))
include "basic_1/drop/props.ma".
-theorem getl_refl:
+lemma getl_refl:
\forall (b: B).(\forall (c: C).(\forall (u: T).(getl O (CHead c (Bind b) u)
(CHead c (Bind b) u))))
\def
b) u) (CHead c (Bind b) u) (CHead c (Bind b) u) (drop_refl (CHead c (Bind b)
u)) (clear_bind b c u)))).
-theorem getl_head:
+lemma getl_head:
\forall (k: K).(\forall (h: nat).(\forall (c: C).(\forall (e: C).((getl (r k
h) c e) \to (\forall (u: T).(getl (S h) (CHead c k u) e))))))
\def
(drop (r k h) O c x)).(\lambda (H2: (clear x e)).(getl_intro (S h) (CHead c k
u) e x (drop_drop k h c x H1 u) H2)))) H0))))))).
-theorem getl_flat:
+lemma getl_flat:
\forall (c: C).(\forall (e: C).(\forall (h: nat).((getl h c e) \to (\forall
(f: F).(\forall (u: T).(getl h (CHead c (Flat f) u) e))))))
\def
(S h0) O c x)).(getl_intro (S h0) (CHead c (Flat f) u) e x (drop_drop (Flat
f) h0 c x H3 u) H2)))) h H1)))) H0))))))).
-theorem getl_ctail:
+lemma getl_ctail:
\forall (b: B).(\forall (c: C).(\forall (d: C).(\forall (u: T).(\forall (i:
nat).((getl i c (CHead d (Bind b) u)) \to (\forall (k: K).(\forall (v:
T).(getl i (CTail k v c) (CHead (CTail k v d) (Bind b) u)))))))))
include "basic_1/tlist/defs.ma".
-theorem iso_ind:
+implied lemma iso_ind:
\forall (P: ((T \to (T \to Prop)))).(((\forall (n1: nat).(\forall (n2:
nat).(P (TSort n1) (TSort n2))))) \to (((\forall (i1: nat).(\forall (i2:
nat).(P (TLRef i1) (TLRef i2))))) \to (((\forall (v1: T).(\forall (v2:
(f x x0) | (iso_lref x x0) \Rightarrow (f0 x x0) | (iso_head x x0 x1 x2 x3)
\Rightarrow (f1 x x0 x1 x2 x3)]))))))).
-theorem iso_gen_sort:
+lemma iso_gen_sort:
\forall (u2: T).(\forall (n1: nat).((iso (TSort n1) u2) \to (ex nat (\lambda
(n2: nat).(eq T u2 (TSort n2))))))
\def
n1) H1) in (False_ind (ex nat (\lambda (n2: nat).(eq T (THead k v2 t2) (TSort
n2)))) H2)))))))) y u2 H0))) H))).
-theorem iso_gen_lref:
+lemma iso_gen_lref:
\forall (u2: T).(\forall (n1: nat).((iso (TLRef n1) u2) \to (ex nat (\lambda
(n2: nat).(eq T u2 (TLRef n2))))))
\def
_ _) \Rightarrow True])) I (TLRef n1) H1) in (False_ind (ex nat (\lambda (n2:
nat).(eq T (THead k v2 t2) (TLRef n2)))) H2)))))))) y u2 H0))) H))).
-theorem iso_gen_head:
+lemma iso_gen_head:
\forall (k: K).(\forall (v1: T).(\forall (t1: T).(\forall (u2: T).((iso
(THead k v1 t1) u2) \to (ex_2 T T (\lambda (v2: T).(\lambda (t2: T).(eq T u2
(THead k v2 t2)))))))))
T).(eq T (THead k v2 t2) (THead k v3 t3)))) v2 t2 (refl_equal T (THead k v2
t2))) k0 H6)))) H3)) H2)))))))) y u2 H0))) H))))).
-theorem iso_flats_lref_bind_false:
+lemma iso_flats_lref_bind_false:
\forall (f: F).(\forall (b: B).(\forall (i: nat).(\forall (v: T).(\forall
(t: T).(\forall (vs: TList).((iso (THeads (Flat f) vs (TLRef i)) (THead (Bind
b) v t)) \to (\forall (P: Prop).P)))))))
\Rightarrow True | (Flat _) \Rightarrow False])])) I (THead (Flat f) x0 x1)
H2) in (False_ind P H3))))) H1)))))))) vs)))))).
-theorem iso_flats_flat_bind_false:
+lemma iso_flats_flat_bind_false:
\forall (f1: F).(\forall (f2: F).(\forall (b: B).(\forall (v: T).(\forall
(v2: T).(\forall (t: T).(\forall (t2: T).(\forall (vs: TList).((iso (THeads
(Flat f1) vs (THead (Flat f2) v2 t2)) (THead (Bind b) v t)) \to (\forall (P:
include "basic_1/iso/fwd.ma".
-theorem iso_refl:
+lemma iso_refl:
\forall (t: T).(iso t t)
\def
\lambda (t: T).(T_ind (\lambda (t0: T).(iso t0 t0)) (\lambda (n:
include "basic_1/leq/props.ma".
-theorem asucc_repl:
+lemma asucc_repl:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).((leq g a1 a2) \to (leq g
(asucc g a1) (asucc g a2)))))
\def
(leq g a5 a6)).(\lambda (H3: (leq g (asucc g a5) (asucc g a6))).(leq_head g
a3 a4 H0 (asucc g a5) (asucc g a6) H3))))))))) a1 a2 H)))).
-theorem asucc_inj:
+lemma asucc_inj:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).((leq g (asucc g a1) (asucc
g a2)) \to (leq g a1 a2))))
\def
(\lambda (a5: A).(leq g a a5)) H5 a3 H10) in (leq_head g a a3 H12 a0 a4 (H0
a4 H11)))))) H8))))))) H4)))))))) a2)))))) a1)).
-theorem leq_asucc:
+lemma leq_asucc:
\forall (g: G).(\forall (a: A).(ex A (\lambda (a0: A).(leq g a (asucc g
a0)))))
\def
(AHead a0 x) (leq_head g a0 a0 (leq_refl g a0) a1 (asucc g x) H2)))) H1))))))
a)).
-theorem leq_ahead_asucc_false:
+lemma leq_ahead_asucc_false:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).((leq g (AHead a1 a2)
(asucc g a1)) \to (\forall (P: Prop).P))))
\def
(eq_ind_r A x0 (\lambda (a3: A).(leq g (AHead a a0) a3)) H3 a H8) in
(leq_ahead_false_1 g a a0 H10 P))))) H6))))))) H2)))))))))) a1)).
-theorem leq_asucc_false:
+lemma leq_asucc_false:
\forall (g: G).(\forall (a: A).((leq g (asucc g a) a) \to (\forall (P:
Prop).P)))
\def
include "basic_1/leq/defs.ma".
-let rec leq_ind (g: G) (P: (A \to (A \to Prop))) (f: (\forall (h1:
+implied let rec leq_ind (g: G) (P: (A \to (A \to Prop))) (f: (\forall (h1:
nat).(\forall (h2: nat).(\forall (n1: nat).(\forall (n2: nat).(\forall (k:
nat).((eq A (aplus g (ASort h1 n1) k) (aplus g (ASort h2 n2) k)) \to (P
(ASort h1 n1) (ASort h2 n2))))))))) (f0: (\forall (a1: A).(\forall (a2:
\Rightarrow (f0 a1 a2 l0 ((leq_ind g P f f0) a1 a2 l0) a3 a4 l1 ((leq_ind g P
f f0) a3 a4 l1))].
-theorem leq_gen_sort1:
+lemma leq_gen_sort1:
\forall (g: G).(\forall (h1: nat).(\forall (n1: nat).(\forall (a2: A).((leq
g (ASort h1 n1) a2) \to (ex2_3 nat nat nat (\lambda (n2: nat).(\lambda (h2:
nat).(\lambda (k: nat).(eq A (aplus g (ASort h1 n1) k) (aplus g (ASort h2 n2)
(\lambda (n2: nat).(\lambda (h2: nat).(\lambda (_: nat).(eq A (AHead a3 a5)
(ASort h2 n2)))))) H6))))))))))) y a2 H0))) H))))).
-theorem leq_gen_head1:
+lemma leq_gen_head1:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).(\forall (a: A).((leq g
(AHead a1 a2) a) \to (ex3_2 A A (\lambda (a3: A).(\lambda (_: A).(leq g a1
a3))) (\lambda (_: A).(\lambda (a4: A).(leq g a2 a4))) (\lambda (a3:
(a6: A).(\lambda (a7: A).(eq A (AHead a3 a5) (AHead a6 a7)))) a3 a5 H12 H10
(refl_equal A (AHead a3 a5))))))))) H6))))))))))) y a H0))) H))))).
-theorem leq_gen_sort2:
+lemma leq_gen_sort2:
\forall (g: G).(\forall (h1: nat).(\forall (n1: nat).(\forall (a2: A).((leq
g a2 (ASort h1 n1)) \to (ex2_3 nat nat nat (\lambda (n2: nat).(\lambda (h2:
nat).(\lambda (k: nat).(eq A (aplus g (ASort h2 n2) k) (aplus g (ASort h1 n1)
(\lambda (n2: nat).(\lambda (h2: nat).(\lambda (_: nat).(eq A (AHead a1 a4)
(ASort h2 n2)))))) H6))))))))))) a2 y H0))) H))))).
-theorem leq_gen_head2:
+lemma leq_gen_head2:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).(\forall (a: A).((leq g a
(AHead a1 a2)) \to (ex3_2 A A (\lambda (a3: A).(\lambda (_: A).(leq g a3
a1))) (\lambda (_: A).(\lambda (a4: A).(leq g a4 a2))) (\lambda (a3:
(a6: A).(\lambda (a7: A).(eq A (AHead a0 a4) (AHead a6 a7)))) a0 a4 H12 H10
(refl_equal A (AHead a0 a4))))))))) H6))))))))))) a y H0))) H))))).
-theorem ahead_inj_snd:
+lemma ahead_inj_snd:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).(\forall (a3: A).(\forall
(a4: A).((leq g (AHead a1 a2) (AHead a3 a4)) \to (leq g a2 a4))))))
\def
include "basic_1/aplus/props.ma".
-theorem leq_refl:
+lemma leq_refl:
\forall (g: G).(\forall (a: A).(leq g a a))
\def
\lambda (g: G).(\lambda (a: A).(A_ind (\lambda (a0: A).(leq g a0 a0))
a0)).(\lambda (a1: A).(\lambda (H0: (leq g a1 a1)).(leq_head g a0 a0 H a1 a1
H0))))) a)).
-theorem leq_eq:
+lemma leq_eq:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).((eq A a1 a2) \to (leq g a1
a2))))
\def
\lambda (g: G).(\lambda (a1: A).(\lambda (a2: A).(\lambda (H: (eq A a1
a2)).(eq_ind A a1 (\lambda (a: A).(leq g a1 a)) (leq_refl g a1) a2 H)))).
-theorem leq_sym:
+lemma leq_sym:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).((leq g a1 a2) \to (leq g
a2 a1))))
\def
a3 a5) a)) (leq_head g a3 x0 (H1 x0 H6) a5 x1 (H3 x1 H7)) a0 H9)))))))
H5))))))))))))) a1 a2 H)))).
-theorem leq_ahead_false_1:
+lemma leq_ahead_false_1:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).((leq g (AHead a1 a2) a1)
\to (\forall (P: Prop).P))))
\def
H4 a0 H7) in (let H10 \def (eq_ind_r A x0 (\lambda (a3: A).(leq g (AHead a
a0) a3)) H3 a H8) in (H a0 H10 P))))) H6))))))) H2)))))))))) a1)).
-theorem leq_ahead_false_2:
+lemma leq_ahead_false_2:
\forall (g: G).(\forall (a2: A).(\forall (a1: A).((leq g (AHead a1 a2) a2)
\to (\forall (P: Prop).P))))
\def
include "basic_1/lift/props.ma".
-theorem lift_gen_sort:
+lemma lift_gen_sort:
\forall (h: nat).(\forall (d: nat).(\forall (n: nat).(\forall (t: T).((eq T
(TSort n) (lift h d t)) \to (eq T t (TSort n))))))
\def
(lift h d t0) (lift h (s k d) t1)) H2) in (False_ind (eq T (THead k t0 t1)
(TSort n)) H3))))))))) t)))).
-theorem lift_gen_lref:
+lemma lift_gen_lref:
\forall (t: T).(\forall (d: nat).(\forall (h: nat).(\forall (i: nat).((eq T
(TLRef i) (lift h d t)) \to (or (land (lt i d) (eq T t (TLRef i))) (land (le
(plus d h) i) (eq T t (TLRef (minus i h)))))))))
i))) (land (le (plus d h) i) (eq T (THead k t0 t1) (TLRef (minus i h)))))
H3)))))))))))) t).
-theorem lift_gen_lref_lt:
+lemma lift_gen_lref_lt:
\forall (h: nat).(\forall (d: nat).(\forall (n: nat).((lt n d) \to (\forall
(t: T).((eq T (TLRef n) (lift h d t)) \to (eq T t (TLRef n)))))))
\def
T (TLRef (minus n h)) (TLRef n)) H3 (lt_le_S n (plus d h) (le_plus_trans (S
n) d h H))) t H4))) H2)) H1)))))))).
-theorem lift_gen_lref_false:
+lemma lift_gen_lref_false:
\forall (h: nat).(\forall (d: nat).(\forall (n: nat).((le d n) \to ((lt n
(plus d h)) \to (\forall (t: T).((eq T (TLRef n) (lift h d t)) \to (\forall
(P: Prop).P)))))))
(H4: (le (plus d h) n)).(\lambda (_: (eq T t (TLRef (minus n h)))).(le_false
(plus d h) n P H4 H0))) H3)) H2)))))))))).
-theorem lift_gen_lref_ge:
+lemma lift_gen_lref_ge:
\forall (h: nat).(\forall (d: nat).(\forall (n: nat).((le d n) \to (\forall
(t: T).((eq T (TLRef (plus n h)) (lift h d t)) \to (eq T t (TLRef n)))))))
\def
h) h)) (\lambda (t0: T).(eq T t0 (TLRef n))) (f_equal nat T TLRef (minus
(plus n h) h) n (minus_plus_r n h)) t H4))) H2)) H1)))))))).
-theorem lift_gen_head:
+lemma lift_gen_head:
\forall (k: K).(\forall (u: T).(\forall (t: T).(\forall (x: T).(\forall (h:
nat).(\forall (d: nat).((eq T (THead k u t) (lift h d x)) \to (ex3_2 T T
(\lambda (y: T).(\lambda (z: T).(eq T x (THead k y z)))) (\lambda (y:
(lift h d t0)) (refl_equal T (lift h (s k d) t1))) u H6))) t H8))) k0 H7)))))
H4)) H3))))))))))) x)))).
-theorem lift_gen_bind:
+lemma lift_gen_bind:
\forall (b: B).(\forall (u: T).(\forall (t: T).(\forall (x: T).(\forall (h:
nat).(\forall (d: nat).((eq T (THead (Bind b) u t) (lift h d x)) \to (ex3_2 T
T (\lambda (y: T).(\lambda (z: T).(eq T x (THead (Bind b) y z)))) (\lambda
z)))) x0 x1 (refl_equal T (THead (Bind b) x0 x1)) (refl_equal T (lift h d
x0)) (refl_equal T (lift h (S d) x1))) u H2) t H3) x H1)))))) H0))))))))).
-theorem lift_gen_flat:
+lemma lift_gen_flat:
\forall (f: F).(\forall (u: T).(\forall (t: T).(\forall (x: T).(\forall (h:
nat).(\forall (d: nat).((eq T (THead (Flat f) u t) (lift h d x)) \to (ex3_2 T
T (\lambda (y: T).(\lambda (z: T).(eq T x (THead (Flat f) y z)))) (\lambda
(refl_equal T (lift h d x0)) (refl_equal T (lift h d x1))) u H2) t H3) x
H1)))))) H0))))))))).
-theorem lift_inj:
+lemma lift_inj:
\forall (x: T).(\forall (t: T).(\forall (h: nat).(\forall (d: nat).((eq T
(lift h d x) (lift h d t)) \to (eq T x t)))))
\def
(refl_equal K (Flat f)) (H x0 h d H4) (H0 x1 h d H5)))) t1 H3))))))
(lift_gen_flat f (lift h d t) (lift h d t0) t1 h d H2)))))))))))) k)) x).
-theorem lift_gen_lift:
+lemma lift_gen_lift:
\forall (t1: T).(\forall (x: T).(\forall (h1: nat).(\forall (h2:
nat).(\forall (d1: nat).(\forall (d2: nat).((le d1 d2) \to ((eq T (lift h1 d1
t1) (lift h2 (plus d2 h1) x)) \to (ex2 T (\lambda (t2: T).(eq T x (lift h1 d1
(lift h1 d1 t) (lift h1 d1 t0) x h2 (plus d2 h1) H4))))) k H2)))))))))))))
t1).
-theorem lifts_inj:
+lemma lifts_inj:
\forall (xs: TList).(\forall (ts: TList).(\forall (h: nat).(\forall (d:
nat).((eq TList (lifts h d xs) (lifts h d ts)) \to (eq TList xs ts)))))
\def
include "basic_1/T/fwd.ma".
-theorem lift_sort:
+lemma lift_sort:
\forall (n: nat).(\forall (h: nat).(\forall (d: nat).(eq T (lift h d (TSort
n)) (TSort n))))
\def
\lambda (n: nat).(\lambda (_: nat).(\lambda (_: nat).(refl_equal T (TSort
n)))).
-theorem lift_lref_lt:
+lemma lift_lref_lt:
\forall (n: nat).(\forall (h: nat).(\forall (d: nat).((lt n d) \to (eq T
(lift h d (TLRef n)) (TLRef n)))))
\def
\Rightarrow n | false \Rightarrow (plus n h)])) (TLRef n))) (refl_equal T
(TLRef n)) (blt n d) (sym_eq bool (blt n d) true (lt_blt d n H)))))).
-theorem lift_lref_ge:
+lemma lift_lref_ge:
\forall (n: nat).(\forall (h: nat).(\forall (d: nat).((le d n) \to (eq T
(lift h d (TLRef n)) (TLRef (plus n h))))))
\def
(refl_equal T (TLRef (plus n h))) (blt n d) (sym_eq bool (blt n d) false
(le_bge d n H)))))).
-theorem lift_head:
+lemma lift_head:
\forall (k: K).(\forall (u: T).(\forall (t: T).(\forall (h: nat).(\forall
(d: nat).(eq T (lift h d (THead k u t)) (THead k (lift h d u) (lift h (s k d)
t)))))))
\lambda (k: K).(\lambda (u: T).(\lambda (t: T).(\lambda (h: nat).(\lambda
(d: nat).(refl_equal T (THead k (lift h d u) (lift h (s k d) t))))))).
-theorem lift_bind:
+lemma lift_bind:
\forall (b: B).(\forall (u: T).(\forall (t: T).(\forall (h: nat).(\forall
(d: nat).(eq T (lift h d (THead (Bind b) u t)) (THead (Bind b) (lift h d u)
(lift h (S d) t)))))))
\lambda (b: B).(\lambda (u: T).(\lambda (t: T).(\lambda (h: nat).(\lambda
(d: nat).(refl_equal T (THead (Bind b) (lift h d u) (lift h (S d) t))))))).
-theorem lift_flat:
+lemma lift_flat:
\forall (f: F).(\forall (u: T).(\forall (t: T).(\forall (h: nat).(\forall
(d: nat).(eq T (lift h d (THead (Flat f) u t)) (THead (Flat f) (lift h d u)
(lift h d t)))))))
\lambda (f: F).(\lambda (u: T).(\lambda (t: T).(\lambda (h: nat).(\lambda
(d: nat).(refl_equal T (THead (Flat f) (lift h d u) (lift h d t))))))).
-theorem thead_x_lift_y_y:
+lemma thead_x_lift_y_y:
\forall (k: K).(\forall (t: T).(\forall (v: T).(\forall (h: nat).(\forall
(d: nat).((eq T (THead k v (lift h d t)) t) \to (\forall (P: Prop).P))))))
\def
k0 t0 t1 h d)) in (H7 (lift h d t0) h (s k0 d) H8 P)))))) H3)) H2))))))))))))
t)).
-theorem lift_r:
+lemma lift_r:
\forall (t: T).(\forall (d: nat).(eq T (lift O d t) t))
\def
\lambda (t: T).(T_ind (\lambda (t0: T).(\forall (d: nat).(eq T (lift O d t0)
t0) t0 (lift O (s k d) t1) t1 (refl_equal K k) (H d) (H0 (s k d))))) (lift O
d (THead k t0 t1)) (lift_head k t0 t1 O d)))))))) t).
-theorem lift_lref_gt:
+lemma lift_lref_gt:
\forall (d: nat).(\forall (n: nat).((lt d n) \to (eq T (lift (S O) d (TLRef
(pred n))) (TLRef n))))
\def
(pred n) (eq_ind nat n (\lambda (n0: nat).(le (S d) n0)) H (S (pred n))
(S_pred n d H))))))).
-theorem lift_tle:
+lemma lift_tle:
\forall (t: T).(\forall (h: nat).(\forall (d: nat).(tle t (lift h d t))))
\def
\lambda (t: T).(T_ind (\lambda (t0: T).(\forall (h: nat).(\forall (d:
(\lambda (x: nat).(plus x h)) d t0)) (tweight t1) (tweight (lref_map (\lambda
(x: nat).(plus x h)) (s k d) t1)) H_y H_y0))))))))))) t).
-theorem lifts_tapp:
+lemma lifts_tapp:
\forall (h: nat).(\forall (d: nat).(\forall (v: T).(\forall (vs: TList).(eq
TList (lifts h d (TApp vs v)) (TApp (lifts h d vs) (lift h d v))))))
\def
(TCons (lift h d t) (TApp (lifts h d t0) (lift h d v)))) (lifts h d (TApp t0
v)) H)))) vs)))).
-theorem lift_free:
+lemma lift_free:
\forall (t: T).(\forall (h: nat).(\forall (k: nat).(\forall (d:
nat).(\forall (e: nat).((le e (plus d h)) \to ((le d e) \to (eq T (lift k e
(lift h d t)) (lift (plus k h) d t))))))))
h (s k d) t1) k0 e)) (lift h d (THead k t0 t1)) (lift_head k t0 t1 h
d))))))))))))) t).
-theorem lift_free_sym:
+lemma lift_free_sym:
\forall (t: T).(\forall (h: nat).(\forall (k: nat).(\forall (d:
nat).(\forall (e: nat).((le e (plus d h)) \to ((le d e) \to (eq T (lift k e
(lift h d t)) (lift (plus h k) d t))))))))
e)).(eq_ind_r nat (plus k h) (\lambda (n: nat).(eq T (lift k e (lift h d t))
(lift n d t))) (lift_free t h k d e H H0) (plus h k) (plus_sym h k)))))))).
-theorem lift_d:
+lemma lift_d:
\forall (t: T).(\forall (h: nat).(\forall (k: nat).(\forall (d:
nat).(\forall (e: nat).((le e d) \to (eq T (lift h (plus k d) (lift k e t))
(lift k e (lift h d t))))))))
include "basic_1/tlt/props.ma".
-theorem lift_weight_map:
+lemma lift_weight_map:
\forall (t: T).(\forall (h: nat).(\forall (d: nat).(\forall (f: ((nat \to
nat))).(((\forall (m: nat).((le d m) \to (eq nat (f m) O)))) \to (eq nat
(weight_map f (lift h d t)) (weight_map f t))))))
(lift h d (THead (Flat f0) t0 t1)) (lift_head (Flat f0) t0 t1 h d)))
k)))))))))) t).
-theorem lift_weight:
+lemma lift_weight:
\forall (t: T).(\forall (h: nat).(\forall (d: nat).(eq nat (weight (lift h d
t)) (weight t))))
\def
(\lambda (_: nat).O) (\lambda (m: nat).(\lambda (_: (le d m)).(refl_equal nat
O)))))).
-theorem lift_weight_add:
+lemma lift_weight_add:
\forall (w: nat).(\forall (t: T).(\forall (h: nat).(\forall (d:
nat).(\forall (f: ((nat \to nat))).(\forall (g: ((nat \to nat))).(((\forall
(m: nat).((lt m d) \to (eq nat (g m) (f m))))) \to ((eq nat (g d) w) \to
(lift h d (THead (Flat f0) t0 t1)) (lift_head (Flat f0) t0 t1 h d)))
k))))))))))))) t)).
-theorem lift_weight_add_O:
+lemma lift_weight_add_O:
\forall (w: nat).(\forall (t: T).(\forall (h: nat).(\forall (f: ((nat \to
nat))).(eq nat (weight_map f (lift h O t)) (weight_map (wadd f w) (lift (S h)
O t))))))
(minus_n_O (wadd f w O)) (\lambda (m: nat).(\lambda (_: (le O m)).(refl_equal
nat (f m)))))))).
-theorem lift_tlt_dx:
+lemma lift_tlt_dx:
\forall (k: K).(\forall (u: T).(\forall (t: T).(\forall (h: nat).(\forall
(d: nat).(tlt t (THead k u (lift h d t)))))))
\def
include "basic_1/drop1/defs.ma".
-theorem lift1_lift1:
+lemma lift1_lift1:
\forall (is1: PList).(\forall (is2: PList).(\forall (t: T).(eq T (lift1 is1
(lift1 is2 t)) (lift1 (papp is1 is2) t))))
\def
T T lift n n n0 n0 (lift1 p (lift1 is2 t)) (lift1 (papp p is2) t) (refl_equal
nat n) (refl_equal nat n0) (H is2 t)))))))) is1).
-theorem lift1_xhg:
+lemma lift1_xhg:
\forall (hds: PList).(\forall (t: T).(eq T (lift1 (Ss hds) (lift (S O) O t))
(lift (S O) O (lift1 hds t))))
\def
p t))) (lift_d (lift1 p t) h (S O) d O (le_O_n d))) (S d) (refl_equal nat (S
d))) (lift1 (Ss p) (lift (S O) O t)) (H t))))))) hds).
-theorem lifts1_xhg:
+lemma lifts1_xhg:
\forall (hds: PList).(\forall (ts: TList).(eq TList (lifts1 (Ss hds) (lifts
(S O) O ts)) (lifts (S O) O (lifts1 hds ts))))
\def
(lifts1 hds t0)))) (lifts1 (Ss hds) (lifts (S O) O t0)) H) (lift1 (Ss hds)
(lift (S O) O t)) (lift1_xhg hds t))))) ts)).
-theorem lift1_free:
+lemma lift1_free:
\forall (hds: PList).(\forall (i: nat).(\forall (t: T).(eq T (lift1 hds
(lift (S i) O t)) (lift (S (trans hds i)) O (lift1 (ptrans hds i) t)))))
\def
include "basic_1/lift/props.ma".
-theorem lift1_sort:
+lemma lift1_sort:
\forall (n: nat).(\forall (is: PList).(eq T (lift1 is (TSort n)) (TSort n)))
\def
\lambda (n: nat).(\lambda (is: PList).(PList_ind (\lambda (p: PList).(eq T
(TSort n)) (TSort n))).(eq_ind_r T (TSort n) (\lambda (t: T).(eq T (lift n0
n1 t) (TSort n))) (refl_equal T (TSort n)) (lift1 p (TSort n)) H))))) is)).
-theorem lift1_lref:
+lemma lift1_lref:
\forall (hds: PList).(\forall (i: nat).(eq T (lift1 hds (TLRef i)) (TLRef
(trans hds i))))
\def
(TLRef (match (blt (trans p i) n0) with [true \Rightarrow (trans p i) | false
\Rightarrow (plus (trans p i) n)]))) (lift1 p (TLRef i)) (H i))))))) hds).
-theorem lift1_bind:
+lemma lift1_bind:
\forall (b: B).(\forall (hds: PList).(\forall (u: T).(\forall (t: T).(eq T
(lift1 hds (THead (Bind b) u t)) (THead (Bind b) (lift1 hds u) (lift1 (Ss
hds) t))))))
(lift_bind b (lift1 p u) (lift1 (Ss p) t) n n0)) (lift1 p (THead (Bind b) u
t)) (H u t)))))))) hds)).
-theorem lift1_flat:
+lemma lift1_flat:
\forall (f: F).(\forall (hds: PList).(\forall (u: T).(\forall (t: T).(eq T
(lift1 hds (THead (Flat f) u t)) (THead (Flat f) (lift1 hds u) (lift1 hds
t))))))
(lift1 p u) (lift1 p t))) (lift_flat f (lift1 p u) (lift1 p t) n n0)) (lift1
p (THead (Flat f) u t)) (H u t)))))))) hds)).
-theorem lift1_cons_tail:
+lemma lift1_cons_tail:
\forall (t: T).(\forall (h: nat).(\forall (d: nat).(\forall (hds: PList).(eq
T (lift1 (PConsTail hds h d) t) (lift1 hds (lift h d t))))))
\def
t))))) (refl_equal T (lift n n0 (lift1 p (lift h d t)))) (lift1 (PConsTail p
h d) t) H))))) hds)))).
-theorem lifts1_flat:
+lemma lifts1_flat:
\forall (f: F).(\forall (hds: PList).(\forall (t: T).(\forall (ts:
TList).(eq T (lift1 hds (THeads (Flat f) ts t)) (THeads (Flat f) (lifts1 hds
ts) (lift1 hds t))))))
(lift1 hds (THead (Flat f) t0 (THeads (Flat f) t1 t))) (lift1_flat f hds t0
(THeads (Flat f) t1 t)))))) ts)))).
-theorem lifts1_nil:
+lemma lifts1_nil:
\forall (ts: TList).(eq TList (lifts1 PNil ts) ts)
\def
\lambda (ts: TList).(TList_ind (\lambda (t: TList).(eq TList (lifts1 PNil t)
TList (TCons t t1) (TCons t t0))) (refl_equal TList (TCons t t0)) (lifts1
PNil t0) H)))) ts).
-theorem lifts1_cons:
+lemma lifts1_cons:
\forall (h: nat).(\forall (d: nat).(\forall (hds: PList).(\forall (ts:
TList).(eq TList (lifts1 (PCons h d hds) ts) (lifts h d (lifts1 hds ts))))))
\def
include "basic_1/llt/defs.ma".
-theorem llt_wf__q_ind:
+fact llt_wf__q_ind:
\forall (P: ((A \to Prop))).(((\forall (n: nat).((\lambda (P0: ((A \to
Prop))).(\lambda (n0: nat).(\forall (a: A).((eq nat (lweight a) n0) \to (P0
a))))) P n))) \to (\forall (a: A).(P a)))
n) \to (P a)))))).(\lambda (a: A).(H (lweight a) a (refl_equal nat (lweight
a)))))).
-theorem llt_wf_ind:
+lemma llt_wf_ind:
\forall (P: ((A \to Prop))).(((\forall (a2: A).(((\forall (a1: A).((llt a1
a2) \to (P a1)))) \to (P a2)))) \to (\forall (a: A).(P a)))
\def
include "basic_1/leq/fwd.ma".
-theorem lweight_repl:
+lemma lweight_repl:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).((leq g a1 a2) \to (eq nat
(lweight a1) (lweight a2)))))
\def
(lweight a0) (lweight a3) (lweight a4) (lweight a5) H1 H3)))))))))) a1 a2
H)))).
-theorem llt_repl:
+lemma llt_repl:
\forall (g: G).(\forall (a1: A).(\forall (a2: A).((leq g a1 a2) \to (\forall
(a3: A).((llt a1 a3) \to (llt a2 a3))))))
\def
a1) (lweight a2))).(\lambda (H0: (lt (lweight a2) (lweight a3))).(lt_trans
(lweight a1) (lweight a2) (lweight a3) H H0))))).
-theorem llt_head_sx:
+lemma llt_head_sx:
\forall (a1: A).(\forall (a2: A).(llt a1 (AHead a1 a2)))
\def
\lambda (a1: A).(\lambda (a2: A).(le_n_S (lweight a1) (plus (lweight a1)
(lweight a2)) (le_plus_l (lweight a1) (lweight a2)))).
-theorem llt_head_dx:
+lemma llt_head_dx:
\forall (a1: A).(\forall (a2: A).(llt a2 (AHead a1 a2)))
\def
\lambda (a1: A).(\lambda (a2: A).(le_n_S (lweight a2) (plus (lweight a1)
include "basic_1/next_plus/defs.ma".
-theorem next_plus_assoc:
+lemma next_plus_assoc:
\forall (g: G).(\forall (n: nat).(\forall (h1: nat).(\forall (h2: nat).(eq
nat (next_plus g (next_plus g n h1) h2) (next_plus g n (plus h1 h2))))))
\def
(next g (next_plus g n n0)) n1) (next g (next_plus g n (plus n0 n1))) H0)
(plus n0 (S n1)) (plus_n_Sm n0 n1)))) h2)))) h1))).
-theorem next_plus_next:
+lemma next_plus_next:
\forall (g: G).(\forall (n: nat).(\forall (h: nat).(eq nat (next_plus g
(next g n) h) (next g (next_plus g n h)))))
\def
h)))) (refl_equal nat (next g (next_plus g n h))) (next_plus g (next_plus g n
(S O)) h) (next_plus_assoc g n (S O) h)))).
-theorem next_plus_lt:
+lemma next_plus_lt:
\forall (g: G).(\forall (h: nat).(\forall (n: nat).(lt n (next_plus g (next
g n) h))))
\def
include "basic_1/arity/subst0.ma".
-theorem arity_nf2_inv_all:
+lemma arity_nf2_inv_all:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (a: A).((arity g c t
a) \to ((nf2 c t) \to (or3 (ex3_2 T T (\lambda (w: T).(\lambda (u: T).(eq T t
(THead (Bind Abst) w u)))) (\lambda (w: T).(\lambda (_: T).(nf2 c w)))
include "basic_1/C/props.ma".
-theorem nf2_dec:
+lemma nf2_dec:
\forall (c: C).(\forall (t1: T).(or (nf2 c t1) (ex2 T (\lambda (t2: T).((eq
T t1 t2) \to (\forall (P: Prop).P))) (\lambda (t2: T).(pr2 c t1 t2)))))
\def
include "basic_1/T/props.ma".
-theorem nf2_gen_lref:
+lemma nf2_gen_lref:
\forall (c: C).(\forall (d: C).(\forall (u: T).(\forall (i: nat).((getl i c
(CHead d (Bind Abbr) u)) \to ((nf2 c (TLRef i)) \to (\forall (P: Prop).P))))))
\def
(lift (S i) O u) (pr2_delta c d u i H (TLRef i) (TLRef i) (pr0_refl (TLRef
i)) (lift (S i) O u) (subst0_lref u i))) P))))))).
-theorem nf2_gen_abst:
+lemma nf2_gen_abst:
\forall (c: C).(\forall (u: T).(\forall (t: T).((nf2 c (THead (Bind Abst) u
t)) \to (land (nf2 c u) (nf2 (CHead c (Bind Abst) u) t)))))
\def
(\lambda (t0: T).(pr2 (CHead c (Bind Abst) u) t t0)) H0 t H1) in (eq_ind T t
(\lambda (t0: T).(eq T t t0)) (refl_equal T t) t2 H1))))))))).
-theorem nf2_gen_cast:
+lemma nf2_gen_cast:
\forall (c: C).(\forall (u: T).(\forall (t: T).((nf2 c (THead (Flat Cast) u
t)) \to (\forall (P: Prop).P))))
\def
(Flat Cast) u t))).(\lambda (P: Prop).(thead_x_y_y (Flat Cast) u t (H t
(pr2_free c (THead (Flat Cast) u t) t (pr0_tau t t (pr0_refl t) u))) P))))).
-theorem nf2_gen_beta:
+lemma nf2_gen_beta:
\forall (c: C).(\forall (u: T).(\forall (v: T).(\forall (t: T).((nf2 c
(THead (Flat Appl) u (THead (Bind Abst) v t))) \to (\forall (P: Prop).P)))))
\def
Abst) v t)) (THead (Bind Abbr) u t) (pr0_beta v u u (pr0_refl u) t t
(pr0_refl t))))) in (False_ind P H0))))))).
-theorem nf2_gen_flat:
+lemma nf2_gen_flat:
\forall (f: F).(\forall (c: C).(\forall (u: T).(\forall (t: T).((nf2 c
(THead (Flat f) u t)) \to (land (nf2 c u) (nf2 c t))))))
\def
(THead (Flat f) u t) (THead (Flat f) u t2) (H (THead (Flat f) u t2)
(pr2_head_2 c u t t2 (Flat f) (pr2_cflat c t t2 H0 f u)))) in H1)))))))).
-theorem nf2_gen__nf2_gen_aux:
+fact nf2_gen__nf2_gen_aux:
\forall (b: B).(\forall (x: T).(\forall (u: T).(\forall (d: nat).((eq T
(THead (Bind b) u (lift (S O) d x)) x) \to (\forall (P: Prop).P)))))
\def
(S O) (S d) t0)) (lift_bind b t t0 (S O) d)) in (H0 (lift (S O) d t) (S d) H8
P)))))) H3)) H2))))))))))) x)).
-theorem nf2_gen_abbr:
+lemma nf2_gen_abbr:
\forall (c: C).(\forall (u: T).(\forall (t: T).((nf2 c (THead (Bind Abbr) u
t)) \to (\forall (P: Prop).P))))
\def
(S O) O x)) x (pr0_zeta Abbr not_abbr_abst x x (pr0_refl x) u))) P))) H1)))
H0))))))).
-theorem nf2_gen_void:
+lemma nf2_gen_void:
\forall (c: C).(\forall (u: T).(\forall (t: T).((nf2 c (THead (Bind Void) u
(lift (S O) O t))) \to (\forall (P: Prop).P))))
\def
include "basic_1/iso/props.ma".
-theorem nf2_iso_appls_lref:
+lemma nf2_iso_appls_lref:
\forall (c: C).(\forall (i: nat).((nf2 c (TLRef i)) \to (\forall (vs:
TList).(\forall (u: T).((pr3 c (THeads (Flat Appl) vs (TLRef i)) u) \to (iso
(THeads (Flat Appl) vs (TLRef i)) u))))))
include "basic_1/drop1/fwd.ma".
-theorem nf2_lift1:
+lemma nf2_lift1:
\forall (e: C).(\forall (hds: PList).(\forall (c: C).(\forall (t: T).((drop1
hds c e) \to ((nf2 e t) \to (nf2 c (lift1 hds t)))))))
\def
include "basic_1/pr3/pr3.ma".
-theorem nf2_pr3_unfold:
+lemma nf2_pr3_unfold:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr3 c t1 t2) \to ((nf2 c
t1) \to (eq T t1 t2)))))
\def
include "basic_1/pr2/fwd.ma".
-theorem nf2_sort:
+lemma nf2_sort:
\forall (c: C).(\forall (n: nat).(nf2 c (TSort n)))
\def
\lambda (c: C).(\lambda (n: nat).(\lambda (t2: T).(\lambda (H: (pr2 c (TSort
n) t2)).(eq_ind_r T (TSort n) (\lambda (t: T).(eq T (TSort n) t)) (refl_equal
T (TSort n)) t2 (pr2_gen_sort c t2 n H))))).
-theorem nf2_csort_lref:
+lemma nf2_csort_lref:
\forall (n: nat).(\forall (i: nat).(nf2 (CSort n) (TLRef i)))
\def
\lambda (n: nat).(\lambda (i: nat).(\lambda (t2: T).(\lambda (H: (pr2 (CSort
u x0 t x1 (refl_equal K (Bind Abst)) (H x0 H4) (H0 x1 (H5 Abst u))) t2
H3)))))) H2)))))))).
-theorem nfs2_tapp:
+lemma nfs2_tapp:
\forall (c: C).(\forall (t: T).(\forall (ts: TList).((nfs2 c (TApp ts t))
\to (land (nfs2 c ts) (nf2 c t)))))
\def
(land (nf2 c t0) (nfs2 c t1)) (nf2 c t) (conj (nf2 c t0) (nfs2 c t1) H2 H5)
H6))) H4))))) H1)))))) ts))).
-theorem nf2_appls_lref:
+lemma nf2_appls_lref:
\forall (c: C).(\forall (i: nat).((nf2 c (TLRef i)) \to (\forall (vs:
TList).((nfs2 c vs) \to (nf2 c (THeads (Flat Appl) vs (TLRef i)))))))
\def
nat).(\lambda (H0: (nf2 c (TLRef i))).(let H_y \def (nf2_appls_lref c i H0
(TCons u TNil)) in (H_y (conj (nf2 c u) True H I))))))).
-theorem nf2_lref_abst:
+lemma nf2_lref_abst:
\forall (c: C).(\forall (e: C).(\forall (u: T).(\forall (i: nat).((getl i c
(CHead e (Bind Abst) u)) \to (nf2 c (TLRef i))))))
\def
(Bind Abst) u) i H (CHead x0 (Bind Abbr) x1) H3)) in (False_ind (eq T (TLRef
i) (lift (S i) O x1)) H6))) t2 H4))))) H2)) H1)))))))).
-theorem nf2_lift:
+lemma nf2_lift:
\forall (d: C).(\forall (t: T).((nf2 d t) \to (\forall (c: C).(\forall (h:
nat).(\forall (i: nat).((drop h i c d) \to (nf2 c (lift h i t))))))))
\def
include "basic_1/pr1/pr1.ma".
-theorem pc1_pr0_r:
+lemma pc1_pr0_r:
\forall (t1: T).(\forall (t2: T).((pr0 t1 t2) \to (pc1 t1 t2)))
\def
\lambda (t1: T).(\lambda (t2: T).(\lambda (H: (pr0 t1 t2)).(ex_intro2 T
(\lambda (t: T).(pr1 t1 t)) (\lambda (t: T).(pr1 t2 t)) t2 (pr1_pr0 t1 t2 H)
(pr1_refl t2)))).
-theorem pc1_pr0_x:
+lemma pc1_pr0_x:
\forall (t1: T).(\forall (t2: T).((pr0 t2 t1) \to (pc1 t1 t2)))
\def
\lambda (t1: T).(\lambda (t2: T).(\lambda (H: (pr0 t2 t1)).(ex_intro2 T
(\lambda (t: T).(pr1 t1 t)) (\lambda (t: T).(pr1 t2 t)) t1 (pr1_refl t1)
(pr1_pr0 t2 t1 H)))).
-theorem pc1_refl:
+lemma pc1_refl:
\forall (t: T).(pc1 t t)
\def
\lambda (t: T).(ex_intro2 T (\lambda (t0: T).(pr1 t t0)) (\lambda (t0:
T).(pr1 t t0)) t (pr1_refl t) (pr1_refl t)).
-theorem pc1_pr0_u:
+lemma pc1_pr0_u:
\forall (t2: T).(\forall (t1: T).((pr0 t1 t2) \to (\forall (t3: T).((pc1 t2
t3) \to (pc1 t1 t3)))))
\def
(t: T).(pr1 t1 t)) (\lambda (t: T).(pr1 t3 t)) x (pr1_sing t2 t1 H x H2)
H3)))) H1)))))).
-theorem pc1_s:
+lemma pc1_s:
\forall (t2: T).(\forall (t1: T).((pc1 t1 t2) \to (pc1 t2 t1)))
\def
\lambda (t2: T).(\lambda (t1: T).(\lambda (H: (pc1 t1 t2)).(let H0 \def H in
x)).(ex_intro2 T (\lambda (t: T).(pr1 t2 t)) (\lambda (t: T).(pr1 t1 t)) x H2
H1)))) H0)))).
-theorem pc1_head_1:
+lemma pc1_head_1:
\forall (u1: T).(\forall (u2: T).((pc1 u1 u2) \to (\forall (t: T).(\forall
(k: K).(pc1 (THead k u1 t) (THead k u2 t))))))
\def
t) t0)) (THead k x t) (pr1_head_1 u1 x H1 t k) (pr1_head_1 u2 x H2 t k)))))
H0)))))).
-theorem pc1_head_2:
+lemma pc1_head_2:
\forall (t1: T).(\forall (t2: T).((pc1 t1 t2) \to (\forall (u: T).(\forall
(k: K).(pc1 (THead k u t1) (THead k u t2))))))
\def
(pr1_t x0 t1 H5 x1 H7) (pr1_t x t3 H3 x1 H8))))) (pr1_confluence t2 x0 H6 x
H2))))) H4))))) H1)))))).
-theorem pc1_pr0_u2:
+lemma pc1_pr0_u2:
\forall (t0: T).(\forall (t1: T).((pr0 t0 t1) \to (\forall (t2: T).((pc1 t0
t2) \to (pc1 t1 t2)))))
\def
include "basic_1/csubst0/getl.ma".
-theorem pc3_pr2_fsubst0:
+lemma pc3_pr2_fsubst0:
\forall (c1: C).(\forall (t1: T).(\forall (t: T).((pr2 c1 t1 t) \to (\forall
(i: nat).(\forall (u: T).(\forall (c2: C).(\forall (t2: T).((fsubst0 i u c1
t1 c2 t2) \to (\forall (e: C).((getl i c1 (CHead e (Bind Abbr) u)) \to (pc3
(csubst0_getl_ge i0 i H7 c c0 u0 H5 (CHead d (Bind Abbr) u) H0) t2 t3 H1 t0
H2))))))))))) c2 t4 H3)))))))))))))))) c1 t1 t H)))).
-theorem pc3_pr2_fsubst0_back:
+lemma pc3_pr2_fsubst0_back:
\forall (c1: C).(\forall (t: T).(\forall (t1: T).((pr2 c1 t t1) \to (\forall
(i: nat).(\forall (u: T).(\forall (c2: C).(\forall (t2: T).((fsubst0 i u c1
t1 c2 t2) \to (\forall (e: C).((getl i c1 (CHead e (Bind Abbr) u)) \to (pc3
i0) c c0 u0 H5 (CHead e (Bind Abbr) u0) H6) t0 t0 (pr0_refl t0) t5
H4))))))))))) c2 t4 H3)))))))))))))))) c1 t t1 H)))).
-theorem pc3_fsubst0:
+lemma pc3_fsubst0:
\forall (c1: C).(\forall (t1: T).(\forall (t: T).((pc3 c1 t1 t) \to (\forall
(i: nat).(\forall (u: T).(\forall (c2: C).(\forall (t2: T).((fsubst0 i u c1
t1 c2 t2) \to (\forall (e: C).((getl i c1 (CHead e (Bind Abbr) u)) \to (pc3
include "basic_1/pc3/props.ma".
-theorem pc3_gen_sort:
+lemma pc3_gen_sort:
\forall (c: C).(\forall (m: nat).(\forall (n: nat).((pc3 c (TSort m) (TSort
n)) \to (eq nat m n))))
\def
\Rightarrow m | (THead _ _ _) \Rightarrow m])) (TSort m) (TSort n) H3) in
H4))))) H0))))).
-theorem pc3_gen_abst:
+lemma pc3_gen_abst:
\forall (c: C).(\forall (u1: T).(\forall (u2: T).(\forall (t1: T).(\forall
(t2: T).((pc3 c (THead (Bind Abst) u1 t1) (THead (Bind Abst) u2 t2)) \to
(land (pc3 c u1 u2) (\forall (b: B).(\forall (u: T).(pc3 (CHead c (Bind b) u)
(CHead c (Bind b) u) t1 x1 (H15 b u) t2 (H6 b u))))))))) H12))))))))
H7))))))) H3))))) H0))))))).
-theorem pc3_gen_abst_shift:
+lemma pc3_gen_abst_shift:
\forall (c: C).(\forall (u: T).(\forall (t1: T).(\forall (t2: T).((pc3 c
(THead (Bind Abst) u t1) (THead (Bind Abst) u t2)) \to (pc3 (CHead c (Bind
Abst) u) t1 t2)))))
((\forall (b: B).(\forall (u0: T).(pc3 (CHead c (Bind b) u0) t1 t2))))).(H2
Abst u))) H0))))))).
-theorem pc3_gen_lift:
+lemma pc3_gen_lift:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).(\forall (h: nat).(\forall
(d: nat).((pc3 c (lift h d t1) (lift h d t2)) \to (\forall (e: C).((drop h d
c e) \to (pc3 e t1 t2))))))))
(\lambda (t: T).(pr3 e t1 t)) H9 x0 (lift_inj x1 x0 h d H10)) in (pc3_pr3_t e
t1 x0 H11 t2 H6)))))) H7))))) H4))))) H1))))))))).
-theorem pc3_gen_not_abst:
+lemma pc3_gen_not_abst:
\forall (b: B).((not (eq B b Abst)) \to (\forall (c: C).(\forall (t1:
T).(\forall (t2: T).(\forall (u1: T).(\forall (u2: T).((pc3 c (THead (Bind b)
u1 t1) (THead (Bind Abst) u2 t2)) \to (pc3 (CHead c (Bind b) u1) t1 (lift (S
t2) (THead (Bind Abst) x0 x1) (pr3_head_12 c u2 x0 H8 (Bind Abst) t2 x1 (H9
Abst x0)))))))))) H6))) H4))))) H1))))))))) b).
-theorem pc3_gen_lift_abst:
+lemma pc3_gen_lift_abst:
\forall (c: C).(\forall (t: T).(\forall (t2: T).(\forall (u2: T).(\forall
(h: nat).(\forall (d: nat).((pc3 c (lift h d t) (THead (Bind Abst) u2 t2))
\to (\forall (e: C).((drop h d c e) \to (ex3_2 T T (\lambda (u1: T).(\lambda
t2 (lift h (S d) t1)))))) x3 x4 H17 H16 H15))))))))) (lift_gen_bind Abst x1
x2 x0 h d H11)))))))) H7))))) H4))))) H1)))))))))).
-theorem pc3_gen_sort_abst:
+lemma pc3_gen_sort_abst:
\forall (c: C).(\forall (u: T).(\forall (t: T).(\forall (n: nat).((pc3 c
(TSort n) (THead (Bind Abst) u t)) \to (\forall (P: Prop).P)))))
\def
include "basic_1/pc3/props.ma".
-let rec pc3_left_ind (c: C) (P: (T \to (T \to Prop))) (f: (\forall (t: T).(P
-t t))) (f0: (\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall
-(t3: T).((pc3_left c t2 t3) \to ((P t2 t3) \to (P t1 t3)))))))) (f1: (\forall
-(t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall (t3: T).((pc3_left c t1
-t3) \to ((P t1 t3) \to (P t2 t3)))))))) (t: T) (t0: T) (p: pc3_left c t t0)
-on p: P t t0 \def match p with [(pc3_left_r t1) \Rightarrow (f t1) |
-(pc3_left_ur t1 t2 p0 t3 p1) \Rightarrow (f0 t1 t2 p0 t3 p1 ((pc3_left_ind c
-P f f0 f1) t2 t3 p1)) | (pc3_left_ux t1 t2 p0 t3 p1) \Rightarrow (f1 t1 t2 p0
-t3 p1 ((pc3_left_ind c P f f0 f1) t1 t3 p1))].
+implied let rec pc3_left_ind (c: C) (P: (T \to (T \to Prop))) (f: (\forall
+(t: T).(P t t))) (f0: (\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to
+(\forall (t3: T).((pc3_left c t2 t3) \to ((P t2 t3) \to (P t1 t3)))))))) (f1:
+(\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall (t3:
+T).((pc3_left c t1 t3) \to ((P t1 t3) \to (P t2 t3)))))))) (t: T) (t0: T) (p:
+pc3_left c t t0) on p: P t t0 \def match p with [(pc3_left_r t1) \Rightarrow
+(f t1) | (pc3_left_ur t1 t2 p0 t3 p1) \Rightarrow (f0 t1 t2 p0 t3 p1
+((pc3_left_ind c P f f0 f1) t2 t3 p1)) | (pc3_left_ux t1 t2 p0 t3 p1)
+\Rightarrow (f1 t1 t2 p0 t3 p1 ((pc3_left_ind c P f f0 f1) t1 t3 p1))].
-theorem pc3_ind_left__pc3_left_pr3:
+fact pc3_ind_left__pc3_left_pr3:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr3 c t1 t2) \to
(pc3_left c t1 t2))))
\def
c t3 t0)).(\lambda (t4: T).(\lambda (_: (pr3 c t0 t4)).(\lambda (H2:
(pc3_left c t0 t4)).(pc3_left_ur c t3 t0 H0 t4 H2))))))) t1 t2 H)))).
-theorem pc3_ind_left__pc3_left_trans:
+fact pc3_ind_left__pc3_left_trans:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pc3_left c t1 t2) \to
(\forall (t3: T).((pc3_left c t2 t3) \to (pc3_left c t1 t3))))))
\def
t5))))).(\lambda (t5: T).(\lambda (H3: (pc3_left c t4 t5)).(pc3_left_ux c t0
t3 H0 t5 (H2 t5 H3)))))))))) t1 t2 H)))).
-theorem pc3_ind_left__pc3_left_sym:
+fact pc3_ind_left__pc3_left_sym:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pc3_left c t1 t2) \to
(pc3_left c t2 t1))))
\def
t0)).(pc3_ind_left__pc3_left_trans c t4 t0 H2 t3 (pc3_left_ur c t0 t3 H0 t3
(pc3_left_r c t3))))))))) t1 t2 H)))).
-theorem pc3_ind_left__pc3_left_pc3:
+fact pc3_ind_left__pc3_left_pc3:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pc3 c t1 t2) \to
(pc3_left c t1 t2))))
\def
(pc3_ind_left__pc3_left_pr3 c t1 x H1) t2 (pc3_ind_left__pc3_left_sym c t2 x
(pc3_ind_left__pc3_left_pr3 c t2 x H2)))))) H0))))).
-theorem pc3_ind_left__pc3_pc3_left:
+fact pc3_ind_left__pc3_pc3_left:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pc3_left c t1 t2) \to
(pc3 c t1 t2))))
\def
(_: (pc3_left c t0 t4)).(\lambda (H2: (pc3 c t0 t4)).(pc3_t t0 c t3
(pc3_pr2_x c t3 t0 H0) t4 H2))))))) t1 t2 H)))).
-theorem pc3_ind_left:
+lemma pc3_ind_left:
\forall (c: C).(\forall (P: ((T \to (T \to Prop)))).(((\forall (t: T).(P t
t))) \to (((\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall (t3:
T).((pc3 c t2 t3) \to ((P t2 t3) \to (P t1 t3)))))))) \to (((\forall (t1:
include "basic_1/nf2/pr3.ma".
-theorem pc3_nf2:
+lemma pc3_nf2:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pc3 c t1 t2) \to ((nf2 c
t1) \to ((nf2 c t2) \to (eq T t1 t2))))))
\def
H_y0) in (eq_ind_r T t1 (\lambda (t: T).(eq T t1 t)) (refl_equal T t1) t2
H_y0))))))))) H2))))))).
-theorem pc3_nf2_unfold:
+lemma pc3_nf2_unfold:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pc3 c t1 t2) \to ((nf2 c
t2) \to (pr3 c t1 t2)))))
\def
include "basic_1/pr3/pr1.ma".
-theorem pc3_pc1:
+lemma pc3_pc1:
\forall (t1: T).(\forall (t2: T).((pc1 t1 t2) \to (\forall (c: C).(pc3 c t1
t2))))
\def
include "basic_1/pr3/pr3.ma".
-theorem clear_pc3_trans:
+lemma clear_pc3_trans:
\forall (c2: C).(\forall (t1: T).(\forall (t2: T).((pc3 c2 t1 t2) \to
(\forall (c1: C).((clear c1 c2) \to (pc3 c1 t1 t2))))))
\def
t)) x (clear_pr3_trans c2 t1 x H2 c1 H0) (clear_pr3_trans c2 t2 x H3 c1
H0))))) H1))))))).
-theorem pc3_pr2_r:
+lemma pc3_pr2_r:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (pc3 c
t1 t2))))
\def
t2)).(ex_intro2 T (\lambda (t: T).(pr3 c t1 t)) (\lambda (t: T).(pr3 c t2 t))
t2 (pr3_pr2 c t1 t2 H) (pr3_refl c t2))))).
-theorem pc3_pr2_x:
+lemma pc3_pr2_x:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr2 c t2 t1) \to (pc3 c
t1 t2))))
\def
t1)).(ex_intro2 T (\lambda (t: T).(pr3 c t1 t)) (\lambda (t: T).(pr3 c t2 t))
t1 (pr3_refl c t1) (pr3_pr2 c t2 t1 H))))).
-theorem pc3_pr3_r:
+lemma pc3_pr3_r:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr3 c t1 t2) \to (pc3 c
t1 t2))))
\def
t2)).(ex_intro2 T (\lambda (t: T).(pr3 c t1 t)) (\lambda (t: T).(pr3 c t2 t))
t2 H (pr3_refl c t2))))).
-theorem pc3_pr3_x:
+lemma pc3_pr3_x:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr3 c t2 t1) \to (pc3 c
t1 t2))))
\def
t1)).(ex_intro2 T (\lambda (t: T).(pr3 c t1 t)) (\lambda (t: T).(pr3 c t2 t))
t1 (pr3_refl c t1) H)))).
-theorem pc3_pr3_t:
+lemma pc3_pr3_t:
\forall (c: C).(\forall (t1: T).(\forall (t0: T).((pr3 c t1 t0) \to (\forall
(t2: T).((pr3 c t2 t0) \to (pc3 c t1 t2))))))
\def
t0)).(\lambda (t2: T).(\lambda (H0: (pr3 c t2 t0)).(ex_intro2 T (\lambda (t:
T).(pr3 c t1 t)) (\lambda (t: T).(pr3 c t2 t)) t0 H H0)))))).
-theorem pc3_refl:
+lemma pc3_refl:
\forall (c: C).(\forall (t: T).(pc3 c t t))
\def
\lambda (c: C).(\lambda (t: T).(ex_intro2 T (\lambda (t0: T).(pr3 c t t0))
(\lambda (t0: T).(pr3 c t t0)) t (pr3_refl c t) (pr3_refl c t))).
-theorem pc3_s:
+lemma pc3_s:
\forall (c: C).(\forall (t2: T).(\forall (t1: T).((pc3 c t1 t2) \to (pc3 c
t2 t1))))
\def
x)).(\lambda (H2: (pr3 c t2 x)).(ex_intro2 T (\lambda (t: T).(pr3 c t2 t))
(\lambda (t: T).(pr3 c t1 t)) x H2 H1)))) H0))))).
-theorem pc3_thin_dx:
+lemma pc3_thin_dx:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pc3 c t1 t2) \to (\forall
(u: T).(\forall (f: F).(pc3 c (THead (Flat f) u t1) (THead (Flat f) u
t2)))))))
(Flat f) u x) (pr3_thin_dx c t1 x H1 u f) (pr3_thin_dx c t2 x H2 u f)))))
H0))))))).
-theorem pc3_head_1:
+lemma pc3_head_1:
\forall (c: C).(\forall (u1: T).(\forall (u2: T).((pc3 c u1 u2) \to (\forall
(k: K).(\forall (t: T).(pc3 c (THead k u1 t) (THead k u2 t)))))))
\def
H1 k t t (pr3_refl (CHead c k x) t)) (pr3_head_12 c u2 x H2 k t t (pr3_refl
(CHead c k x) t)))))) H0))))))).
-theorem pc3_head_2:
+lemma pc3_head_2:
\forall (c: C).(\forall (u: T).(\forall (t1: T).(\forall (t2: T).(\forall
(k: K).((pc3 (CHead c k u) t1 t2) \to (pc3 c (THead k u t1) (THead k u
t2)))))))
t2) t)) (THead k u x) (pr3_head_12 c u u (pr3_refl c u) k t1 x H1)
(pr3_head_12 c u u (pr3_refl c u) k t2 x H2))))) H0))))))).
-theorem pc3_pr2_u:
+lemma pc3_pr2_u:
\forall (c: C).(\forall (t2: T).(\forall (t1: T).((pr2 c t1 t2) \to (\forall
(t3: T).((pc3 c t2 t3) \to (pc3 c t1 t3))))))
\def
H5 x1 H7) t3 (pr3_t x t3 c H3 x1 H8))))) (pr3_confluence c t2 x0 H6 x H2)))))
H4))))) H1))))))).
-theorem pc3_pr2_u2:
+lemma pc3_pr2_u2:
\forall (c: C).(\forall (t0: T).(\forall (t1: T).((pr2 c t0 t1) \to (\forall
(t2: T).((pc3 c t0 t2) \to (pc3 c t1 t2))))))
\def
t1)).(\lambda (t2: T).(\lambda (H0: (pc3 c t0 t2)).(pc3_t t0 c t1 (pc3_pr2_x
c t1 t0 H) t2 H0)))))).
-theorem pc3_pr3_conf:
+lemma pc3_pr3_conf:
\forall (c: C).(\forall (t: T).(\forall (t1: T).((pc3 c t t1) \to (\forall
(t2: T).((pr3 c t t2) \to (pc3 c t2 t1))))))
\def
(CHead c k u1) t1 t2)).(pc3_t (THead k u1 t2) c (THead k u1 t1) (pc3_head_2 c
u1 t1 t2 k H0) (THead k u2 t2) (pc3_head_1 c u1 u2 H k t2))))))))).
-theorem pc3_pr0_pr2_t:
+lemma pc3_pr0_pr2_t:
\forall (u1: T).(\forall (u2: T).((pr0 u2 u1) \to (\forall (c: C).(\forall
(t1: T).(\forall (t2: T).(\forall (k: K).((pr2 (CHead c k u2) t1 t2) \to (pc3
(CHead c k u1) t1 t2))))))))
(r (Flat f) i0) H10 t3 t4 H3 t H9) f u1))))) k IHi (getl_gen_S k c (CHead d
(Bind Abbr) u) u2 i0 H8)))))) i H7 H4)))))))))))))) y t1 t2 H1))) H0)))))))).
-theorem pc3_pr2_pr2_t:
+lemma pc3_pr2_pr2_t:
\forall (c: C).(\forall (u1: T).(\forall (u2: T).((pr2 c u2 u1) \to (\forall
(t1: T).(\forall (t2: T).(\forall (k: K).((pr2 (CHead c k u2) t1 t2) \to (pc3
(CHead c k u1) t1 t2))))))))
(getl_gen_S k c0 (CHead d0 (Bind Abbr) u0) t1 i1 H10)))))) i0 H9
H7))))))))))))) y t0 t3 H4))) H3))))))))))))))) c u2 u1 H)))).
-theorem pc3_pr2_pr3_t:
+lemma pc3_pr2_pr3_t:
\forall (c: C).(\forall (u2: T).(\forall (t1: T).(\forall (t2: T).(\forall
(k: K).((pr3 (CHead c k u2) t1 t2) \to (\forall (u1: T).((pr2 c u2 u1) \to
(pc3 (CHead c k u1) t1 t2))))))))
u1)).(pc3_t t0 (CHead c k u1) t3 (pc3_pr2_pr2_t c u1 u2 H3 t3 t0 k H0) t4 (H2
u1 H3)))))))))) t1 t2 H)))))).
-theorem pc3_pr3_pc3_t:
+lemma pc3_pr3_pc3_t:
\forall (c: C).(\forall (u1: T).(\forall (u2: T).((pr3 c u2 u1) \to (\forall
(t1: T).(\forall (t2: T).(\forall (k: K).((pc3 (CHead c k u2) t1 t2) \to (pc3
(CHead c k u1) t1 t2))))))))
x k H5 t2 H0) t4 (pc3_s (CHead c k t2) x t4 (pc3_pr2_pr3_t c t1 t4 x k H6 t2
H0)))))) H4))))))))))))) u2 u1 H)))).
-theorem pc3_lift:
+lemma pc3_lift:
\forall (c: C).(\forall (e: C).(\forall (h: nat).(\forall (d: nat).((drop h
d c e) \to (\forall (t1: T).(\forall (t2: T).((pc3 e t1 t2) \to (pc3 c (lift
h d t1) (lift h d t2)))))))))
(lift h d x) (pr3_lift c e h d H t1 x H2) (lift h d t2) (pr3_lift c e h d H
t2 x H3))))) H1))))))))).
-theorem pc3_eta:
+lemma pc3_eta:
\forall (c: C).(\forall (t: T).(\forall (w: T).(\forall (u: T).((pc3 c t
(THead (Bind Abst) w u)) \to (\forall (v: T).((pc3 c v w) \to (pc3 c (THead
(Bind Abst) v (THead (Flat Appl) (TLRef O) (lift (S O) O t))) t)))))))
include "basic_1/pr3/subst1.ma".
-theorem pc3_gen_cabbr:
+lemma pc3_gen_cabbr:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pc3 c t1 t2) \to (\forall
(e: C).(\forall (u: T).(\forall (d: nat).((getl d c (CHead e (Bind Abbr) u))
\to (\forall (a0: C).((csubst1 d u c a0) \to (\forall (a: C).((drop (S O) d
include "basic_1/wcpr0/getl.ma".
-theorem pc3_wcpr0__pc3_wcpr0_t_aux:
+fact pc3_wcpr0__pc3_wcpr0_t_aux:
\forall (c1: C).(\forall (c2: C).((wcpr0 c1 c2) \to (\forall (k: K).(\forall
(u: T).(\forall (t1: T).(\forall (t2: T).((pr3 (CHead c1 k u) t1 t2) \to (pc3
(CHead c2 k u) t1 t2))))))))
u) (CHead c2 k u) (wcpr0_comp c1 c2 H u u (pr0_refl u) k) i d u0 (Bind Abbr)
H9)))))))))))))) y t4 t3 H4))) H1) t5 H3))))))) t1 t2 H0)))))))).
-theorem pc3_wcpr0_t:
+lemma pc3_wcpr0_t:
\forall (c1: C).(\forall (c2: C).((wcpr0 c1 c2) \to (\forall (t1:
T).(\forall (t2: T).((pr3 c1 t1 t2) \to (pc3 c2 t1 t2))))))
\def
(pc3_s (CHead c3 k u2) x t2 (pc3_wcpr0__pc3_wcpr0_t_aux c0 c3 H0 k u2 t2 x
H6)))))) H4))))))))))))) c1 c2 H))).
-theorem pc3_wcpr0:
+lemma pc3_wcpr0:
\forall (c1: C).(\forall (c2: C).((wcpr0 c1 c2) \to (\forall (t1:
T).(\forall (t2: T).((pc3 c1 t1 t2) \to (pc3 c2 t1 t2))))))
\def
include "basic_1/T/props.ma".
-theorem nf0_dec:
+lemma nf0_dec:
\forall (t1: T).(or (\forall (t2: T).((pr0 t1 t2) \to (eq T t1 t2))) (ex2 T
(\lambda (t2: T).((eq T t1 t2) \to (\forall (P: Prop).P))) (\lambda (t2:
T).(pr0 t1 t2))))
include "basic_1/subst0/fwd.ma".
-let rec pr0_ind (P: (T \to (T \to Prop))) (f: (\forall (t: T).(P t t))) (f0:
-(\forall (u1: T).(\forall (u2: T).((pr0 u1 u2) \to ((P u1 u2) \to (\forall
-(t1: T).(\forall (t2: T).((pr0 t1 t2) \to ((P t1 t2) \to (\forall (k: K).(P
-(THead k u1 t1) (THead k u2 t2)))))))))))) (f1: (\forall (u: T).(\forall (v1:
-T).(\forall (v2: T).((pr0 v1 v2) \to ((P v1 v2) \to (\forall (t1: T).(\forall
-(t2: T).((pr0 t1 t2) \to ((P t1 t2) \to (P (THead (Flat Appl) v1 (THead (Bind
-Abst) u t1)) (THead (Bind Abbr) v2 t2)))))))))))) (f2: (\forall (b: B).((not
-(eq B b Abst)) \to (\forall (v1: T).(\forall (v2: T).((pr0 v1 v2) \to ((P v1
-v2) \to (\forall (u1: T).(\forall (u2: T).((pr0 u1 u2) \to ((P u1 u2) \to
-(\forall (t1: T).(\forall (t2: T).((pr0 t1 t2) \to ((P t1 t2) \to (P (THead
-(Flat Appl) v1 (THead (Bind b) u1 t1)) (THead (Bind b) u2 (THead (Flat Appl)
-(lift (S O) O v2) t2)))))))))))))))))) (f3: (\forall (u1: T).(\forall (u2:
-T).((pr0 u1 u2) \to ((P u1 u2) \to (\forall (t1: T).(\forall (t2: T).((pr0 t1
-t2) \to ((P t1 t2) \to (\forall (w: T).((subst0 O u2 t2 w) \to (P (THead
-(Bind Abbr) u1 t1) (THead (Bind Abbr) u2 w))))))))))))) (f4: (\forall (b:
-B).((not (eq B b Abst)) \to (\forall (t1: T).(\forall (t2: T).((pr0 t1 t2)
-\to ((P t1 t2) \to (\forall (u: T).(P (THead (Bind b) u (lift (S O) O t1))
-t2))))))))) (f5: (\forall (t1: T).(\forall (t2: T).((pr0 t1 t2) \to ((P t1
-t2) \to (\forall (u: T).(P (THead (Flat Cast) u t1) t2))))))) (t: T) (t0: T)
-(p: pr0 t t0) on p: P t t0 \def match p with [(pr0_refl t1) \Rightarrow (f
-t1) | (pr0_comp u1 u2 p0 t1 t2 p1 k) \Rightarrow (f0 u1 u2 p0 ((pr0_ind P f
-f0 f1 f2 f3 f4 f5) u1 u2 p0) t1 t2 p1 ((pr0_ind P f f0 f1 f2 f3 f4 f5) t1 t2
-p1) k) | (pr0_beta u v1 v2 p0 t1 t2 p1) \Rightarrow (f1 u v1 v2 p0 ((pr0_ind
-P f f0 f1 f2 f3 f4 f5) v1 v2 p0) t1 t2 p1 ((pr0_ind P f f0 f1 f2 f3 f4 f5) t1
-t2 p1)) | (pr0_upsilon b n v1 v2 p0 u1 u2 p1 t1 t2 p2) \Rightarrow (f2 b n v1
-v2 p0 ((pr0_ind P f f0 f1 f2 f3 f4 f5) v1 v2 p0) u1 u2 p1 ((pr0_ind P f f0 f1
-f2 f3 f4 f5) u1 u2 p1) t1 t2 p2 ((pr0_ind P f f0 f1 f2 f3 f4 f5) t1 t2 p2)) |
-(pr0_delta u1 u2 p0 t1 t2 p1 w s0) \Rightarrow (f3 u1 u2 p0 ((pr0_ind P f f0
-f1 f2 f3 f4 f5) u1 u2 p0) t1 t2 p1 ((pr0_ind P f f0 f1 f2 f3 f4 f5) t1 t2 p1)
-w s0) | (pr0_zeta b n t1 t2 p0 u) \Rightarrow (f4 b n t1 t2 p0 ((pr0_ind P f
-f0 f1 f2 f3 f4 f5) t1 t2 p0) u) | (pr0_tau t1 t2 p0 u) \Rightarrow (f5 t1 t2
-p0 ((pr0_ind P f f0 f1 f2 f3 f4 f5) t1 t2 p0) u)].
+implied let rec pr0_ind (P: (T \to (T \to Prop))) (f: (\forall (t: T).(P t
+t))) (f0: (\forall (u1: T).(\forall (u2: T).((pr0 u1 u2) \to ((P u1 u2) \to
+(\forall (t1: T).(\forall (t2: T).((pr0 t1 t2) \to ((P t1 t2) \to (\forall
+(k: K).(P (THead k u1 t1) (THead k u2 t2)))))))))))) (f1: (\forall (u:
+T).(\forall (v1: T).(\forall (v2: T).((pr0 v1 v2) \to ((P v1 v2) \to (\forall
+(t1: T).(\forall (t2: T).((pr0 t1 t2) \to ((P t1 t2) \to (P (THead (Flat
+Appl) v1 (THead (Bind Abst) u t1)) (THead (Bind Abbr) v2 t2)))))))))))) (f2:
+(\forall (b: B).((not (eq B b Abst)) \to (\forall (v1: T).(\forall (v2:
+T).((pr0 v1 v2) \to ((P v1 v2) \to (\forall (u1: T).(\forall (u2: T).((pr0 u1
+u2) \to ((P u1 u2) \to (\forall (t1: T).(\forall (t2: T).((pr0 t1 t2) \to ((P
+t1 t2) \to (P (THead (Flat Appl) v1 (THead (Bind b) u1 t1)) (THead (Bind b)
+u2 (THead (Flat Appl) (lift (S O) O v2) t2)))))))))))))))))) (f3: (\forall
+(u1: T).(\forall (u2: T).((pr0 u1 u2) \to ((P u1 u2) \to (\forall (t1:
+T).(\forall (t2: T).((pr0 t1 t2) \to ((P t1 t2) \to (\forall (w: T).((subst0
+O u2 t2 w) \to (P (THead (Bind Abbr) u1 t1) (THead (Bind Abbr) u2
+w))))))))))))) (f4: (\forall (b: B).((not (eq B b Abst)) \to (\forall (t1:
+T).(\forall (t2: T).((pr0 t1 t2) \to ((P t1 t2) \to (\forall (u: T).(P (THead
+(Bind b) u (lift (S O) O t1)) t2))))))))) (f5: (\forall (t1: T).(\forall (t2:
+T).((pr0 t1 t2) \to ((P t1 t2) \to (\forall (u: T).(P (THead (Flat Cast) u
+t1) t2))))))) (t: T) (t0: T) (p: pr0 t t0) on p: P t t0 \def match p with
+[(pr0_refl t1) \Rightarrow (f t1) | (pr0_comp u1 u2 p0 t1 t2 p1 k)
+\Rightarrow (f0 u1 u2 p0 ((pr0_ind P f f0 f1 f2 f3 f4 f5) u1 u2 p0) t1 t2 p1
+((pr0_ind P f f0 f1 f2 f3 f4 f5) t1 t2 p1) k) | (pr0_beta u v1 v2 p0 t1 t2
+p1) \Rightarrow (f1 u v1 v2 p0 ((pr0_ind P f f0 f1 f2 f3 f4 f5) v1 v2 p0) t1
+t2 p1 ((pr0_ind P f f0 f1 f2 f3 f4 f5) t1 t2 p1)) | (pr0_upsilon b n v1 v2 p0
+u1 u2 p1 t1 t2 p2) \Rightarrow (f2 b n v1 v2 p0 ((pr0_ind P f f0 f1 f2 f3 f4
+f5) v1 v2 p0) u1 u2 p1 ((pr0_ind P f f0 f1 f2 f3 f4 f5) u1 u2 p1) t1 t2 p2
+((pr0_ind P f f0 f1 f2 f3 f4 f5) t1 t2 p2)) | (pr0_delta u1 u2 p0 t1 t2 p1 w
+s0) \Rightarrow (f3 u1 u2 p0 ((pr0_ind P f f0 f1 f2 f3 f4 f5) u1 u2 p0) t1 t2
+p1 ((pr0_ind P f f0 f1 f2 f3 f4 f5) t1 t2 p1) w s0) | (pr0_zeta b n t1 t2 p0
+u) \Rightarrow (f4 b n t1 t2 p0 ((pr0_ind P f f0 f1 f2 f3 f4 f5) t1 t2 p0) u)
+| (pr0_tau t1 t2 p0 u) \Rightarrow (f5 t1 t2 p0 ((pr0_ind P f f0 f1 f2 f3 f4
+f5) t1 t2 p0) u)].
-theorem pr0_gen_sort:
+lemma pr0_gen_sort:
\forall (x: T).(\forall (n: nat).((pr0 (TSort n) x) \to (eq T x (TSort n))))
\def
\lambda (x: T).(\lambda (n: nat).(\lambda (H: (pr0 (TSort n) x)).(insert_eq
True])) I (TSort n) H3) in (False_ind (eq T t2 (THead (Flat Cast) u t1))
H4)))))))) y x H0))) H))).
-theorem pr0_gen_lref:
+lemma pr0_gen_lref:
\forall (x: T).(\forall (n: nat).((pr0 (TLRef n) x) \to (eq T x (TLRef n))))
\def
\lambda (x: T).(\lambda (n: nat).(\lambda (H: (pr0 (TLRef n) x)).(insert_eq
True])) I (TLRef n) H3) in (False_ind (eq T t2 (THead (Flat Cast) u t1))
H4)))))))) y x H0))) H))).
-theorem pr0_gen_abst:
+lemma pr0_gen_abst:
\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr0 (THead (Bind Abst) u1
t1) x) \to (ex3_2 T T (\lambda (u2: T).(\lambda (t2: T).(eq T x (THead (Bind
Abst) u2 t2)))) (\lambda (u2: T).(\lambda (_: T).(pr0 u1 u2))) (\lambda (_:
T).(\lambda (_: T).(pr0 u1 u2))) (\lambda (_: T).(\lambda (t3: T).(pr0 t1
t3)))) H4)))))))) y x H0))) H)))).
-theorem pr0_gen_appl:
+lemma pr0_gen_appl:
\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr0 (THead (Flat Appl) u1
t1) x) \to (or3 (ex3_2 T T (\lambda (u2: T).(\lambda (t2: T).(eq T x (THead
(Flat Appl) u2 t2)))) (\lambda (u2: T).(\lambda (_: T).(pr0 u1 u2))) (\lambda
B).(\lambda (_: T).(\lambda (z1: T).(\lambda (_: T).(\lambda (_: T).(\lambda
(t3: T).(pr0 z1 t3))))))))) H4)))))))) y x H0))) H)))).
-theorem pr0_gen_cast:
+lemma pr0_gen_cast:
\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr0 (THead (Flat Cast) u1
t1) x) \to (or (ex3_2 T T (\lambda (u2: T).(\lambda (t2: T).(eq T x (THead
(Flat Cast) u2 t2)))) (\lambda (u2: T).(\lambda (_: T).(pr0 u1 u2))) (\lambda
T).(pr0 u1 u2))) (\lambda (_: T).(\lambda (t3: T).(pr0 t1 t3)))) (pr0 t1 t2)
H8))))) H4)))))))) y x H0))) H)))).
-theorem pr0_gen_lift:
+lemma pr0_gen_lift:
\forall (t1: T).(\forall (x: T).(\forall (h: nat).(\forall (d: nat).((pr0
(lift h d t1) x) \to (ex2 T (\lambda (t2: T).(eq T x (lift h d t2))) (\lambda
(t2: T).(pr0 t1 t2)))))))
include "basic_1/tlt/fwd.ma".
-theorem pr0_confluence__pr0_cong_upsilon_refl:
+fact pr0_confluence__pr0_cong_upsilon_refl:
\forall (b: B).((not (eq B b Abst)) \to (\forall (u0: T).(\forall (u3:
T).((pr0 u0 u3) \to (\forall (t4: T).(\forall (t5: T).((pr0 t4 t5) \to
(\forall (u2: T).(\forall (v2: T).(\forall (x: T).((pr0 u2 x) \to ((pr0 v2 x)
O) O x) (pr0_lift v2 x H3 (S O) O) t5 t5 (pr0_refl t5) (Flat Appl)) (Bind
b))))))))))))))).
-theorem pr0_confluence__pr0_cong_upsilon_cong:
+fact pr0_confluence__pr0_cong_upsilon_cong:
\forall (b: B).((not (eq B b Abst)) \to (\forall (u2: T).(\forall (v2:
T).(\forall (x: T).((pr0 u2 x) \to ((pr0 v2 x) \to (\forall (t2: T).(\forall
(t5: T).(\forall (x0: T).((pr0 t2 x0) \to ((pr0 t5 x0) \to (\forall (u5:
(lift (S O) O v2) (lift (S O) O x) (pr0_lift v2 x H1 (S O) O) t5 x0 H3 (Flat
Appl)) (Bind b))))))))))))))))))).
-theorem pr0_confluence__pr0_cong_upsilon_delta:
+fact pr0_confluence__pr0_cong_upsilon_delta:
(not (eq B Abbr Abst)) \to (\forall (u5: T).(\forall (t2: T).(\forall (w:
T).((subst0 O u5 t2 w) \to (\forall (u2: T).(\forall (v2: T).(\forall (x:
T).((pr0 u2 x) \to ((pr0 v2 x) \to (\forall (t5: T).(\forall (x0: T).((pr0 t2
(lift (S O) O x))))))) H7)) (pr0_subst0 t2 x0 H3 u5 w O H0 x1
H5))))))))))))))))))).
-theorem pr0_confluence__pr0_cong_upsilon_zeta:
+fact pr0_confluence__pr0_cong_upsilon_zeta:
\forall (b: B).((not (eq B b Abst)) \to (\forall (u0: T).(\forall (u3:
T).((pr0 u0 u3) \to (\forall (u2: T).(\forall (v2: T).(\forall (x0: T).((pr0
u2 x0) \to ((pr0 v2 x0) \to (\forall (x: T).(\forall (t3: T).(\forall (x1:
(lift (S O) O v2) (lift (S O) O x)) (lift_flat Appl v2 x (S O)
O)))))))))))))))).
-theorem pr0_confluence__pr0_cong_delta:
+fact pr0_confluence__pr0_cong_delta:
\forall (u3: T).(\forall (t5: T).(\forall (w: T).((subst0 O u3 t5 w) \to
(\forall (u2: T).(\forall (x: T).((pr0 u2 x) \to ((pr0 u3 x) \to (\forall
(t3: T).(\forall (x0: T).((pr0 t3 x0) \to ((pr0 t5 x0) \to (ex2 T (\lambda
u2 x H0 t3 x0 H2 x1 H6) (pr0_comp u3 x H1 w x1 H5 (Bind Abbr)))))) H4))
(pr0_subst0 t5 x0 H3 u3 w O H x H1))))))))))))).
-theorem pr0_confluence__pr0_upsilon_upsilon:
+fact pr0_confluence__pr0_upsilon_upsilon:
\forall (b: B).((not (eq B b Abst)) \to (\forall (v1: T).(\forall (v2:
T).(\forall (x0: T).((pr0 v1 x0) \to ((pr0 v2 x0) \to (\forall (u1:
T).(\forall (u2: T).(\forall (x1: T).((pr0 u1 x1) \to ((pr0 u2 x1) \to
x0) x2) (pr0_comp (lift (S O) O v2) (lift (S O) O x0) (pr0_lift v2 x0 H1 (S
O) O) t2 x2 H5 (Flat Appl)) (Bind b))))))))))))))))))).
-theorem pr0_confluence__pr0_delta_delta:
+fact pr0_confluence__pr0_delta_delta:
\forall (u2: T).(\forall (t3: T).(\forall (w: T).((subst0 O u2 t3 w) \to
(\forall (u3: T).(\forall (t5: T).(\forall (w0: T).((subst0 O u3 t5 w0) \to
(\forall (x: T).((pr0 u2 x) \to ((pr0 u3 x) \to (\forall (x0: T).((pr0 t3 x0)
x2 x O H10 x1 H7))))) H8)) (pr0_subst0 t3 x0 H3 u2 w O H x H1))))) H5))
(pr0_subst0 t5 x0 H4 u3 w0 O H0 x H2))))))))))))))).
-theorem pr0_confluence__pr0_delta_tau:
+fact pr0_confluence__pr0_delta_tau:
\forall (u2: T).(\forall (t3: T).(\forall (w: T).((subst0 O u2 t3 w) \to
(\forall (t4: T).((pr0 (lift (S O) O t4) t3) \to (\forall (t2: T).(ex2 T
(\lambda (t: T).(pr0 (THead (Bind Abbr) u2 w) t)) (\lambda (t: T).(pr0 t2
include "basic_1/subst0/props.ma".
-theorem pr0_lift:
+lemma pr0_lift:
\forall (t1: T).(\forall (t2: T).((pr0 t1 t2) \to (\forall (h: nat).(\forall
(d: nat).(pr0 (lift h d t1) (lift h d t2))))))
\def
(lift h d (THead (Flat Cast) u t3)) (lift_head (Flat Cast) u t3 h d)))))))))
t1 t2 H))).
-theorem pr0_gen_abbr:
+lemma pr0_gen_abbr:
\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr0 (THead (Bind Abbr) u1
t1) x) \to (or (ex3_2 T T (\lambda (u2: T).(\lambda (t2: T).(eq T x (THead
(Bind Abbr) u2 t2)))) (\lambda (u2: T).(\lambda (_: T).(pr0 u1 u2))) (\lambda
(\lambda (y0: T).(subst0 O u2 y0 t3))))))) (pr0 t1 (lift (S O) O t2)))
H4)))))))) y x H0))) H)))).
-theorem pr0_gen_void:
+lemma pr0_gen_void:
\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr0 (THead (Bind Void) u1
t1) x) \to (or (ex3_2 T T (\lambda (u2: T).(\lambda (t2: T).(eq T x (THead
(Bind Void) u2 t2)))) (\lambda (u2: T).(\lambda (_: T).(pr0 u1 u2))) (\lambda
include "basic_1/subst0/subst0.ma".
-theorem pr0_subst0_back:
+lemma pr0_subst0_back:
\forall (u2: T).(\forall (t1: T).(\forall (t2: T).(\forall (i: nat).((subst0
i u2 t1 t2) \to (\forall (u1: T).((pr0 u1 u2) \to (ex2 T (\lambda (t:
T).(subst0 i u1 t1 t)) (\lambda (t: T).(pr0 t t2)))))))))
t4))) (THead k x0 x) (subst0_both u0 u1 x0 i0 H7 k t3 x H5) (pr0_comp x0 u3
H8 x t4 H6 k))))) (H1 u0 H4))))) (H3 u0 H4))))))))))))))) i u2 t1 t2 H))))).
-theorem pr0_subst0_fwd:
+lemma pr0_subst0_fwd:
\forall (u2: T).(\forall (t1: T).(\forall (t2: T).(\forall (i: nat).((subst0
i u2 t1 t2) \to (\forall (u1: T).((pr0 u2 u1) \to (ex2 T (\lambda (t:
T).(subst0 i u1 t1 t)) (\lambda (t: T).(pr0 t2 t)))))))))
include "basic_1/subst1/fwd.ma".
-theorem pr0_delta1:
+lemma pr0_delta1:
\forall (u1: T).(\forall (u2: T).((pr0 u1 u2) \to (\forall (t1: T).(\forall
(t2: T).((pr0 t1 t2) \to (\forall (w: T).((subst1 O u2 t2 w) \to (pr0 (THead
(Bind Abbr) u1 t1) (THead (Bind Abbr) u2 w)))))))))
Abbr)) (\lambda (t0: T).(\lambda (H2: (subst0 O u2 t2 t0)).(pr0_delta u1 u2 H
t1 t2 H0 t0 H2))) w H1)))))))).
-theorem pr0_subst1_back:
+lemma pr0_subst1_back:
\forall (u2: T).(\forall (t1: T).(\forall (t2: T).(\forall (i: nat).((subst1
i u2 t1 t2) \to (\forall (u1: T).((pr0 u1 u2) \to (ex2 T (\lambda (t:
T).(subst1 i u1 t1 t)) (\lambda (t: T).(pr0 t t2)))))))))
T).(subst1 i u1 t1 t)) (\lambda (t: T).(pr0 t t0)) x (subst1_single i u1 t1 x
H2) H3)))) (pr0_subst0_back u2 t1 t0 i H0 u1 H1)))))) t2 H))))).
-theorem pr0_subst1_fwd:
+lemma pr0_subst1_fwd:
\forall (u2: T).(\forall (t1: T).(\forall (t2: T).(\forall (i: nat).((subst1
i u2 t1 t2) \to (\forall (u1: T).((pr0 u2 u1) \to (ex2 T (\lambda (t:
T).(subst1 i u1 t1 t)) (\lambda (t: T).(pr0 t2 t)))))))))
include "basic_1/pr1/defs.ma".
-let rec pr1_ind (P: (T \to (T \to Prop))) (f: (\forall (t: T).(P t t))) (f0:
-(\forall (t2: T).(\forall (t1: T).((pr0 t1 t2) \to (\forall (t3: T).((pr1 t2
-t3) \to ((P t2 t3) \to (P t1 t3)))))))) (t: T) (t0: T) (p: pr1 t t0) on p: P
-t t0 \def match p with [(pr1_refl t1) \Rightarrow (f t1) | (pr1_sing t2 t1 p0
-t3 p1) \Rightarrow (f0 t2 t1 p0 t3 p1 ((pr1_ind P f f0) t2 t3 p1))].
+implied let rec pr1_ind (P: (T \to (T \to Prop))) (f: (\forall (t: T).(P t
+t))) (f0: (\forall (t2: T).(\forall (t1: T).((pr0 t1 t2) \to (\forall (t3:
+T).((pr1 t2 t3) \to ((P t2 t3) \to (P t1 t3)))))))) (t: T) (t0: T) (p: pr1 t
+t0) on p: P t t0 \def match p with [(pr1_refl t1) \Rightarrow (f t1) |
+(pr1_sing t2 t1 p0 t3 p1) \Rightarrow (f0 t2 t1 p0 t3 p1 ((pr1_ind P f f0) t2
+t3 p1))].
include "basic_1/pr0/pr0.ma".
-theorem pr1_strip:
+lemma pr1_strip:
\forall (t0: T).(\forall (t1: T).((pr1 t0 t1) \to (\forall (t2: T).((pr0 t0
t2) \to (ex2 T (\lambda (t: T).(pr1 t1 t)) (\lambda (t: T).(pr1 t2 t)))))))
\def
include "basic_1/T/props.ma".
-theorem pr1_pr0:
+lemma pr1_pr0:
\forall (t1: T).(\forall (t2: T).((pr0 t1 t2) \to (pr1 t1 t2)))
\def
\lambda (t1: T).(\lambda (t2: T).(\lambda (H: (pr0 t1 t2)).(pr1_sing t2 t1 H
t5))))).(\lambda (t5: T).(\lambda (H3: (pr1 t4 t5)).(pr1_sing t0 t3 H0 t5 (H2
t5 H3)))))))))) t1 t2 H))).
-theorem pr1_head_1:
+lemma pr1_head_1:
\forall (u1: T).(\forall (u2: T).((pr1 u1 u2) \to (\forall (t: T).(\forall
(k: K).(pr1 (THead k u1 t) (THead k u2 t))))))
\def
(THead k t2 t) (THead k t1 t) (pr0_comp t1 t2 H0 t t (pr0_refl t) k) (THead k
t3 t) H2))))))) u1 u2 H))))).
-theorem pr1_head_2:
+lemma pr1_head_2:
\forall (t1: T).(\forall (t2: T).((pr1 t1 t2) \to (\forall (u: T).(\forall
(k: K).(pr1 (THead k u t1) (THead k u t2))))))
\def
t1 t1 (pr0_refl t1) t4 t0 H4 k) (THead k t3 t5) H6))))))) t u H3))))))))))) v
w H))).
-theorem pr1_eta:
+lemma pr1_eta:
\forall (w: T).(\forall (u: T).(let t \def (THead (Bind Abst) w u) in
(\forall (v: T).((pr1 v w) \to (pr1 (THead (Bind Abst) v (THead (Flat Appl)
(TLRef O) (lift (S O) O t))) t)))))
include "basic_1/clen/getl.ma".
-theorem pr2_gen_ctail:
+lemma pr2_gen_ctail:
\forall (k: K).(\forall (c: C).(\forall (u: T).(\forall (t1: T).(\forall
(t2: T).((pr2 (CTail k u c) t1 t2) \to (or (pr2 c t1 t2) (ex3 T (\lambda (_:
T).(eq K k (Bind Abbr))) (\lambda (t: T).(pr0 t1 t)) (\lambda (t: T).(subst0
(refl_equal K (Bind Abbr)) H2 H13)) k H9)))))))) H7)) H6))))))))))))))) y t1
t2 H0))) H)))))).
-theorem pr2_gen_cbind:
+lemma pr2_gen_cbind:
\forall (b: B).(\forall (c: C).(\forall (v: T).(\forall (t1: T).(\forall
(t2: T).((pr2 (CHead c (Bind b) v) t1 t2) \to (pr2 c (THead (Bind b) v t1)
(THead (Bind b) v t2)))))))
Abbr) u) x H9) t3 t4 H2 t H11))))))) H7)) H6))))))))))))))) y t1 t2 H0)))
H)))))).
-theorem pr2_gen_cflat:
+lemma pr2_gen_cflat:
\forall (f: F).(\forall (c: C).(\forall (v: T).(\forall (t1: T).(\forall
(t2: T).((pr2 (CHead c (Flat f) v) t1 t2) \to (pr2 c t1 t2))))))
\def
include "basic_1/getl/drop.ma".
-theorem pr2_ind:
+implied lemma pr2_ind:
\forall (P: ((C \to (T \to (T \to Prop))))).(((\forall (c: C).(\forall (t1:
T).(\forall (t2: T).((pr0 t1 t2) \to (P c t1 t2)))))) \to (((\forall (c:
C).(\forall (d: C).(\forall (u: T).(\forall (i: nat).((getl i c (CHead d
\Rightarrow (f x x0 x1 x2) | (pr2_delta x x0 x1 x2 x3 x4 x5 x6 x7 x8)
\Rightarrow (f0 x x0 x1 x2 x3 x4 x5 x6 x7 x8)]))))))).
-theorem pr2_gen_sort:
+lemma pr2_gen_sort:
\forall (c: C).(\forall (x: T).(\forall (n: nat).((pr2 c (TSort n) x) \to
(eq T x (TSort n)))))
\def
(pr0_gen_sort t2 n H5)) in (subst0_gen_sort u t i n H6 (eq T t (TSort n))))
t1 H4))))))))))))) c y x H0))) H)))).
-theorem pr2_gen_lref:
+lemma pr2_gen_lref:
\forall (c: C).(\forall (x: T).(\forall (n: nat).((pr2 c (TLRef n) x) \to
(or (eq T x (TLRef n)) (ex2_2 C T (\lambda (d: C).(\lambda (u: T).(getl n c
(CHead d (Bind Abbr) u)))) (\lambda (_: C).(\lambda (u: T).(eq T x (lift (S
n) O u0)))) d u H9 (refl_equal T (lift (S n) O u))))) t H8)))
(subst0_gen_lref u t i n H6))) t1 H4))))))))))))) c y x H0))) H)))).
-theorem pr2_gen_abst:
+lemma pr2_gen_abst:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr2 c
(THead (Bind Abst) u1 t1) x) \to (ex3_2 T T (\lambda (u2: T).(\lambda (t2:
T).(eq T x (THead (Bind Abst) u2 t2)))) (\lambda (u2: T).(\lambda (_: T).(pr2
H13)))) t H11)))))) H10)) (subst0_gen_head (Bind Abst) u x0 x1 t i H9))))))))
(pr0_gen_abst u1 t1 t2 H5)))))))))))))) c y x H0))) H))))).
-theorem pr2_gen_cast:
+lemma pr2_gen_cast:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr2 c
(THead (Flat Cast) u1 t1) x) \to (or (ex3_2 T T (\lambda (u2: T).(\lambda
(t2: T).(eq T x (THead (Flat Cast) u2 t2)))) (\lambda (u2: T).(\lambda (_:
(pr2_delta c0 d u i H1 t1 t2 H6 t H3))) (pr0_gen_cast u1 t1 t2
H5)))))))))))))) c y x H0))) H))))).
-theorem pr2_gen_csort:
+lemma pr2_gen_csort:
\forall (t1: T).(\forall (t2: T).(\forall (n: nat).((pr2 (CSort n) t1 t2)
\to (pr0 t1 t2))))
\def
(Bind Abbr) u))) H1 (CSort n) H4) in (getl_gen_sort n i (CHead d (Bind Abbr)
u) H5 (pr0 t3 t)))))))))))))) y t1 t2 H0))) H)))).
-theorem pr2_gen_appl:
+lemma pr2_gen_appl:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr2 c
(THead (Flat Appl) u1 t1) x) \to (or3 (ex3_2 T T (\lambda (u2: T).(\lambda
(t2: T).(eq T x (THead (Flat Appl) u2 t2)))) (\lambda (u2: T).(\lambda (_:
(Flat Appl) (lift (S O) O x3) x5) t i H13)) t1 H8)))))))))))))) H6))
(pr0_gen_appl u1 t1 t2 H5)))))))))))))) c y x H0))) H))))).
-theorem pr2_gen_lift:
+lemma pr2_gen_lift:
\forall (c: C).(\forall (t1: T).(\forall (x: T).(\forall (h: nat).(\forall
(d: nat).((pr2 c (lift h d t1) x) \to (\forall (e: C).((drop h d c e) \to
(ex2 T (\lambda (t2: T).(eq T x (lift h d t2))) (\lambda (t2: T).(pr2 e t1
include "basic_1/getl/fwd.ma".
-theorem pr2_confluence__pr2_free_free:
+fact pr2_confluence__pr2_free_free:
\forall (c: C).(\forall (t0: T).(\forall (t1: T).(\forall (t2: T).((pr0 t0
t1) \to ((pr0 t0 t2) \to (ex2 T (\lambda (t: T).(pr2 c t1 t)) (\lambda (t:
T).(pr2 c t2 t))))))))
(\lambda (t: T).(pr2 c t2 t)) x (pr2_free c t1 x H2) (pr2_free c t2 x H1)))))
(pr0_confluence t0 t2 H0 t1 H))))))).
-theorem pr2_confluence__pr2_free_delta:
+fact pr2_confluence__pr2_free_delta:
\forall (c: C).(\forall (d: C).(\forall (t0: T).(\forall (t1: T).(\forall
(t2: T).(\forall (t4: T).(\forall (u: T).(\forall (i: nat).((pr0 t0 t1) \to
((getl i c (CHead d (Bind Abbr) u)) \to ((pr0 t0 t4) \to ((subst0 i u t4 t2)
H6))))) H5)) (pr0_subst0 t4 x H3 u t2 i H2 u (pr0_refl u))))))
(pr0_confluence t0 t4 H1 t1 H))))))))))))).
-theorem pr2_confluence__pr2_delta_delta:
+fact pr2_confluence__pr2_delta_delta:
\forall (c: C).(\forall (d: C).(\forall (d0: C).(\forall (t0: T).(\forall
(t1: T).(\forall (t2: T).(\forall (t3: T).(\forall (t4: T).(\forall (u:
T).(\forall (u0: T).(\forall (i: nat).(\forall (i0: nat).((getl i c (CHead d
include "basic_1/pr0/subst0.ma".
-theorem pr2_thin_dx:
+lemma pr2_thin_dx:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall
(u: T).(\forall (f: F).(pr2 c (THead (Flat f) u t1) (THead (Flat f) u
t2)))))))
t3 H1 (Flat f)) (THead (Flat f) u t) (subst0_snd (Flat f) u0 t t3 i H2
u)))))))))))) c t1 t2 H)))))).
-theorem pr2_head_1:
+lemma pr2_head_1:
\forall (c: C).(\forall (u1: T).(\forall (u2: T).((pr2 c u1 u2) \to (\forall
(k: K).(\forall (t: T).(pr2 c (THead k u1 t) (THead k u2 t)))))))
\def
t t (pr0_refl t) k) (THead k t0 t) (subst0_fst u t0 t2 i H2 t k)))))))))))) c
u1 u2 H)))))).
-theorem pr2_head_2:
+lemma pr2_head_2:
\forall (c: C).(\forall (u: T).(\forall (t1: T).(\forall (t2: T).(\forall
(k: K).((pr2 (CHead c k u) t1 t2) \to (pr2 c (THead k u t1) (THead k u
t2)))))))
H3 (Flat f)) (THead (Flat f) u t) (subst0_snd (Flat f) u0 t t4 (r (Flat f) n)
H4 u))))))))))))) i)))))) k) y t1 t2 H0))) H)))))).
-theorem clear_pr2_trans:
+lemma clear_pr2_trans:
\forall (c2: C).(\forall (t1: T).(\forall (t2: T).((pr2 c2 t1 t2) \to
(\forall (c1: C).((clear c1 c2) \to (pr2 c1 t1 t2))))))
\def
C).(\lambda (H3: (clear c1 c)).(pr2_delta c1 d u i (clear_getl_trans i c
(CHead d (Bind Abbr) u) H0 c1 H3) t3 t4 H1 t H2))))))))))))) c2 t1 t2 H)))).
-theorem pr2_cflat:
+lemma pr2_cflat:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall
(f: F).(\forall (v: T).(pr2 (CHead c (Flat f) v) t1 t2))))))
\def
i (getl_flat c0 (CHead d (Bind Abbr) u) i H0 f v) t3 t4 H1 t H2))))))))))) c
t1 t2 H)))))).
-theorem pr2_ctail:
+lemma pr2_ctail:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall
(k: K).(\forall (u: T).(pr2 (CTail k u c) t1 t2))))))
\def
(subst0 i u0 t4 t)).(pr2_delta (CTail k u c0) (CTail k u d) u0 i (getl_ctail
Abbr c0 d u0 i H0 k u) t3 t4 H1 t H2))))))))))) c t1 t2 H)))))).
-theorem pr2_change:
+lemma pr2_change:
\forall (b: B).((not (eq B b Abbr)) \to (\forall (c: C).(\forall (v1:
T).(\forall (t1: T).(\forall (t2: T).((pr2 (CHead c (Bind b) v1) t1 t2) \to
(\forall (v2: T).(pr2 (CHead c (Bind b) v2) t1 t2))))))))
(CHead d (Bind Abbr) u) v1 i0 H7) v2) t3 t4 H3 t H8))))) i H6 H4)))))))))))))
y t1 t2 H1))) H0)))))))).
-theorem pr2_lift:
+lemma pr2_lift:
\forall (c: C).(\forall (e: C).(\forall (h: nat).(\forall (d: nat).((drop h
d c e) \to (\forall (t1: T).(\forall (t2: T).((pr2 e t1 t2) \to (pr2 c (lift
h d t1) (lift h d t2)))))))))
(lift h d t4) (pr0_lift t3 t4 H3 h d) (lift h d t) (subst0_lift_ge t4 t u i h
H4 d H7)))))))))))))))) y t1 t2 H1))) H0)))))))).
-theorem pr2_gen_abbr:
+lemma pr2_gen_abbr:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr2 c
(THead (Bind Abbr) u1 t1) x) \to (or (ex3_2 T T (\lambda (u2: T).(\lambda
(t2: T).(eq T x (THead (Bind Abbr) u2 t2)))) (\lambda (u2: T).(\lambda (_:
H6 (lift (S O) O t) (subst0_lift_ge_S t2 t u i H3 O (le_O_n i)))))))
(pr0_gen_abbr u1 t1 t2 H5)))))))))))))) c y x H0))) H))))).
-theorem pr2_gen_void:
+lemma pr2_gen_void:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr2 c
(THead (Bind Void) u1 t1) x) \to (or (ex3_2 T T (\lambda (u2: T).(\lambda
(t2: T).(eq T x (THead (Bind Void) u2 t2)))) (\lambda (u2: T).(\lambda (_:
include "basic_1/subst1/subst1.ma".
-theorem pr2_delta1:
+lemma pr2_delta1:
\forall (c: C).(\forall (d: C).(\forall (u: T).(\forall (i: nat).((getl i c
(CHead d (Bind Abbr) u)) \to (\forall (t1: T).(\forall (t2: T).((pr0 t1 t2)
\to (\forall (t: T).((subst1 i u t2 t) \to (pr2 c t1 t))))))))))
(\lambda (t0: T).(\lambda (H2: (subst0 i u t2 t0)).(pr2_delta c d u i H t1 t2
H0 t0 H2))) t H1)))))))))).
-theorem pr2_subst1:
+lemma pr2_subst1:
\forall (c: C).(\forall (e: C).(\forall (v: T).(\forall (i: nat).((getl i c
(CHead e (Bind Abbr) v)) \to (\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2)
\to (\forall (w1: T).((subst1 i v t1 w1) \to (ex2 T (\lambda (w2: T).(pr2 c
H14)))))))))) (pr0_subst1 t3 t4 H3 v w1 i H6 v (pr0_refl v))) c0
H5))))))))))))))) y t1 t2 H1))) H0)))))))).
-theorem pr2_gen_cabbr:
+lemma pr2_gen_cabbr:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall
(e: C).(\forall (u: T).(\forall (d: nat).((getl d c (CHead e (Bind Abbr) u))
\to (\forall (a0: C).((csubst1 d u c a0) \to (\forall (a: C).((drop (S O) d
include "basic_1/pr2/fwd.ma".
-let rec pr3_ind (c: C) (P: (T \to (T \to Prop))) (f: (\forall (t: T).(P t
-t))) (f0: (\forall (t2: T).(\forall (t1: T).((pr2 c t1 t2) \to (\forall (t3:
-T).((pr3 c t2 t3) \to ((P t2 t3) \to (P t1 t3)))))))) (t: T) (t0: T) (p: pr3
-c t t0) on p: P t t0 \def match p with [(pr3_refl t1) \Rightarrow (f t1) |
-(pr3_sing t2 t1 p0 t3 p1) \Rightarrow (f0 t2 t1 p0 t3 p1 ((pr3_ind c P f f0)
-t2 t3 p1))].
+implied let rec pr3_ind (c: C) (P: (T \to (T \to Prop))) (f: (\forall (t:
+T).(P t t))) (f0: (\forall (t2: T).(\forall (t1: T).((pr2 c t1 t2) \to
+(\forall (t3: T).((pr3 c t2 t3) \to ((P t2 t3) \to (P t1 t3)))))))) (t: T)
+(t0: T) (p: pr3 c t t0) on p: P t t0 \def match p with [(pr3_refl t1)
+\Rightarrow (f t1) | (pr3_sing t2 t1 p0 t3 p1) \Rightarrow (f0 t2 t1 p0 t3 p1
+((pr3_ind c P f f0) t2 t3 p1))].
-theorem pr3_gen_sort:
+lemma pr3_gen_sort:
\forall (c: C).(\forall (x: T).(\forall (n: nat).((pr3 c (TSort n) x) \to
(eq T x (TSort n)))))
\def
\to (eq T t3 t))) H3 (TSort n) (pr2_gen_sort c t2 n H5)) in (H6 (refl_equal T
(TSort n)))) t1 H4))))))))) y x H0))) H)))).
-theorem pr3_gen_abst:
+lemma pr3_gen_abst:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr3 c
(THead (Bind Abst) u1 t1) x) \to (ex3_2 T T (\lambda (u2: T).(\lambda (t2:
T).(eq T x (THead (Bind Abst) u2 t2)))) (\lambda (u2: T).(\lambda (_: T).(pr3
T).(pr3_sing (CHead c (Bind b) u) x3 x1 (H9 b u) x5 (H14 b u))))))))))
H11)))))))) H6)))))))))))) y x H0))))) H))))).
-theorem pr3_gen_cast:
+lemma pr3_gen_cast:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr3 c
(THead (Flat Cast) u1 t1) x) \to (or (ex3_2 T T (\lambda (u2: T).(\lambda
(t2: T).(eq T x (THead (Flat Cast) u2 t2)))) (\lambda (u2: T).(\lambda (_:
T).(\lambda (t5: T).(pr3 c x1 t5)))) (pr3 c x1 t4) (pr3_sing c t2 x1 H7 t4
H2))) H6)))))))))))) y x H0))))) H))))).
-theorem pr3_gen_lift:
+lemma pr3_gen_lift:
\forall (c: C).(\forall (t1: T).(\forall (x: T).(\forall (h: nat).(\forall
(d: nat).((pr3 c (lift h d t1) x) \to (\forall (e: C).((drop h d c e) \to
(ex2 T (\lambda (t2: T).(eq T x (lift h d t2))) (\lambda (t2: T).(pr3 e t1
H10 (pr3_sing e x1 x0 H9 x2 H11))))) (H3 x1 H8 e H5))))) H7))))))))))))) y x
H0)))) H)))))).
-theorem pr3_gen_lref:
+lemma pr3_gen_lref:
\forall (c: C).(\forall (x: T).(\forall (n: nat).((pr3 c (TLRef n) x) \to
(or (eq T x (TLRef n)) (ex3_3 C T T (\lambda (d: C).(\lambda (u: T).(\lambda
(_: T).(getl n c (CHead d (Bind Abbr) u))))) (\lambda (d: C).(\lambda (u:
include "basic_1/tlist/fwd.ma".
-theorem pr3_iso_appls_abbr:
+lemma pr3_iso_appls_abbr:
\forall (c: C).(\forall (d: C).(\forall (w: T).(\forall (i: nat).((getl i c
(CHead d (Bind Abbr) w)) \to (\forall (vs: TList).(let u1 \def (THeads (Flat
Appl) vs (TLRef i)) in (\forall (u2: T).((pr3 c u1 u2) \to ((((iso u1 u2) \to
Appl) (lift (S O) O x4) x3) (pr3_thin_dx (CHead c (Bind x0) x5) x2 x3 H10
(lift (S O) O x4) Appl)) u2 H7)))))))))))))) H4)) H3)))))))) vs)))))).
-theorem pr3_iso_appls_cast:
+lemma pr3_iso_appls_cast:
\forall (c: C).(\forall (v: T).(\forall (t: T).(\forall (vs: TList).(let u1
\def (THeads (Flat Appl) vs (THead (Flat Cast) v t)) in (\forall (u2:
T).((pr3 c u1 u2) \to ((((iso u1 u2) \to (\forall (P: Prop).P))) \to (pr3 c
(THead (Flat Appl) (lift (S O) O x4) x3) (pr3_thin_dx (CHead c (Bind x0) x5)
x2 x3 H9 (lift (S O) O x4) Appl)) u2 H6)))))))))))))) H3)) H2)))))))) vs)))).
-theorem pr3_iso_appl_bind:
+lemma pr3_iso_appl_bind:
\forall (b: B).((not (eq B b Abst)) \to (\forall (v1: T).(\forall (v2:
T).(\forall (t: T).(let u1 \def (THead (Flat Appl) v1 (THead (Bind b) v2 t))
in (\forall (c: C).(\forall (u2: T).((pr3 c u1 u2) \to ((((iso u1 u2) \to
(Bind x0) x1 x2))) (lift_flat Appl x4 (THead (Bind x0) x1 x2) (S O) O))))
H10))) u2 H6))))))))))))) H3)) H2)))))))))).
-theorem pr3_iso_appls_appl_bind:
+lemma pr3_iso_appls_appl_bind:
\forall (b: B).((not (eq B b Abst)) \to (\forall (v: T).(\forall (u:
T).(\forall (t: T).(\forall (vs: TList).(let u1 \def (THeads (Flat Appl) vs
(THead (Flat Appl) v (THead (Bind b) u t))) in (\forall (c: C).(\forall (u2:
(pr3_thin_dx (CHead c (Bind x0) x5) x2 x3 H10 (lift (S O) O x4) Appl)) u2
H7)))))))))))))) H4)) H3))))))))) vs)))))).
-theorem pr3_iso_appls_bind:
+lemma pr3_iso_appls_bind:
\forall (b: B).((not (eq B b Abst)) \to (\forall (vs: TList).(\forall (u:
T).(\forall (t: T).(let u1 \def (THeads (Flat Appl) vs (THead (Bind b) u t))
in (\forall (c: C).(\forall (u2: T).((pr3 c u1 u2) \to ((((iso u1 u2) \to
Appl) (lift (S O) O t) t0 (lifts (S O) O ts))) (lifts (S O) O (TApp ts t))
(lifts_tapp (S O) O t ts))))))))))) vs))).
-theorem pr3_iso_beta:
+lemma pr3_iso_beta:
\forall (v: T).(\forall (w: T).(\forall (t: T).(let u1 \def (THead (Flat
Appl) v (THead (Bind Abst) w t)) in (\forall (c: C).(\forall (u2: T).((pr3 c
u1 u2) \to ((((iso u1 u2) \to (\forall (P: Prop).P))) \to (pr3 c (THead (Bind
False with []) in H23))))))))) H14)) H13))))))) H9)))))))))))))) H2))
H1)))))))).
-theorem pr3_iso_appls_beta:
+lemma pr3_iso_appls_beta:
\forall (us: TList).(\forall (v: T).(\forall (w: T).(\forall (t: T).(let u1
\def (THeads (Flat Appl) us (THead (Flat Appl) v (THead (Bind Abst) w t))) in
(\forall (c: C).(\forall (u2: T).((pr3 c u1 u2) \to ((((iso u1 u2) \to
include "basic_1/pr1/fwd.ma".
-theorem pr3_pr1:
+lemma pr3_pr1:
\forall (t1: T).(\forall (t2: T).((pr1 t1 t2) \to (\forall (c: C).(pr3 c t1
t2))))
\def
include "basic_1/pr2/pr2.ma".
-theorem pr3_strip:
+lemma pr3_strip:
\forall (c: C).(\forall (t0: T).(\forall (t1: T).((pr3 c t0 t1) \to (\forall
(t2: T).((pr2 c t0 t2) \to (ex2 T (\lambda (t: T).(pr3 c t1 t)) (\lambda (t:
T).(pr3 c t2 t))))))))
include "basic_1/pr1/props.ma".
-theorem clear_pr3_trans:
+lemma clear_pr3_trans:
\forall (c2: C).(\forall (t1: T).(\forall (t2: T).((pr3 c2 t1 t2) \to
(\forall (c1: C).((clear c1 c2) \to (pr3 c1 t1 t2))))))
\def
T).(\lambda (_: (pr3 c2 t3 t5)).(\lambda (H3: (pr3 c1 t3 t5)).(pr3_sing c1 t3
t4 (clear_pr2_trans c2 t4 t3 H1 c1 H0) t5 H3))))))) t1 t2 H)))))).
-theorem pr3_pr2:
+lemma pr3_pr2:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (pr3 c
t1 t2))))
\def
(t5: T).((pr3 c t4 t5) \to (pr3 c t0 t5))))).(\lambda (t5: T).(\lambda (H3:
(pr3 c t4 t5)).(pr3_sing c t0 t3 H0 t5 (H2 t5 H3)))))))))) t1 t2 H)))).
-theorem pr3_thin_dx:
+lemma pr3_thin_dx:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr3 c t1 t2) \to (\forall
(u: T).(\forall (f: F).(pr3 c (THead (Flat f) u t1) (THead (Flat f) u
t2)))))))
t4))).(pr3_sing c (THead (Flat f) u t0) (THead (Flat f) u t3) (pr2_thin_dx c
t3 t0 H0 u f) (THead (Flat f) u t4) H2))))))) t1 t2 H)))))).
-theorem pr3_head_1:
+lemma pr3_head_1:
\forall (c: C).(\forall (u1: T).(\forall (u2: T).((pr3 c u1 u2) \to (\forall
(k: K).(\forall (t: T).(pr3 c (THead k u1 t) (THead k u2 t)))))))
\def
c (THead k t2 t) (THead k t1 t) (pr2_head_1 c t1 t2 H0 k t) (THead k t3 t)
(H2 k t)))))))))) u1 u2 H)))).
-theorem pr3_head_2:
+lemma pr3_head_2:
\forall (c: C).(\forall (u: T).(\forall (t1: T).(\forall (t2: T).(\forall
(k: K).((pr3 (CHead c k u) t1 t2) \to (pr3 c (THead k u t1) (THead k u
t2)))))))
(CHead c k u2) t1 t2)).(pr3_t (THead k u2 t1) (THead k u1 t1) c (pr3_head_1 c
u1 u2 H k t1) (THead k u2 t2) (pr3_head_2 c u2 t1 t2 k H0))))))))).
-theorem pr3_cflat:
+lemma pr3_cflat:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr3 c t1 t2) \to (\forall
(f: F).(\forall (v: T).(pr3 (CHead c (Flat f) v) t1 t2))))))
\def
(f: F).(pr3_head_12 c u1 u2 H (Flat f) t1 t2 (pr3_cflat c t1 t2 H0 f
u2))))))))).
-theorem pr3_pr0_pr2_t:
+lemma pr3_pr0_pr2_t:
\forall (u1: T).(\forall (u2: T).((pr0 u1 u2) \to (\forall (c: C).(\forall
(t1: T).(\forall (t2: T).(\forall (k: K).((pr2 (CHead c k u2) t1 t2) \to (pr3
(CHead c k u1) t1 t2))))))))
(getl_gen_S (Flat f) c (CHead d (Bind Abbr) u) u2 i0 H9) t3 t4 H3 t H8) f
u1))))) k H7 IHi))))) i H6 H4))))))))))))) y t1 t2 H1))) H0)))))))).
-theorem pr3_pr2_pr2_t:
+lemma pr3_pr2_pr2_t:
\forall (c: C).(\forall (u1: T).(\forall (u2: T).((pr2 c u1 u2) \to (\forall
(t1: T).(\forall (t2: T).(\forall (k: K).((pr2 (CHead c k u2) t1 t2) \to (pr3
(CHead c k u1) t1 t2))))))))
f t1)))) k H10))))) i0 H9 H7))))))))))))) y t0 t3 H4))) H3))))))))))))))) c
u1 u2 H)))).
-theorem pr3_pr2_pr3_t:
+lemma pr3_pr2_pr3_t:
\forall (c: C).(\forall (u2: T).(\forall (t1: T).(\forall (t2: T).(\forall
(k: K).((pr3 (CHead c k u2) t1 t2) \to (\forall (u1: T).((pr2 c u1 u2) \to
(pr3 (CHead c k u1) t1 t2))))))))
T).(\lambda (t4: T).(\lambda (k: K).(\lambda (H3: (pr3 (CHead c k t3) t0
t4)).(pr3_pr2_pr3_t c t2 t0 t4 k (H2 t0 t4 k H3) t1 H0))))))))))) u1 u2 H)))).
-theorem pr3_lift:
+lemma pr3_lift:
\forall (c: C).(\forall (e: C).(\forall (h: nat).(\forall (d: nat).((drop h
d c e) \to (\forall (t1: T).(\forall (t2: T).((pr3 e t1 t2) \to (pr3 c (lift
h d t1) (lift h d t2)))))))))
t0) (lift h d t3) (pr2_lift c e h d H t3 t0 H1) (lift h d t4) H3))))))) t1 t2
H0)))))))).
-theorem pr3_eta:
+lemma pr3_eta:
\forall (c: C).(\forall (w: T).(\forall (u: T).(let t \def (THead (Bind
Abst) w u) in (\forall (v: T).((pr3 c v w) \to (pr3 c (THead (Bind Abst) v
(THead (Flat Appl) (TLRef O) (lift (S O) O t))) t))))))
(pr0_refl u) (TLRef O))))) (CHead c (Bind Abst) w))) (lift (S O) O (THead
(Bind Abst) w u)) (lift_bind Abst w u (S O) O))))))).
-theorem pr3_gen_void:
+lemma pr3_gen_void:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr3 c
(THead (Bind Void) u1 t1) x) \to (or (ex3_2 T T (\lambda (u2: T).(\lambda
(t2: T).(eq T x (THead (Bind Void) u2 t2)))) (\lambda (u2: T).(\lambda (_:
O) O t4) (pr3_lift (CHead c (Bind Void) x0) c (S O) O (drop_drop (Bind Void)
O c c (drop_refl c) x0) t2 t4 H2)))) H6)))))))))))) y x H0))))) H))))).
-theorem pr3_gen_abbr:
+lemma pr3_gen_abbr:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr3 c
(THead (Bind Abbr) u1 t1) x) \to (or (ex3_2 T T (\lambda (u2: T).(\lambda
(t2: T).(eq T x (THead (Bind Abbr) u2 t2)))) (\lambda (u2: T).(\lambda (_:
(pr3_lift (CHead c (Bind Abbr) x0) c (S O) O (drop_drop (Bind Abbr) O c c
(drop_refl c) x0) t2 t4 H2)))) H6)))))))))))) y x H0))))) H))))).
-theorem pr3_gen_appl:
+lemma pr3_gen_appl:
\forall (c: C).(\forall (u1: T).(\forall (t1: T).(\forall (x: T).((pr3 c
(THead (Flat Appl) u1 t1) x) \to (or3 (ex3_2 T T (\lambda (u2: T).(\lambda
(t2: T).(eq T x (THead (Flat Appl) u2 t2)))) (\lambda (u2: T).(\lambda (_:
H15 (pr3_pr2 c x0 x6 H11) (pr3_pr2 c x3 x7 H12) (pr3_pr2 (CHead c (Bind x2)
x7) x4 x5 H13))))) x1 H9))))))))))))) H7)) H6)))))))))))) y x H0))))) H))))).
-theorem pr3_gen_bind:
+lemma pr3_gen_bind:
\forall (b: B).((not (eq B b Abst)) \to (\forall (c: C).(\forall (u1:
T).(\forall (t1: T).(\forall (x: T).((pr3 c (THead (Bind b) u1 t1) x) \to (or
(ex3_2 T T (\lambda (u2: T).(\lambda (t2: T).(eq T x (THead (Bind b) u2
include "basic_1/pr2/subst1.ma".
-theorem pr3_subst1:
+lemma pr3_subst1:
\forall (c: C).(\forall (e: C).(\forall (v: T).(\forall (i: nat).((getl i c
(CHead e (Bind Abbr) v)) \to (\forall (t1: T).(\forall (t2: T).((pr3 c t1 t2)
\to (\forall (w1: T).((subst1 i v t1 w1) \to (ex2 T (\lambda (w2: T).(pr3 c
i v t5 w2)) x0 (pr3_sing c x w1 H5 x0 H7) H8)))) (H3 x H6))))) (pr2_subst1 c
e v i H t4 t3 H1 w1 H4)))))))))) t1 t2 H0)))))))).
-theorem pr3_gen_cabbr:
+lemma pr3_gen_cabbr:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr3 c t1 t2) \to (\forall
(e: C).(\forall (u: T).(\forall (d: nat).((getl d c (CHead e (Bind Abbr) u))
\to (\forall (a0: C).((csubst1 d u c a0) \to (\forall (a: C).((drop (S O) d
include "basic_1/wcpr0/getl.ma".
-theorem pr3_wcpr0_t:
+lemma pr3_wcpr0_t:
\forall (c1: C).(\forall (c2: C).((wcpr0 c2 c1) \to (\forall (t1:
T).(\forall (t2: T).((pr3 c1 t1 t2) \to (pr3 c2 t1 t2))))))
\def
include "basic_1/s/defs.ma".
-theorem r_S:
+lemma r_S:
\forall (k: K).(\forall (i: nat).(eq nat (r k (S i)) (S (r k i))))
\def
\lambda (k: K).(K_ind (\lambda (k0: K).(\forall (i: nat).(eq nat (r k0 (S
(Bind b) i))))) (\lambda (f: F).(\lambda (i: nat).(refl_equal nat (S (r (Flat
f) i))))) k).
-theorem r_plus:
+lemma r_plus:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).(eq nat (r k (plus i j))
(plus (r k i) j))))
\def
(\lambda (f: F).(\lambda (i: nat).(\lambda (j: nat).(refl_equal nat (plus (r
(Flat f) i) j))))) k).
-theorem r_plus_sym:
+lemma r_plus_sym:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).(eq nat (r k (plus i j))
(plus i (r k j)))))
\def
(i: nat).(\lambda (j: nat).(refl_equal nat (plus i j))))) (\lambda (_:
F).(\lambda (i: nat).(\lambda (j: nat).(plus_n_Sm i j)))) k).
-theorem r_minus:
+lemma r_minus:
\forall (i: nat).(\forall (n: nat).((lt n i) \to (\forall (k: K).(eq nat
(minus (r k i) (S n)) (r k (minus i (S n)))))))
\def
n))))) (\lambda (_: B).(refl_equal nat (minus i (S n)))) (\lambda (_:
F).(minus_x_Sy i n H)) k)))).
-theorem r_dis:
+lemma r_dis:
\forall (k: K).(\forall (P: Prop).(((((\forall (i: nat).(eq nat (r k i) i)))
\to P)) \to (((((\forall (i: nat).(eq nat (r k i) (S i)))) \to P)) \to P)))
\def
((((\forall (i: nat).(eq nat (r (Flat f) i) (S i)))) \to P))).(H0 (\lambda
(i: nat).(refl_equal nat (S i)))))))) k).
-theorem s_r:
+lemma s_r:
\forall (k: K).(\forall (i: nat).(eq nat (s k (r k i)) (S i)))
\def
\lambda (k: K).(K_ind (\lambda (k0: K).(\forall (i: nat).(eq nat (s k0 (r k0
i)) (S i)))) (\lambda (_: B).(\lambda (i: nat).(refl_equal nat (S i))))
(\lambda (_: F).(\lambda (i: nat).(refl_equal nat (S i)))) k).
-theorem r_arith0:
+lemma r_arith0:
\forall (k: K).(\forall (i: nat).(eq nat (minus (r k (S i)) (S O)) (r k i)))
\def
\lambda (k: K).(\lambda (i: nat).(eq_ind_r nat (S (r k i)) (\lambda (n:
nat).(eq nat n (r k i))) (refl_equal nat (r k i)) (minus (S (r k i)) (S O))
(minus_Sx_SO (r k i))) (r k (S i)) (r_S k i))).
-theorem r_arith1:
+lemma r_arith1:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).(eq nat (minus (r k (S
i)) (S j)) (minus (r k i) j))))
\def
(\lambda (n: nat).(eq nat (minus n (S j)) (minus (r k i) j))) (refl_equal nat
(minus (r k i) j)) (r k (S i)) (r_S k i)))).
-theorem r_arith2:
+lemma r_arith2:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((le (S i) (s k j)) \to
(le (r k i) j))))
\def
(le_S_n i j H) in H_y))))) (\lambda (_: F).(\lambda (i: nat).(\lambda (j:
nat).(\lambda (H: (le (S i) j)).H)))) k).
-theorem r_arith3:
+lemma r_arith3:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((le (s k j) (S i)) \to
(le j (r k i)))))
\def
(le_S_n j i H) in H_y))))) (\lambda (_: F).(\lambda (i: nat).(\lambda (j:
nat).(\lambda (H: (le j (S i))).H)))) k).
-theorem r_arith4:
+lemma r_arith4:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).(eq nat (minus (S i) (s k
j)) (minus (r k i) j))))
\def
j))))) (\lambda (f: F).(\lambda (i: nat).(\lambda (j: nat).(refl_equal nat
(minus (r (Flat f) i) j))))) k).
-theorem r_arith5:
+lemma r_arith5:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((lt (s k j) (S i)) \to
(lt j (r k i)))))
\def
(\lambda (_: F).(\lambda (i: nat).(\lambda (j: nat).(\lambda (H: (lt j (S
i))).H)))) k).
-theorem r_arith6:
+lemma r_arith6:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).(eq nat (minus (r k i) (S
j)) (minus i (s k j)))))
\def
j)))))) (\lambda (f: F).(\lambda (i: nat).(\lambda (j: nat).(refl_equal nat
(minus i (s (Flat f) j)))))) k).
-theorem r_arith7:
+lemma r_arith7:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((eq nat (S i) (s k j))
\to (eq nat (r k i) j))))
\def
include "basic_1/s/defs.ma".
-theorem s_inj:
+lemma s_inj:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((eq nat (s k i) (s k j))
\to (eq nat i j))))
\def
(Bind b) j))).(eq_add_S i j H))))) (\lambda (f: F).(\lambda (i: nat).(\lambda
(j: nat).(\lambda (H: (eq nat (s (Flat f) i) (s (Flat f) j))).H)))) k).
-theorem s_le_gen:
+lemma s_le_gen:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((le (s k i) (s k j)) \to
(le i j))))
\def
j))).(le_S_n i j H))))) (\lambda (f: F).(\lambda (i: nat).(\lambda (j:
nat).(\lambda (H: (le (s (Flat f) i) (s (Flat f) j))).H)))) k).
-theorem s_lt_gen:
+lemma s_lt_gen:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((lt (s k i) (s k j)) \to
(lt i j))))
\def
include "basic_1/s/defs.ma".
-theorem s_S:
+lemma s_S:
\forall (k: K).(\forall (i: nat).(eq nat (s k (S i)) (S (s k i))))
\def
\lambda (k: K).(K_ind (\lambda (k0: K).(\forall (i: nat).(eq nat (s k0 (S
(Bind b) i))))) (\lambda (f: F).(\lambda (i: nat).(refl_equal nat (S (s (Flat
f) i))))) k).
-theorem s_plus:
+lemma s_plus:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).(eq nat (s k (plus i j))
(plus (s k i) j))))
\def
(\lambda (f: F).(\lambda (i: nat).(\lambda (j: nat).(refl_equal nat (plus (s
(Flat f) i) j))))) k).
-theorem s_plus_sym:
+lemma s_plus_sym:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).(eq nat (s k (plus i j))
(plus i (s k j)))))
\def
(plus_n_Sm i j))))) (\lambda (f: F).(\lambda (i: nat).(\lambda (j:
nat).(refl_equal nat (plus i (s (Flat f) j)))))) k).
-theorem s_minus:
+lemma s_minus:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((le j i) \to (eq nat (s
k (minus i j)) (minus (s k i) j)))))
\def
(\lambda (f: F).(\lambda (i: nat).(\lambda (j: nat).(\lambda (_: (le j
i)).(refl_equal nat (minus (s (Flat f) i) j)))))) k).
-theorem minus_s_s:
+lemma minus_s_s:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).(eq nat (minus (s k i) (s
k j)) (minus i j))))
\def
(\lambda (_: F).(\lambda (i: nat).(\lambda (j: nat).(refl_equal nat (minus i
j))))) k).
-theorem s_le:
+lemma s_le:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((le i j) \to (le (s k i)
(s k j)))))
\def
nat).(\lambda (j: nat).(\lambda (H: (le i j)).(le_n_S i j H))))) (\lambda (_:
F).(\lambda (i: nat).(\lambda (j: nat).(\lambda (H: (le i j)).H)))) k).
-theorem s_lt:
+lemma s_lt:
\forall (k: K).(\forall (i: nat).(\forall (j: nat).((lt i j) \to (lt (s k i)
(s k j)))))
\def
nat).(\lambda (j: nat).(\lambda (H: (lt i j)).(lt_n_S i j H))))) (\lambda (_:
F).(\lambda (i: nat).(\lambda (j: nat).(\lambda (H: (lt i j)).H)))) k).
-theorem s_inc:
+lemma s_inc:
\forall (k: K).(\forall (i: nat).(le i (s k i)))
\def
\lambda (k: K).(K_ind (\lambda (k0: K).(\forall (i: nat).(le i (s k0 i))))
(S (S (s (Bind b) i))) (le_n (S (S (s (Bind b) i)))))))))) (\lambda (f:
F).(\lambda (i: nat).(le_n (s (Flat f) i)))) k).
-theorem s_arith0:
+lemma s_arith0:
\forall (k: K).(\forall (i: nat).(eq nat (minus (s k i) (s k O)) i))
\def
\lambda (k: K).(\lambda (i: nat).(eq_ind_r nat (minus i O) (\lambda (n:
nat).(eq nat n i)) (eq_ind nat i (\lambda (n: nat).(eq nat n i)) (refl_equal
nat i) (minus i O) (minus_n_O i)) (minus (s k i) (s k O)) (minus_s_s k i O))).
-theorem s_arith1:
+lemma s_arith1:
\forall (b: B).(\forall (i: nat).(eq nat (minus (s (Bind b) i) (S O)) i))
\def
\lambda (_: B).(\lambda (i: nat).(eq_ind nat i (\lambda (n: nat).(eq nat n
include "basic_1/csubc/props.ma".
-theorem sc3_arity_csubc:
+lemma sc3_arity_csubc:
\forall (g: G).(\forall (c1: C).(\forall (t: T).(\forall (a: A).((arity g c1
t a) \to (\forall (d1: C).(\forall (is: PList).((drop1 is d1 c1) \to (\forall
(c2: C).((csubc g d1 c2) \to (sc3 g a c2 (lift1 is t)))))))))))
c)).(\lambda (c2: C).(\lambda (H4: (csubc g d1 c2)).(sc3_repl g a1 c2 (lift1
is t0) (H1 d1 is H3 c2 H4) a2 H2))))))))))))) c1 t a H))))).
-theorem sc3_arity:
+lemma sc3_arity:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (a: A).((arity g c t
a) \to (sc3 g a c t)))))
\def
include "basic_1/lift1/drop1.ma".
-theorem sc3_arity_gen:
+lemma sc3_arity_gen:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (a: A).((sc3 g a c
t) \to (arity g c t a)))))
\def
a0 d w) \to (\forall (is: PList).((drop1 is d c) \to (sc3 g a1 d (THead (Flat
Appl) w (lift1 is t)))))))))).H3)) H2))))))) a)))).
-theorem sc3_repl:
+lemma sc3_repl:
\forall (g: G).(\forall (a1: A).(\forall (c: C).(\forall (t: T).((sc3 g a1 c
t) \to (\forall (a2: A).((leq g a1 a2) \to (sc3 g a2 c t)))))))
\def
(llt_head_sx a a0)) d w H12 a (leq_sym g a x0 H8)) is H13) x1 H9))))))) a3
H11))))))) H7))))) H4)))))))))))) a2)) a1)).
-theorem sc3_lift:
+lemma sc3_lift:
\forall (g: G).(\forall (a: A).(\forall (e: C).(\forall (t: T).((sc3 g a e
t) \to (\forall (c: C).(\forall (h: nat).(\forall (d: nat).((drop h d c e)
\to (sc3 g a c (lift h d t))))))))))
t0))) (H_y (drop1_cons_tail c e h d H2 is d0 H7)) (lift1 is (lift h d t))
(lift1_cons_tail t h d is))))))))))) H3))))))))))))) a)).
-theorem sc3_lift1:
+lemma sc3_lift1:
\forall (g: G).(\forall (e: C).(\forall (a: A).(\forall (hds:
PList).(\forall (c: C).(\forall (t: T).((sc3 g a e t) \to ((drop1 hds c e)
\to (sc3 g a c (lift1 hds t)))))))))
n0 c x)).(\lambda (H4: (drop1 p x e)).(sc3_lift g a x (lift1 p t) (H x t H0
H4) c n n0 H3)))) H2))))))))))) hds)))).
-theorem sc3_abbr:
+lemma sc3_abbr:
\forall (g: G).(\forall (a: A).(\forall (vs: TList).(\forall (i:
nat).(\forall (d: C).(\forall (v: T).(\forall (c: C).((sc3 g a c (THeads
(Flat Appl) vs (lift (S i) O v))) \to ((getl i c (CHead d (Bind Abbr) v)) \to
(lift1 is (THeads (Flat Appl) vs (THead (Flat Cast) u t))) (lifts1_flat Appl
is (THead (Flat Cast) u t) vs))))))))))) H6)))) H3)))))))))))) a)).
-theorem sc3_props__sc3_sn3_abst:
+fact sc3_props__sc3_sn3_abst:
\forall (g: G).(\forall (a: A).(land (\forall (c: C).(\forall (t: T).((sc3 g
a c t) \to (sn3 c t)))) (\forall (vs: TList).(\forall (i: nat).(let t \def
(THeads (Flat Appl) vs (TLRef i)) in (\forall (c: C).((arity g c t a) \to
(lift1_lref is i)) (lift1 is (THeads (Flat Appl) vs (TLRef i))) (lifts1_flat
Appl is (TLRef i) vs))))) H9)))) H6))))))))))))))))))) a)).
-theorem sc3_sn3:
+lemma sc3_sn3:
\forall (g: G).(\forall (a: A).(\forall (c: C).(\forall (t: T).((sc3 g a c
t) \to (sn3 c t)))))
\def
\to ((sns3 c0 vs) \to (sc3 g a c0 (THeads (Flat Appl) vs (TLRef
i))))))))))).(H1 c t H))) H0))))))).
-theorem sc3_abst:
+lemma sc3_abst:
\forall (g: G).(\forall (a: A).(\forall (vs: TList).(\forall (c: C).(\forall
(i: nat).((arity g c (THeads (Flat Appl) vs (TLRef i)) a) \to ((nf2 c (TLRef
i)) \to ((sns3 c vs) \to (sc3 g a c (THeads (Flat Appl) vs (TLRef i))))))))))
include "basic_1/pr3/props.ma".
-let rec sn3_ind (c: C) (P: (T \to Prop)) (f: (\forall (t1: T).(((\forall (t2:
-T).((((eq T t1 t2) \to (\forall (P0: Prop).P0))) \to ((pr3 c t1 t2) \to (sn3
-c t2))))) \to (((\forall (t2: T).((((eq T t1 t2) \to (\forall (P0:
-Prop).P0))) \to ((pr3 c t1 t2) \to (P t2))))) \to (P t1))))) (t: T) (s0: sn3
-c t) on s0: P t \def match s0 with [(sn3_sing t1 s1) \Rightarrow (f t1 s1
-(\lambda (t2: T).(\lambda (p: (((eq T t1 t2) \to (\forall (P0:
-Prop).P0)))).(\lambda (p0: (pr3 c t1 t2)).((sn3_ind c P f) t2 (s1 t2 p
-p0))))))].
+implied let rec sn3_ind (c: C) (P: (T \to Prop)) (f: (\forall (t1:
+T).(((\forall (t2: T).((((eq T t1 t2) \to (\forall (P0: Prop).P0))) \to ((pr3
+c t1 t2) \to (sn3 c t2))))) \to (((\forall (t2: T).((((eq T t1 t2) \to
+(\forall (P0: Prop).P0))) \to ((pr3 c t1 t2) \to (P t2))))) \to (P t1)))))
+(t: T) (s0: sn3 c t) on s0: P t \def match s0 with [(sn3_sing t1 s1)
+\Rightarrow (f t1 s1 (\lambda (t2: T).(\lambda (p: (((eq T t1 t2) \to
+(\forall (P0: Prop).P0)))).(\lambda (p0: (pr3 c t1 t2)).((sn3_ind c P f) t2
+(s1 t2 p p0))))))].
-theorem sn3_gen_bind:
+lemma sn3_gen_bind:
\forall (b: B).(\forall (c: C).(\forall (u: T).(\forall (t: T).((sn3 c
(THead (Bind b) u t)) \to (land (sn3 c u) (sn3 (CHead c (Bind b) u) t))))))
\def
x) t2) (\lambda (_: (sn3 c x)).(\lambda (H10: (sn3 (CHead c (Bind b) x)
t2)).H10)) H8))))))))))))))) y H0))))) H))))).
-theorem sn3_gen_flat:
+lemma sn3_gen_flat:
\forall (f: F).(\forall (c: C).(\forall (u: T).(\forall (t: T).((sn3 c
(THead (Flat f) u t)) \to (land (sn3 c u) (sn3 c t))))))
\def
(sn3 c t2) (sn3 c t2) (\lambda (_: (sn3 c x)).(\lambda (H10: (sn3 c
t2)).H10)) H8))))))))))))))) y H0))))) H))))).
-theorem sn3_gen_head:
+lemma sn3_gen_head:
\forall (k: K).(\forall (c: C).(\forall (u: T).(\forall (t: T).((sn3 c
(THead k u t)) \to (sn3 c u)))))
\def
(land_ind (sn3 c u) (sn3 c t) (sn3 c u) (\lambda (H1: (sn3 c u)).(\lambda (_:
(sn3 c t)).H1)) H0)))))))) k).
-theorem sn3_gen_cflat:
+lemma sn3_gen_cflat:
\forall (f: F).(\forall (c: C).(\forall (u: T).(\forall (t: T).((sn3 (CHead
c (Flat f) u) t) \to (sn3 c t)))))
\def
\to (\forall (P: Prop).P)))).(\lambda (H3: (pr3 c t1 t2)).(H1 t2 H2
(pr3_cflat c t1 t2 H3 f u))))))))) t H))))).
-theorem sn3_gen_lift:
+lemma sn3_gen_lift:
\forall (c1: C).(\forall (t: T).(\forall (h: nat).(\forall (d: nat).((sn3 c1
(lift h d t)) \to (\forall (c2: C).((drop h d c1 c2) \to (sn3 c2 t)))))))
\def
include "basic_1/lift1/props.ma".
-theorem sns3_lifts1:
+lemma sns3_lifts1:
\forall (e: C).(\forall (hds: PList).(\forall (c: C).((drop1 hds c e) \to
(\forall (ts: TList).((sns3 e ts) \to (sns3 c (lifts1 hds ts)))))))
\def
include "basic_1/nf2/pr3.ma".
-theorem sn3_nf2:
+lemma sn3_nf2:
\forall (c: C).(\forall (t: T).((nf2 c t) \to (sn3 c t)))
\def
\lambda (c: C).(\lambda (t: T).(\lambda (H: (nf2 c t)).(sn3_sing c t
Prop).P))) H0 t H_y) in (eq_ind T t (\lambda (t0: T).(sn3 c t0)) (H3
(refl_equal T t) (sn3 c t)) t2 H_y)))))))))).
-theorem nf2_sn3:
+lemma nf2_sn3:
\forall (c: C).(\forall (t: T).((sn3 c t) \to (ex2 T (\lambda (u: T).(pr3 c
t u)) (\lambda (u: T).(nf2 c u)))))
\def
include "basic_1/pr3/iso.ma".
-theorem sn3_pr3_trans:
+lemma sn3_pr3_trans:
\forall (c: C).(\forall (t1: T).((sn3 c t1) \to (\forall (t2: T).((pr3 c t1
t2) \to (sn3 c t2)))))
\def
H6) in (H0 t0 H8 H7))))) (\lambda (H6: (((eq T t2 t3) \to (\forall (P:
Prop).P)))).(H1 t3 H6 H2 t0 H4)) H5)))))))))))) t1 H))).
-theorem sn3_pr2_intro:
+lemma sn3_pr2_intro:
\forall (c: C).(\forall (t1: T).(((\forall (t2: T).((((eq T t1 t2) \to
(\forall (P: Prop).P))) \to ((pr2 c t1 t2) \to (sn3 c t2))))) \to (sn3 c t1)))
\def
t2)).(sn3_pr3_trans c t0 (sn3_sing c t0 H3) t2 (pr3_pr2 c t0 t2 H8)))
H7))))))))) t H2)))))) u H))).
-theorem sn3_cflat:
+lemma sn3_cflat:
\forall (c: C).(\forall (t: T).((sn3 c t) \to (\forall (f: F).(\forall (u:
T).(sn3 (CHead c (Flat f) u) t)))))
\def
Prop).P)))).(\lambda (H3: (pr2 (CHead c (Flat f) u) t1 t2)).(H1 t2 H2
(pr3_pr2 c t1 t2 (pr2_gen_cflat f c u t1 t2 H3)))))))))) t H))))).
-theorem sn3_shift:
+lemma sn3_shift:
\forall (b: B).(\forall (c: C).(\forall (v: T).(\forall (t: T).((sn3 c
(THead (Bind b) v t)) \to (sn3 (CHead c (Bind b) v) t)))))
\def
(Bind b) v) t) (\lambda (_: (sn3 c v)).(\lambda (H2: (sn3 (CHead c (Bind b)
v) t)).H2)) H0))))))).
-theorem sn3_change:
+lemma sn3_change:
\forall (b: B).((not (eq B b Abbr)) \to (\forall (c: C).(\forall (v1:
T).(\forall (t: T).((sn3 (CHead c (Bind b) v1) t) \to (\forall (v2: T).(sn3
(CHead c (Bind b) v2) t)))))))
(pr3_pr2 (CHead c (Bind b) v1) t1 t2 (pr2_change b H c v2 t1 t2 H4
v1)))))))))) t H0))))))).
-theorem sn3_gen_def:
+lemma sn3_gen_def:
\forall (c: C).(\forall (d: C).(\forall (v: T).(\forall (i: nat).((getl i c
(CHead d (Bind Abbr) v)) \to ((sn3 c (TLRef i)) \to (sn3 d v))))))
\def
i) (pr0_refl (TLRef i)) (lift (S i) O v) (subst0_lref v i)))) d (getl_drop
Abbr c d v i H))))))).
-theorem sn3_cdelta:
+lemma sn3_cdelta:
\forall (v: T).(\forall (t: T).(\forall (i: nat).(((\forall (w: T).(ex T
(\lambda (u: T).(subst0 i w t u))))) \to (\forall (c: C).(\forall (d:
C).((getl i c (CHead d (Bind Abbr) v)) \to ((sn3 c t) \to (sn3 d v))))))))
(sn3_gen_head k c u1 t1 H7) in (H3 c d H6 H_y))))))))))))))))) i v t x H1)))
H0)))))).
-theorem sn3_cpr3_trans:
+lemma sn3_cpr3_trans:
\forall (c: C).(\forall (u1: T).(\forall (u2: T).((pr3 c u1 u2) \to (\forall
(k: K).(\forall (t: T).((sn3 (CHead c k u1) t) \to (sn3 (CHead c k u2)
t)))))))
False with []) in H29) x1 H24)))))))) H21)) H20)) t3 H15)))))))))))))) H12))
H11))))))))) w H4))))))))))) y H0))))) H)))).
-theorem sn3_appl_lref:
+lemma sn3_appl_lref:
\forall (c: C).(\forall (i: nat).((nf2 c (TLRef i)) \to (\forall (v:
T).((sn3 c v) \to (sn3 c (THead (Flat Appl) v (TLRef i)))))))
\def
(False_ind (sn3 c (THead (Bind x0) x5 (THead (Flat Appl) (lift (S O) O x4)
x3))) H14)) t2 H9)))))))))))))) H6)) H5))))))))) v H0))))).
-theorem sn3_appl_abbr:
+lemma sn3_appl_abbr:
\forall (c: C).(\forall (d: C).(\forall (w: T).(\forall (i: nat).((getl i c
(CHead d (Bind Abbr) w)) \to (\forall (v: T).((sn3 c (THead (Flat Appl) v
(lift (S i) O w))) \to (sn3 c (THead (Flat Appl) v (TLRef i)))))))))
Appl) v2 u2))))))).(sn3_appl_appl v1 (THeads (Flat Appl) vs t1) c H v2 H0
H1))))))))).
-theorem sn3_appls_lref:
+lemma sn3_appls_lref:
\forall (c: C).(\forall (i: nat).((nf2 c (TLRef i)) \to (\forall (us:
TList).((sns3 c us) \to (sn3 c (THeads (Flat Appl) us (TLRef i)))))))
\def
t))) H1 (THead (Flat Appl) u u2) (pr3_thin_dx c (THeads (Flat Appl) (TCons t0
t1) (THead (Bind Abbr) v t)) u2 H8 u Appl))))))))) H3)))))))))) us0))) us)))).
-theorem sn3_lift:
+lemma sn3_lift:
\forall (d: C).(\forall (t: T).((sn3 d t) \to (\forall (c: C).(\forall (h:
nat).(\forall (i: nat).((drop h i c d) \to (sn3 c (lift h i t))))))))
\def
(refl_equal T (lift h i t1)) P))))) (pr3_pr2 d t1 x H7) c h i H2) t2 H6)))))
H5))))))))))))) t H))).
-theorem sn3_abbr:
+lemma sn3_abbr:
\forall (c: C).(\forall (d: C).(\forall (v: T).(\forall (i: nat).((getl i c
(CHead d (Bind Abbr) v)) \to ((sn3 d v) \to (sn3 c (TLRef i)))))))
\def
v))) H12 d H11) in (sn3_lift d v H0 c (S i) O (getl_drop Abbr c d v i H13)))
x1 H10)))) H9))) t2 H6)))))) H4)) H3))))))))))).
-theorem sn3_appls_abbr:
+lemma sn3_appls_abbr:
\forall (c: C).(\forall (d: C).(\forall (w: T).(\forall (i: nat).((getl i c
(CHead d (Bind Abbr) w)) \to (\forall (vs: TList).((sn3 c (THeads (Flat Appl)
vs (lift (S i) O w))) \to (sn3 c (THeads (Flat Appl) vs (TLRef i)))))))))
(pr3_iso_appls_abbr c d w i H (TCons t t0) u2 H6 H7) v Appl))))))))
H3)))))))) vs0))) vs)))))).
-theorem sns3_lifts:
+lemma sns3_lifts:
\forall (c: C).(\forall (d: C).(\forall (h: nat).(\forall (i: nat).((drop h
i c d) \to (\forall (ts: TList).((sns3 d ts) \to (sns3 c (lifts h i ts))))))))
\def
include "basic_1/sty0/defs.ma".
-let rec sty0_ind (g: G) (P: (C \to (T \to (T \to Prop)))) (f: (\forall (c:
-C).(\forall (n: nat).(P c (TSort n) (TSort (next g n)))))) (f0: (\forall (c:
-C).(\forall (d: C).(\forall (v: T).(\forall (i: nat).((getl i c (CHead d
-(Bind Abbr) v)) \to (\forall (w: T).((sty0 g d v w) \to ((P d v w) \to (P c
-(TLRef i) (lift (S i) O w))))))))))) (f1: (\forall (c: C).(\forall (d:
-C).(\forall (v: T).(\forall (i: nat).((getl i c (CHead d (Bind Abst) v)) \to
-(\forall (w: T).((sty0 g d v w) \to ((P d v w) \to (P c (TLRef i) (lift (S i)
-O v))))))))))) (f2: (\forall (b: B).(\forall (c: C).(\forall (v: T).(\forall
-(t1: T).(\forall (t2: T).((sty0 g (CHead c (Bind b) v) t1 t2) \to ((P (CHead
-c (Bind b) v) t1 t2) \to (P c (THead (Bind b) v t1) (THead (Bind b) v
-t2)))))))))) (f3: (\forall (c: C).(\forall (v: T).(\forall (t1: T).(\forall
-(t2: T).((sty0 g c t1 t2) \to ((P c t1 t2) \to (P c (THead (Flat Appl) v t1)
-(THead (Flat Appl) v t2))))))))) (f4: (\forall (c: C).(\forall (v1:
-T).(\forall (v2: T).((sty0 g c v1 v2) \to ((P c v1 v2) \to (\forall (t1:
+implied let rec sty0_ind (g: G) (P: (C \to (T \to (T \to Prop)))) (f:
+(\forall (c: C).(\forall (n: nat).(P c (TSort n) (TSort (next g n)))))) (f0:
+(\forall (c: C).(\forall (d: C).(\forall (v: T).(\forall (i: nat).((getl i c
+(CHead d (Bind Abbr) v)) \to (\forall (w: T).((sty0 g d v w) \to ((P d v w)
+\to (P c (TLRef i) (lift (S i) O w))))))))))) (f1: (\forall (c: C).(\forall
+(d: C).(\forall (v: T).(\forall (i: nat).((getl i c (CHead d (Bind Abst) v))
+\to (\forall (w: T).((sty0 g d v w) \to ((P d v w) \to (P c (TLRef i) (lift
+(S i) O v))))))))))) (f2: (\forall (b: B).(\forall (c: C).(\forall (v:
+T).(\forall (t1: T).(\forall (t2: T).((sty0 g (CHead c (Bind b) v) t1 t2) \to
+((P (CHead c (Bind b) v) t1 t2) \to (P c (THead (Bind b) v t1) (THead (Bind
+b) v t2)))))))))) (f3: (\forall (c: C).(\forall (v: T).(\forall (t1:
+T).(\forall (t2: T).((sty0 g c t1 t2) \to ((P c t1 t2) \to (P c (THead (Flat
+Appl) v t1) (THead (Flat Appl) v t2))))))))) (f4: (\forall (c: C).(\forall
+(v1: T).(\forall (v2: T).((sty0 g c v1 v2) \to ((P c v1 v2) \to (\forall (t1:
T).(\forall (t2: T).((sty0 g c t1 t2) \to ((P c t1 t2) \to (P c (THead (Flat
Cast) v1 t1) (THead (Flat Cast) v2 t2)))))))))))) (c: C) (t: T) (t0: T) (s0:
sty0 g c t t0) on s0: P c t t0 \def match s0 with [(sty0_sort c0 n)
((sty0_ind g P f f0 f1 f2 f3 f4) c0 v1 v2 s1) t1 t2 s2 ((sty0_ind g P f f0 f1
f2 f3 f4) c0 t1 t2 s2))].
-theorem sty0_gen_sort:
+lemma sty0_gen_sort:
\forall (g: G).(\forall (c: C).(\forall (x: T).(\forall (n: nat).((sty0 g c
(TSort n) x) \to (eq T x (TSort (next g n)))))))
\def
(THead (Flat Cast) v2 t2) (TSort (next g n))) H6)))))))))))) c y x H0)))
H))))).
-theorem sty0_gen_lref:
+lemma sty0_gen_lref:
\forall (g: G).(\forall (c: C).(\forall (x: T).(\forall (n: nat).((sty0 g c
(TLRef n) x) \to (or (ex3_3 C T T (\lambda (e: C).(\lambda (u: T).(\lambda
(_: T).(getl n c (CHead e (Bind Abbr) u))))) (\lambda (e: C).(\lambda (u:
(_: T).(eq T (THead (Flat Cast) v2 t2) (lift (S n) O u))))))) H6))))))))))))
c y x H0))) H))))).
-theorem sty0_gen_bind:
+lemma sty0_gen_bind:
\forall (g: G).(\forall (b: B).(\forall (c: C).(\forall (u: T).(\forall (t1:
T).(\forall (x: T).((sty0 g c (THead (Bind b) u t1) x) \to (ex2 T (\lambda
(t2: T).(sty0 g (CHead c (Bind b) u) t1 t2)) (\lambda (t2: T).(eq T x (THead
t3)) (\lambda (t3: T).(eq T (THead (Flat Cast) v2 t2) (THead (Bind b) u
t3)))) H6)))))))))))) c y x H0))) H))))))).
-theorem sty0_gen_appl:
+lemma sty0_gen_appl:
\forall (g: G).(\forall (c: C).(\forall (u: T).(\forall (t1: T).(\forall (x:
T).((sty0 g c (THead (Flat Appl) u t1) x) \to (ex2 T (\lambda (t2: T).(sty0 g
c t1 t2)) (\lambda (t2: T).(eq T x (THead (Flat Appl) u t2)))))))))
T).(sty0 g c0 t1 t3)) (\lambda (t3: T).(eq T (THead (Flat Cast) v2 t2) (THead
(Flat Appl) u t3)))) H6)))))))))))) c y x H0))) H)))))).
-theorem sty0_gen_cast:
+lemma sty0_gen_cast:
\forall (g: G).(\forall (c: C).(\forall (v1: T).(\forall (t1: T).(\forall
(x: T).((sty0 g c (THead (Flat Cast) v1 t1) x) \to (ex3_2 T T (\lambda (v2:
T).(\lambda (_: T).(sty0 g c v1 v2))) (\lambda (_: T).(\lambda (t2: T).(sty0
include "basic_1/getl/drop.ma".
-theorem sty0_lift:
+lemma sty0_lift:
\forall (g: G).(\forall (e: C).(\forall (t1: T).(\forall (t2: T).((sty0 g e
t1 t2) \to (\forall (c: C).(\forall (h: nat).(\forall (d: nat).((drop h d c
e) \to (sty0 g c (lift h d t1) (lift h d t2))))))))))
Cast) v2 t4 h d)) (lift h d (THead (Flat Cast) v1 t3)) (lift_head (Flat Cast)
v1 t3 h d))))))))))))))) e t1 t2 H))))).
-theorem sty0_correct:
+lemma sty0_correct:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t: T).((sty0 g c
t1 t) \to (ex T (\lambda (t2: T).(sty0 g c t t2)))))))
\def
include "basic_1/cnt/props.ma".
-theorem sty1_cnt:
+lemma sty1_cnt:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t: T).((sty0 g c
t1 t) \to (ex2 T (\lambda (t2: T).(sty1 g c t1 t2)) (\lambda (t2: T).(cnt
t2)))))))
include "basic_1/sty1/defs.ma".
-let rec sty1_ind (g: G) (c: C) (t1: T) (P: (T \to Prop)) (f: (\forall (t2:
-T).((sty0 g c t1 t2) \to (P t2)))) (f0: (\forall (t: T).((sty1 g c t1 t) \to
-((P t) \to (\forall (t2: T).((sty0 g c t t2) \to (P t2))))))) (t: T) (s0:
+implied let rec sty1_ind (g: G) (c: C) (t1: T) (P: (T \to Prop)) (f: (\forall
+(t2: T).((sty0 g c t1 t2) \to (P t2)))) (f0: (\forall (t: T).((sty1 g c t1 t)
+\to ((P t) \to (\forall (t2: T).((sty0 g c t t2) \to (P t2))))))) (t: T) (s0:
sty1 g c t1 t) on s0: P t \def match s0 with [(sty1_sty0 t2 s1) \Rightarrow
(f t2 s1) | (sty1_sing t0 s1 t2 s2) \Rightarrow (f0 t0 s1 ((sty1_ind g c t1 P
f f0) t0 s1) t2 s2)].
c t t0)).(\lambda (H2: (sty1 g c t1 t0)).(\lambda (t3: T).(\lambda (H3: (sty0
g c t0 t3)).(sty1_sing g c t1 t0 H2 t3 H3)))))) t2 H0))))))).
-theorem sty1_bind:
+lemma sty1_bind:
\forall (g: G).(\forall (b: B).(\forall (c: C).(\forall (v: T).(\forall (t1:
T).(\forall (t2: T).((sty1 g (CHead c (Bind b) v) t1 t2) \to (sty1 g c (THead
(Bind b) v t1) (THead (Bind b) v t2))))))))
(Bind b) v) t t3)).(sty1_sing g c (THead (Bind b) v t1) (THead (Bind b) v t)
H1 (THead (Bind b) v t3) (sty0_bind g b c v t t3 H2))))))) t2 H))))))).
-theorem sty1_appl:
+lemma sty1_appl:
\forall (g: G).(\forall (c: C).(\forall (v: T).(\forall (t1: T).(\forall
(t2: T).((sty1 g c t1 t2) \to (sty1 g c (THead (Flat Appl) v t1) (THead (Flat
Appl) v t2)))))))
t3)).(sty1_sing g c (THead (Flat Appl) v t1) (THead (Flat Appl) v t) H1
(THead (Flat Appl) v t3) (sty0_appl g c v t t3 H2))))))) t2 H)))))).
-theorem sty1_lift:
+lemma sty1_lift:
\forall (g: G).(\forall (e: C).(\forall (t1: T).(\forall (t2: T).((sty1 g e
t1 t2) \to (\forall (c: C).(\forall (h: nat).(\forall (d: nat).((drop h d c
e) \to (sty1 g c (lift h d t1) (lift h d t2))))))))))
(H3: (drop h d c e)).(sty1_sing g c (lift h d t1) (lift h d t) (H1 c h d H3)
(lift h d t3) (sty0_lift g e t t3 H2 c h d H3))))))))))) t2 H))))).
-theorem sty1_correct:
+lemma sty1_correct:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t: T).((sty1 g c
t1 t) \to (ex T (\lambda (t2: T).(sty0 g c t t2)))))))
\def
t0)).(\lambda (_: (ex T (\lambda (t2: T).(sty0 g c t0 t2)))).(\lambda (t2:
T).(\lambda (H2: (sty0 g c t0 t2)).(sty0_correct g c t0 t2 H2)))))) t H))))).
-theorem sty1_abbr:
+lemma sty1_abbr:
\forall (g: G).(\forall (c: C).(\forall (d: C).(\forall (v: T).(\forall (i:
nat).((getl i c (CHead d (Bind Abbr) v)) \to (\forall (w: T).((sty1 g d v w)
\to (sty1 g c (TLRef i) (lift (S i) O w)))))))))
(sty0_lift g d t t2 H3 c (S i) O (getl_drop Abbr c d v i H)))))))) w
H0)))))))).
-theorem sty1_cast2:
+lemma sty1_cast2:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t2: T).((sty1 g c
t1 t2) \to (\forall (v1: T).(\forall (v2: T).((sty0 g c v1 v2) \to (ex2 T
(\lambda (v3: T).(sty1 g c v1 v3)) (\lambda (v3: T).(sty1 g c (THead (Flat
include "basic_1/subst0/fwd.ma".
-theorem subst_sort:
+lemma subst_sort:
\forall (v: T).(\forall (d: nat).(\forall (k: nat).(eq T (subst d v (TSort
k)) (TSort k))))
\def
\lambda (_: T).(\lambda (_: nat).(\lambda (k: nat).(refl_equal T (TSort
k)))).
-theorem subst_lref_lt:
+lemma subst_lref_lt:
\forall (v: T).(\forall (d: nat).(\forall (i: nat).((lt i d) \to (eq T
(subst d v (TLRef i)) (TLRef i)))))
\def
\Rightarrow (TLRef (pred i)) | false \Rightarrow (lift d O v)])]) (TLRef i)))
(refl_equal T (TLRef i)) (blt i d) (lt_blt d i H))))).
-theorem subst_lref_eq:
+lemma subst_lref_eq:
\forall (v: T).(\forall (i: nat).(eq T (subst i v (TLRef i)) (lift i O v)))
\def
\lambda (v: T).(\lambda (i: nat).(eq_ind_r bool false (\lambda (b: bool).(eq
[true \Rightarrow (TLRef (pred i)) | false \Rightarrow (lift i O v)])]) (lift
i O v))) (refl_equal T (lift i O v)) (blt i i) (le_bge i i (le_n i)))).
-theorem subst_lref_gt:
+lemma subst_lref_gt:
\forall (v: T).(\forall (d: nat).(\forall (i: nat).((lt d i) \to (eq T
(subst d v (TLRef i)) (TLRef (pred i))))))
\def
i)))) (refl_equal T (TLRef (pred i))) (blt d i) (lt_blt i d H)) (blt i d)
(le_bge d i (lt_le_weak d i H)))))).
-theorem subst_head:
+lemma subst_head:
\forall (k: K).(\forall (w: T).(\forall (u: T).(\forall (t: T).(\forall (d:
nat).(eq T (subst d w (THead k u t)) (THead k (subst d w u) (subst (s k d) w
t)))))))
\lambda (k: K).(\lambda (w: T).(\lambda (u: T).(\lambda (t: T).(\lambda (d:
nat).(refl_equal T (THead k (subst d w u) (subst (s k d) w t))))))).
-theorem subst_lift_SO:
+lemma subst_lift_SO:
\forall (v: T).(\forall (t: T).(\forall (d: nat).(eq T (subst d v (lift (S
O) d t)) t)))
\def
(S O) (s k d) t1) d)) (lift (S O) d (THead k t0 t1)) (lift_head k t0 t1 (S O)
d)))))))) t)).
-theorem subst_subst0:
+lemma subst_subst0:
\forall (v: T).(\forall (t1: T).(\forall (t2: T).(\forall (d: nat).((subst0
d v t1 t2) \to (eq T (subst d v t1) (subst d v t2))))))
\def
include "basic_1/lift/props.ma".
-theorem dnf_dec2:
+lemma dnf_dec2:
\forall (t: T).(\forall (d: nat).(or (\forall (w: T).(ex T (\lambda (v:
T).(subst0 d w t (lift (S O) d v))))) (ex T (\lambda (v: T).(eq T t (lift (S
O) d v))))))
(S O) (s k d) x0))) (lift (S O) d (THead k x x0)) (lift_head k x x0 (S O)
d)))) t0 H3) t1 H6))) H5)) H4))))) H2)) H1))))))))) t).
-theorem dnf_dec:
+lemma dnf_dec:
\forall (w: T).(\forall (t: T).(\forall (d: nat).(ex T (\lambda (v: T).(or
(subst0 d w t (lift (S O) d v)) (eq T t (lift (S O) d v)))))))
\def
include "basic_1/lift/fwd.ma".
-let rec subst0_ind (P: (nat \to (T \to (T \to (T \to Prop))))) (f: (\forall
-(v: T).(\forall (i: nat).(P i v (TLRef i) (lift (S i) O v))))) (f0: (\forall
-(v: T).(\forall (u2: T).(\forall (u1: T).(\forall (i: nat).((subst0 i v u1
-u2) \to ((P i v u1 u2) \to (\forall (t: T).(\forall (k: K).(P i v (THead k u1
-t) (THead k u2 t))))))))))) (f1: (\forall (k: K).(\forall (v: T).(\forall
-(t2: T).(\forall (t1: T).(\forall (i: nat).((subst0 (s k i) v t1 t2) \to ((P
-(s k i) v t1 t2) \to (\forall (u: T).(P i v (THead k u t1) (THead k u
-t2))))))))))) (f2: (\forall (v: T).(\forall (u1: T).(\forall (u2: T).(\forall
-(i: nat).((subst0 i v u1 u2) \to ((P i v u1 u2) \to (\forall (k: K).(\forall
-(t1: T).(\forall (t2: T).((subst0 (s k i) v t1 t2) \to ((P (s k i) v t1 t2)
-\to (P i v (THead k u1 t1) (THead k u2 t2)))))))))))))) (n: nat) (t: T) (t0:
-T) (t1: T) (s0: subst0 n t t0 t1) on s0: P n t t0 t1 \def match s0 with
-[(subst0_lref v i) \Rightarrow (f v i) | (subst0_fst v u2 u1 i s1 t2 k)
-\Rightarrow (f0 v u2 u1 i s1 ((subst0_ind P f f0 f1 f2) i v u1 u2 s1) t2 k) |
-(subst0_snd k v t2 t3 i s1 u) \Rightarrow (f1 k v t2 t3 i s1 ((subst0_ind P f
-f0 f1 f2) (s k i) v t3 t2 s1) u) | (subst0_both v u1 u2 i s1 k t2 t3 s2)
-\Rightarrow (f2 v u1 u2 i s1 ((subst0_ind P f f0 f1 f2) i v u1 u2 s1) k t2 t3
-s2 ((subst0_ind P f f0 f1 f2) (s k i) v t2 t3 s2))].
+implied let rec subst0_ind (P: (nat \to (T \to (T \to (T \to Prop))))) (f:
+(\forall (v: T).(\forall (i: nat).(P i v (TLRef i) (lift (S i) O v))))) (f0:
+(\forall (v: T).(\forall (u2: T).(\forall (u1: T).(\forall (i: nat).((subst0
+i v u1 u2) \to ((P i v u1 u2) \to (\forall (t: T).(\forall (k: K).(P i v
+(THead k u1 t) (THead k u2 t))))))))))) (f1: (\forall (k: K).(\forall (v:
+T).(\forall (t2: T).(\forall (t1: T).(\forall (i: nat).((subst0 (s k i) v t1
+t2) \to ((P (s k i) v t1 t2) \to (\forall (u: T).(P i v (THead k u t1) (THead
+k u t2))))))))))) (f2: (\forall (v: T).(\forall (u1: T).(\forall (u2:
+T).(\forall (i: nat).((subst0 i v u1 u2) \to ((P i v u1 u2) \to (\forall (k:
+K).(\forall (t1: T).(\forall (t2: T).((subst0 (s k i) v t1 t2) \to ((P (s k
+i) v t1 t2) \to (P i v (THead k u1 t1) (THead k u2 t2)))))))))))))) (n: nat)
+(t: T) (t0: T) (t1: T) (s0: subst0 n t t0 t1) on s0: P n t t0 t1 \def match
+s0 with [(subst0_lref v i) \Rightarrow (f v i) | (subst0_fst v u2 u1 i s1 t2
+k) \Rightarrow (f0 v u2 u1 i s1 ((subst0_ind P f f0 f1 f2) i v u1 u2 s1) t2
+k) | (subst0_snd k v t2 t3 i s1 u) \Rightarrow (f1 k v t2 t3 i s1
+((subst0_ind P f f0 f1 f2) (s k i) v t3 t2 s1) u) | (subst0_both v u1 u2 i s1
+k t2 t3 s2) \Rightarrow (f2 v u1 u2 i s1 ((subst0_ind P f f0 f1 f2) i v u1 u2
+s1) k t2 t3 s2 ((subst0_ind P f f0 f1 f2) (s k i) v t2 t3 s2))].
-theorem subst0_gen_sort:
+lemma subst0_gen_sort:
\forall (v: T).(\forall (x: T).(\forall (i: nat).(\forall (n: nat).((subst0
i v (TSort n) x) \to (\forall (P: Prop).P)))))
\def
True])) I (TSort n) H5) in (False_ind P H6)))))))))))))) i v y x H0)))
H)))))).
-theorem subst0_gen_lref:
+lemma subst0_gen_lref:
\forall (v: T).(\forall (x: T).(\forall (i: nat).(\forall (n: nat).((subst0
i v (TLRef n) x) \to (land (eq nat n i) (eq T x (lift (S n) O v)))))))
\def
(TLRef n) H5) in (False_ind (land (eq nat n i0) (eq T (THead k u2 t2) (lift
(S n) O v0))) H6)))))))))))))) i v y x H0))) H))))).
-theorem subst0_gen_head:
+lemma subst0_gen_head:
\forall (k: K).(\forall (v: T).(\forall (u1: T).(\forall (t1: T).(\forall
(x: T).(\forall (i: nat).((subst0 i v (THead k u1 t1) x) \to (or3 (ex2 T
(\lambda (u2: T).(eq T x (THead k u2 t1))) (\lambda (u2: T).(subst0 i v u1
u2 t2)) H16 H14)))) k0 H10)))))))) H7)) H6)))))))))))))) i v y x H0)))
H))))))).
-theorem subst0_gen_lift_lt:
+lemma subst0_gen_lift_lt:
\forall (u: T).(\forall (t1: T).(\forall (x: T).(\forall (i: nat).(\forall
(h: nat).(\forall (d: nat).((subst0 i (lift h d u) (lift h (S (plus i d)) t1)
x) \to (ex2 T (\lambda (t2: T).(eq T x (lift h (S (plus i d)) t2))) (\lambda
(lift h d u) (lift h (S (plus i d)) t) (lift h (s k (S (plus i d))) t0) x i
H2))))))))))))) t1)).
-theorem subst0_gen_lift_false:
+lemma subst0_gen_lift_false:
\forall (t: T).(\forall (u: T).(\forall (x: T).(\forall (h: nat).(\forall
(d: nat).(\forall (i: nat).((le d i) \to ((lt i (plus d h)) \to ((subst0 i u
(lift h d t) x) \to (\forall (P: Prop).P)))))))))
i H1 H2 H7 P)))))) H5)) (subst0_gen_head k u (lift h d t0) (lift h (s k d)
t1) x i H4))))))))))))))))) t).
-theorem subst0_gen_lift_ge:
+lemma subst0_gen_lift_ge:
\forall (u: T).(\forall (t1: T).(\forall (x: T).(\forall (i: nat).(\forall
(h: nat).(\forall (d: nat).((subst0 i u (lift h d t1) x) \to ((le (plus d h)
i) \to (ex2 T (\lambda (t2: T).(eq T x (lift h d t2))) (\lambda (t2:
H2) (plus (s k d) h) (s_plus k d h)))) x H5)))))) H4)) (subst0_gen_head k u
(lift h d t) (lift h (s k d) t0) x i H3)))))))))))))) t1)).
-theorem subst0_gen_lift_rev_ge:
+lemma subst0_gen_lift_rev_ge:
\forall (t1: T).(\forall (v: T).(\forall (u2: T).(\forall (i: nat).(\forall
(h: nat).(\forall (d: nat).((subst0 i v t1 (lift h d u2)) \to ((le (plus d h)
i) \to (ex2 T (\lambda (u1: T).(subst0 (minus i h) v u1 u2)) (\lambda (u1:
include "basic_1/subst0/fwd.ma".
-theorem subst0_refl:
+lemma subst0_refl:
\forall (u: T).(\forall (t: T).(\forall (d: nat).((subst0 d u t t) \to
(\forall (P: Prop).P))))
\def
t0 t2)) H4 t0 H8) in (H d H10 P))))) H6))))))) H2)) (subst0_gen_head k u t0
t1 (THead k t0 t1) d H1)))))))))) t)).
-theorem subst0_lift_lt:
+lemma subst0_lift_lt:
\forall (t1: T).(\forall (t2: T).(\forall (u: T).(\forall (i: nat).((subst0
i u t1 t2) \to (\forall (d: nat).((lt i d) \to (\forall (h: nat).(subst0 i
(lift h (minus d (S i)) u) (lift h d t1) (lift h d t2)))))))))
(THead k u2 t3)) (lift_head k u2 t3 h d)) (lift h d (THead k u1 t0))
(lift_head k u1 t0 h d))))))))))))))))) i u t1 t2 H))))).
-theorem subst0_lift_ge:
+lemma subst0_lift_ge:
\forall (t1: T).(\forall (t2: T).(\forall (u: T).(\forall (i: nat).(\forall
(h: nat).((subst0 i u t1 t2) \to (\forall (d: nat).((le d i) \to (subst0
(plus i h) u (lift h d t1) (lift h d t2)))))))))
i0 H4))) (lift h d (THead k u2 t3)) (lift_head k u2 t3 h d)) (lift h d (THead
k u1 t0)) (lift_head k u1 t0 h d)))))))))))))))) i u t1 t2 H)))))).
-theorem subst0_lift_ge_S:
+lemma subst0_lift_ge_S:
\forall (t1: T).(\forall (t2: T).(\forall (u: T).(\forall (i: nat).((subst0
i u t1 t2) \to (\forall (d: nat).((le d i) \to (subst0 (S i) u (lift (S O) d
t1) (lift (S O) d t2))))))))
i) (\lambda (n: nat).(eq nat n (S i))) (le_antisym (plus (S O) i) (S i) (le_n
(S i)) (le_n (plus (S O) i))) (plus i (S O)) (plus_sym i (S O)))))))))).
-theorem subst0_lift_ge_s:
+lemma subst0_lift_ge_s:
\forall (t1: T).(\forall (t2: T).(\forall (u: T).(\forall (i: nat).((subst0
i u t1 t2) \to (\forall (d: nat).((le d i) \to (\forall (b: B).(subst0 (s
(Bind b) i) u (lift (S O) d t1) (lift (S O) d t2)))))))))
include "basic_1/lift/tlt.ma".
-theorem subst0_weight_le:
+lemma subst0_weight_le:
\forall (u: T).(\forall (t: T).(\forall (z: T).(\forall (d: nat).((subst0 d
u t z) \to (\forall (f: ((nat \to nat))).(\forall (g: ((nat \to
nat))).(((\forall (m: nat).(le (f m) (g m)))) \to ((lt (weight_map f (lift (S
(weight_map f0 u2) (weight_map g u1) (weight_map f0 t2) (weight_map g t1) (H1
f0 g H4 H5) (H3 f0 g H4 H5)))))))))))) k)))))))) d u t z H))))).
-theorem subst0_weight_lt:
+lemma subst0_weight_lt:
\forall (u: T).(\forall (t: T).(\forall (z: T).(\forall (d: nat).((subst0 d
u t z) \to (\forall (f: ((nat \to nat))).(\forall (g: ((nat \to
nat))).(((\forall (m: nat).(le (f m) (g m)))) \to ((lt (weight_map f (lift (S
(weight_map g t1) (H1 f0 g H4 H5) (H3 f0 g H4 H5)))))))))))) k)))))))) d u t
z H))))).
-theorem subst0_tlt_head:
+lemma subst0_tlt_head:
\forall (u: T).(\forall (t: T).(\forall (z: T).((subst0 O u t z) \to (tlt
(THead (Bind Abbr) u z) (THead (Bind Abbr) u t)))))
\def
(_: nat).O) u))) (lift (S O) O u)) (lift_weight_add_O (S (weight_map (\lambda
(_: nat).O) u)) u O (\lambda (_: nat).O))))))))).
-theorem subst0_tlt:
+lemma subst0_tlt:
\forall (u: T).(\forall (t: T).(\forall (z: T).((subst0 O u t z) \to (tlt z
(THead (Bind Abbr) u t)))))
\def
include "basic_1/subst0/fwd.ma".
-theorem subst1_ind:
+implied lemma subst1_ind:
\forall (i: nat).(\forall (v: T).(\forall (t1: T).(\forall (P: ((T \to
Prop))).((P t1) \to (((\forall (t2: T).((subst0 i v t1 t2) \to (P t2)))) \to
(\forall (t: T).((subst1 i v t1 t) \to (P t))))))))
with [subst1_refl \Rightarrow f | (subst1_single x x0) \Rightarrow (f0 x
x0)])))))))).
-theorem subst1_gen_sort:
+lemma subst1_gen_sort:
\forall (v: T).(\forall (x: T).(\forall (i: nat).(\forall (n: nat).((subst1
i v (TSort n) x) \to (eq T x (TSort n))))))
\def
i v (TSort n) t2)).(subst0_gen_sort v t2 i n H0 (eq T t2 (TSort n))))) x
H))))).
-theorem subst1_gen_lref:
+lemma subst1_gen_lref:
\forall (v: T).(\forall (x: T).(\forall (i: nat).(\forall (n: nat).((subst1
i v (TLRef n) x) \to (or (eq T x (TLRef n)) (land (eq nat n i) (eq T x (lift
(S n) O v))))))))
(eq T t2 (lift (S n) O v)) H1 H2)))) (subst0_gen_lref v t2 i n H0)))) x
H))))).
-theorem subst1_gen_head:
+lemma subst1_gen_head:
\forall (k: K).(\forall (v: T).(\forall (u1: T).(\forall (t1: T).(\forall
(x: T).(\forall (i: nat).((subst1 i v (THead k u1 t1) x) \to (ex3_2 T T
(\lambda (u2: T).(\lambda (t2: T).(eq T x (THead k u2 t2)))) (\lambda (u2:
x1 H2 (subst1_single i v u1 x0 H3) (subst1_single (s k i) v t1 x1 H4)))))))
H1)) (subst0_gen_head k v u1 t1 t2 i H0)))) x H))))))).
-theorem subst1_gen_lift_lt:
+lemma subst1_gen_lift_lt:
\forall (u: T).(\forall (t1: T).(\forall (x: T).(\forall (i: nat).(\forall
(h: nat).(\forall (d: nat).((subst1 i (lift h d u) (lift h (S (plus i d)) t1)
x) \to (ex2 T (\lambda (t2: T).(eq T x (lift h (S (plus i d)) t2))) (\lambda
(subst1_single i u t1 x0 H2))))) (subst0_gen_lift_lt u t1 t2 i h d H0)))) x
H))))))).
-theorem subst1_gen_lift_eq:
+lemma subst1_gen_lift_eq:
\forall (t: T).(\forall (u: T).(\forall (x: T).(\forall (h: nat).(\forall
(d: nat).(\forall (i: nat).((le d i) \to ((lt i (plus d h)) \to ((subst1 i u
(lift h d t) x) \to (eq T x (lift h d t))))))))))
(t2: T).(\lambda (H2: (subst0 i u (lift h d t) t2)).(subst0_gen_lift_false t
u t2 h d i H H0 H2 (eq T t2 (lift h d t))))) x H1))))))))).
-theorem subst1_gen_lift_ge:
+lemma subst1_gen_lift_ge:
\forall (u: T).(\forall (t1: T).(\forall (x: T).(\forall (i: nat).(\forall
(h: nat).(\forall (d: nat).((subst1 i u (lift h d t1) x) \to ((le (plus d h)
i) \to (ex2 T (\lambda (t2: T).(eq T x (lift h d t2))) (\lambda (t2:
i) v t1 t3)).(subst1_single i v (THead k u1 t1) (THead k t2 t3) (subst0_both
v u1 t2 i H0 k t1 t3 H2)))) t0 H1))))))) u2 H))))).
-theorem subst1_lift_lt:
+lemma subst1_lift_lt:
\forall (t1: T).(\forall (t2: T).(\forall (u: T).(\forall (i: nat).((subst1
i u t1 t2) \to (\forall (d: nat).((lt i d) \to (\forall (h: nat).(subst1 i
(lift h (minus d (S i)) u) (lift h d t1) (lift h d t2)))))))))
(minus d (S i)) u) (lift h d t1) (lift h d t3) (subst0_lift_lt t1 t3 u i H0 d
H1 h))))))) t2 H))))).
-theorem subst1_lift_ge:
+lemma subst1_lift_ge:
\forall (t1: T).(\forall (t2: T).(\forall (u: T).(\forall (i: nat).(\forall
(h: nat).((subst1 i u t1 t2) \to (\forall (d: nat).((le d i) \to (subst1
(plus i h) u (lift h d t1) (lift h d t2)))))))))
(d: nat).(\lambda (H1: (le d i)).(subst1_single (plus i h) u (lift h d t1)
(lift h d t3) (subst0_lift_ge t1 t3 u i h H0 d H1)))))) t2 H)))))).
-theorem subst1_ex:
+lemma subst1_ex:
\forall (u: T).(\forall (t1: T).(\forall (d: nat).(ex T (\lambda (t2:
T).(subst1 d u t1 (lift (S O) d t2))))))
\def
(lift (S O) d x) d H2 k t0 (lift (S O) (s k d) x0) H4) (lift (S O) d (THead k
x x0)) (lift_head k x x0 (S O) d))))) H3))))) H1))))))))) t1)).
-theorem subst1_lift_S:
+lemma subst1_lift_S:
\forall (u: T).(\forall (i: nat).(\forall (h: nat).((le h i) \to (subst1 i
(TLRef h) (lift (S h) (S i) u) (lift (S h) i u)))))
\def
include "basic_1/tlist/props.ma".
-theorem tslt_wf__q_ind:
+fact tslt_wf__q_ind:
\forall (P: ((TList \to Prop))).(((\forall (n: nat).((\lambda (P0: ((TList
\to Prop))).(\lambda (n0: nat).(\forall (ts: TList).((eq nat (tslen ts) n0)
\to (P0 ts))))) P n))) \to (\forall (ts: TList).(P ts)))
ts) n) \to (P ts)))))).(\lambda (ts: TList).(H (tslen ts) ts (refl_equal nat
(tslen ts)))))).
-theorem tslt_wf_ind:
+lemma tslt_wf_ind:
\forall (P: ((TList \to Prop))).(((\forall (ts2: TList).(((\forall (ts1:
TList).((tslt ts1 ts2) \to (P ts1)))) \to (P ts2)))) \to (\forall (ts:
TList).(P ts)))
H1) in (H ts0 (\lambda (ts1: TList).(\lambda (H3: (lt (tslen ts1) (tslen
ts0))).(H2 (tslen ts1) H3 ts1 (refl_equal nat (tslen ts1))))))))))))) ts)))).
-theorem tlist_ind_rev:
+lemma tlist_ind_rev:
\forall (P: ((TList \to Prop))).((P TNil) \to (((\forall (ts:
TList).(\forall (t: T).((P ts) \to (P (TApp ts t)))))) \to (\forall (ts:
TList).(P ts))))
include "basic_1/tlist/defs.ma".
-theorem theads_tapp:
+lemma theads_tapp:
\forall (k: K).(\forall (v: T).(\forall (t: T).(\forall (vs: TList).(eq T
(THeads k (TApp vs v) t) (THeads k vs (THead k v t))))))
\def
(THeads k (TApp t1 v) t)) (THead k t0 t2))) (refl_equal T (THead k t0 (THeads
k (TApp t1 v) t))) (THeads k t1 (THead k v t)) H)))) vs)))).
-theorem tcons_tapp_ex:
+lemma tcons_tapp_ex:
\forall (ts1: TList).(\forall (t1: T).(ex2_2 TList T (\lambda (ts2:
TList).(\lambda (t2: T).(eq TList (TCons t1 ts1) (TApp ts2 t2)))) (\lambda
(ts2: TList).(\lambda (_: T).(eq nat (tslen ts1) (tslen ts2))))))
include "basic_1/tlt/defs.ma".
-theorem tlt_wf__q_ind:
+fact tlt_wf__q_ind:
\forall (P: ((T \to Prop))).(((\forall (n: nat).((\lambda (P0: ((T \to
Prop))).(\lambda (n0: nat).(\forall (t: T).((eq nat (weight t) n0) \to (P0
t))))) P n))) \to (\forall (t: T).(P t)))
n) \to (P t)))))).(\lambda (t: T).(H (weight t) t (refl_equal nat (weight
t)))))).
-theorem tlt_wf_ind:
+lemma tlt_wf_ind:
\forall (P: ((T \to Prop))).(((\forall (t: T).(((\forall (v: T).((tlt v t)
\to (P v)))) \to (P t)))) \to (\forall (t: T).(P t)))
\def
include "basic_1/tlt/defs.ma".
-theorem wadd_le:
+lemma wadd_le:
\forall (f: ((nat \to nat))).(\forall (g: ((nat \to nat))).(((\forall (n:
nat).(le (f n) (g n)))) \to (\forall (v: nat).(\forall (w: nat).((le v w) \to
(\forall (n: nat).(le (wadd f v n) (wadd g w n))))))))
nat).(le (wadd f v n0) (wadd g w n0))) H0 (\lambda (n0: nat).(\lambda (_: (le
(wadd f v n0) (wadd g w n0))).(H n0))) n))))))).
-theorem wadd_lt:
+lemma wadd_lt:
\forall (f: ((nat \to nat))).(\forall (g: ((nat \to nat))).(((\forall (n:
nat).(le (f n) (g n)))) \to (\forall (v: nat).(\forall (w: nat).((lt v w) \to
(\forall (n: nat).(le (wadd f v n) (wadd g w n))))))))
(S (S v)) (S w) (le_n_S (S v) w H0)))) (\lambda (n0: nat).(\lambda (_: (le
(wadd f v n0) (wadd g w n0))).(H n0))) n))))))).
-theorem wadd_O:
+lemma wadd_O:
\forall (n: nat).(eq nat (wadd (\lambda (_: nat).O) O n) O)
\def
\lambda (n: nat).(nat_ind (\lambda (n0: nat).(eq nat (wadd (\lambda (_:
nat).O) O n0) O)) (refl_equal nat O) (\lambda (n0: nat).(\lambda (_: (eq nat
(wadd (\lambda (_: nat).O) O n0) O)).(refl_equal nat O))) n).
-theorem weight_le:
+lemma weight_le:
\forall (t: T).(\forall (f: ((nat \to nat))).(\forall (g: ((nat \to
nat))).(((\forall (n: nat).(le (f n) (g n)))) \to (le (weight_map f t)
(weight_map g t)))))
t1)) (le_plus_plus (weight_map f0 t0) (weight_map g t0) (weight_map f0 t1)
(weight_map g t1) (H f0 g H1) (H0 f0 g H1))))))))))) k)) t).
-theorem weight_eq:
+lemma weight_eq:
\forall (t: T).(\forall (f: ((nat \to nat))).(\forall (g: ((nat \to
nat))).(((\forall (n: nat).(eq nat (f n) (g n)))) \to (eq nat (weight_map f
t) (weight_map g t)))))
(H n)))) (weight_le t g f (\lambda (n: nat).(eq_ind_r nat (g n) (\lambda (n0:
nat).(le (g n) n0)) (le_n (g n)) (f n) (H n)))))))).
-theorem weight_add_O:
+lemma weight_add_O:
\forall (t: T).(eq nat (weight_map (wadd (\lambda (_: nat).O) O) t)
(weight_map (\lambda (_: nat).O) t))
\def
\lambda (t: T).(weight_eq t (wadd (\lambda (_: nat).O) O) (\lambda (_:
nat).O) (\lambda (n: nat).(wadd_O n))).
-theorem weight_add_S:
+lemma weight_add_S:
\forall (t: T).(\forall (m: nat).(le (weight_map (wadd (\lambda (_: nat).O)
O) t) (weight_map (wadd (\lambda (_: nat).O) (S m)) t)))
\def
(weight v))).(\lambda (H0: (lt (weight v) (weight t))).(lt_trans (weight u)
(weight v) (weight t) H H0))))).
-theorem tlt_head_sx:
+lemma tlt_head_sx:
\forall (k: K).(\forall (u: T).(\forall (t: T).(tlt u (THead k u t))))
\def
\lambda (k: K).(K_ind (\lambda (k0: K).(\forall (u: T).(\forall (t: T).(lt
(weight_map (\lambda (_: nat).O) t)) (le_plus_l (weight_map (\lambda (_:
nat).O) u) (weight_map (\lambda (_: nat).O) t)))))) k).
-theorem tlt_head_dx:
+lemma tlt_head_dx:
\forall (k: K).(\forall (u: T).(\forall (t: T).(tlt t (THead k u t))))
\def
\lambda (k: K).(K_ind (\lambda (k0: K).(\forall (u: T).(\forall (t: T).(lt
include "basic_1/asucc/fwd.ma".
-theorem ty3_arity:
+lemma ty3_arity:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t2: T).((ty3 g c
t1 t2) \to (ex2 A (\lambda (a1: A).(arity g c t1 a1)) (\lambda (a1: A).(arity
g c t2 (asucc g a1))))))))
include "basic_1/sc3/arity.ma".
-theorem ty3_predicative:
+lemma ty3_predicative:
\forall (g: G).(\forall (c: C).(\forall (v: T).(\forall (t: T).(\forall (u:
T).((ty3 g c (THead (Bind Abst) v t) u) \to ((pc3 c u v) \to (\forall (P:
Prop).P)))))))
H11)) P)))) H9)))))) H6))))))) H3)))) (ty3_correct g (CHead c (Bind Abst) w)
t (lift (S O) O u2) H0))))))))))).
-theorem ty3_acyclic:
+lemma ty3_acyclic:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (u: T).((ty3 g c t
u) \to ((pc3 c u t) \to (\forall (P: Prop).P))))))
\def
c t x)).(\lambda (H3: (arity g c t (asucc g x))).(leq_asucc_false g x
(arity_mono g c t (asucc g x) H3 x H2) P)))) H1)))))))))).
-theorem ty3_sn3:
+lemma ty3_sn3:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (u: T).((ty3 g c t
u) \to (sn3 c t)))))
\def
include "basic_1/flt/fwd.ma".
-theorem ty3_inference:
+lemma ty3_inference:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(or (ex T (\lambda (t2:
T).(ty3 g c t1 t2))) (\forall (t2: T).((ty3 g c t1 t2) \to False)))))
\def
include "basic_1/getl/getl.ma".
-theorem ty3_fsubst0:
+lemma ty3_fsubst0:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t: T).((ty3 g c1
t1 t) \to (\forall (i: nat).(\forall (u: T).(\forall (c2: C).(\forall (t2:
T).((fsubst0 i u c1 t1 c2 t2) \to (\forall (e: C).((getl i c1 (CHead e (Bind
i u c t3 c3 H6) e H7))) t5 H9)))))) H8)) (subst0_gen_head (Flat Cast) u t3 t2
t5 i H5)))))))) c2 t4 H4)))))))))))))) c1 t1 t H))))).
-theorem ty3_csubst0:
+lemma ty3_csubst0:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t2: T).((ty3 g c1
t1 t2) \to (\forall (e: C).(\forall (u: T).(\forall (i: nat).((getl i c1
(CHead e (Bind Abbr) u)) \to (\forall (c2: C).((csubst0 i u c1 c2) \to (ty3 g
C).(\lambda (H1: (csubst0 i u c1 c2)).(ty3_fsubst0 g c1 t1 t2 H i u c2 t1
(fsubst0_fst i u c1 t1 c2 H1) e H0))))))))))).
-theorem ty3_subst0:
+lemma ty3_subst0:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t: T).((ty3 g c t1
t) \to (\forall (e: C).(\forall (u: T).(\forall (i: nat).((getl i c (CHead e
(Bind Abbr) u)) \to (\forall (t2: T).((subst0 i u t1 t2) \to (ty3 g c t2
include "basic_1/pc3/props.ma".
-let rec ty3_ind (g: G) (P: (C \to (T \to (T \to Prop)))) (f: (\forall (c:
-C).(\forall (t2: T).(\forall (t: T).((ty3 g c t2 t) \to ((P c t2 t) \to
+implied let rec ty3_ind (g: G) (P: (C \to (T \to (T \to Prop)))) (f: (\forall
+(c: C).(\forall (t2: T).(\forall (t: T).((ty3 g c t2 t) \to ((P c t2 t) \to
(\forall (u: T).(\forall (t1: T).((ty3 g c u t1) \to ((P c u t1) \to ((pc3 c
t1 t2) \to (P c u t2)))))))))))) (f0: (\forall (c: C).(\forall (m: nat).(P c
(TSort m) (TSort (next g m)))))) (f1: (\forall (n: nat).(\forall (c:
f1 f2 f3 f4 f5) c0 t2 t3 t4) t5 t6 ((ty3_ind g P f f0 f1 f2 f3 f4 f5) c0 t3
t5 t6))].
-let rec tys3_ind (g: G) (c: C) (P: (TList \to (T \to Prop))) (f: (\forall (u:
-T).(\forall (u0: T).((ty3 g c u u0) \to (P TNil u))))) (f0: (\forall (t:
-T).(\forall (u: T).((ty3 g c t u) \to (\forall (ts: TList).((tys3 g c ts u)
-\to ((P ts u) \to (P (TCons t ts) u)))))))) (t: TList) (t0: T) (t1: tys3 g c
-t t0) on t1: P t t0 \def match t1 with [(tys3_nil u u0 t2) \Rightarrow (f u
-u0 t2) | (tys3_cons t2 u t3 ts t4) \Rightarrow (f0 t2 u t3 ts t4 ((tys3_ind g
-c P f f0) ts u t4))].
+implied let rec tys3_ind (g: G) (c: C) (P: (TList \to (T \to Prop))) (f:
+(\forall (u: T).(\forall (u0: T).((ty3 g c u u0) \to (P TNil u))))) (f0:
+(\forall (t: T).(\forall (u: T).((ty3 g c t u) \to (\forall (ts:
+TList).((tys3 g c ts u) \to ((P ts u) \to (P (TCons t ts) u)))))))) (t:
+TList) (t0: T) (t1: tys3 g c t t0) on t1: P t t0 \def match t1 with
+[(tys3_nil u u0 t2) \Rightarrow (f u u0 t2) | (tys3_cons t2 u t3 ts t4)
+\Rightarrow (f0 t2 u t3 ts t4 ((tys3_ind g c P f f0) ts u t4))].
-theorem ty3_gen_sort:
+lemma ty3_gen_sort:
\forall (g: G).(\forall (c: C).(\forall (x: T).(\forall (n: nat).((ty3 g c
(TSort n) x) \to (pc3 c (TSort (next g n)) x)))))
\def
\Rightarrow True])) I (TSort n) H5) in (False_ind (pc3 c0 (TSort (next g n))
(THead (Flat Cast) t0 t2)) H6))))))))))) c y x H0))) H))))).
-theorem ty3_gen_lref:
+lemma ty3_gen_lref:
\forall (g: G).(\forall (c: C).(\forall (x: T).(\forall (n: nat).((ty3 g c
(TLRef n) x) \to (or (ex3_3 C T T (\lambda (_: C).(\lambda (_: T).(\lambda
(t: T).(pc3 c (lift (S n) O t) x)))) (\lambda (e: C).(\lambda (u: T).(\lambda
(\lambda (e: C).(\lambda (u: T).(\lambda (t: T).(ty3 g e u t))))))
H6))))))))))) c y x H0))) H))))).
-theorem ty3_gen_bind:
+lemma ty3_gen_bind:
\forall (g: G).(\forall (b: B).(\forall (c: C).(\forall (u: T).(\forall (t1:
T).(\forall (x: T).((ty3 g c (THead (Bind b) u t1) x) \to (ex3_2 T T (\lambda
(t2: T).(\lambda (_: T).(pc3 c (THead (Bind b) u t2) x))) (\lambda (_:
(_: T).(\lambda (t: T).(ty3 g c0 u t))) (\lambda (t4: T).(\lambda (_: T).(ty3
g (CHead c0 (Bind b) u) t1 t4)))) H6))))))))))) c y x H0))) H))))))).
-theorem ty3_gen_appl:
+lemma ty3_gen_appl:
\forall (g: G).(\forall (c: C).(\forall (w: T).(\forall (v: T).(\forall (x:
T).((ty3 g c (THead (Flat Appl) w v) x) \to (ex3_2 T T (\lambda (u:
T).(\lambda (t: T).(pc3 c (THead (Flat Appl) w (THead (Bind Abst) u t)) x)))
t)))) (\lambda (u: T).(\lambda (_: T).(ty3 g c0 w u)))) H6))))))))))) c y x
H0))) H)))))).
-theorem ty3_gen_cast:
+lemma ty3_gen_cast:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t2: T).(\forall
(x: T).((ty3 g c (THead (Flat Cast) t2 t1) x) \to (ex3 T (\lambda (t0:
T).(pc3 c (THead (Flat Cast) t0 t2) x)) (\lambda (_: T).(ty3 g c t1 t2))
(THead (Flat Cast) t4 t2)) H14 H10))) t3 H8))))))) H6))))))))))) c y x H0)))
H)))))).
-theorem tys3_gen_nil:
+lemma tys3_gen_nil:
\forall (g: G).(\forall (c: C).(\forall (u: T).((tys3 g c TNil u) \to (ex T
(\lambda (u0: T).(ty3 g c u u0))))))
\def
True])) I TNil H4) in (False_ind (ex T (\lambda (u1: T).(ty3 g c u0 u1)))
H5))))))))) y u H0))) H)))).
-theorem tys3_gen_cons:
+lemma tys3_gen_cons:
\forall (g: G).(\forall (c: C).(\forall (ts: TList).(\forall (t: T).(\forall
(u: T).((tys3 g c (TCons t ts) u) \to (land (ty3 g c t u) (tys3 g c ts
u)))))))
include "basic_1/nf2/fwd.ma".
-theorem ty3_gen_appl_nf2:
+lemma ty3_gen_appl_nf2:
\forall (g: G).(\forall (c: C).(\forall (w: T).(\forall (v: T).(\forall (x:
T).((ty3 g c (THead (Flat Appl) w v) x) \to (ex4_2 T T (\lambda (u:
T).(\lambda (t: T).(pc3 c (THead (Flat Appl) w (THead (Bind Abst) u t)) x)))
(pc3_pr3_r c x0 x5 H13)) H15)))))))) H11))))) H8)))))) H5))))) H3))))))))
(ty3_gen_appl g c w v x H))))))).
-theorem ty3_inv_lref_nf2_pc3:
+lemma ty3_inv_lref_nf2_pc3:
\forall (g: G).(\forall (c: C).(\forall (u1: T).(\forall (i: nat).((ty3 g c
(TLRef i) u1) \to ((nf2 c (TLRef i)) \to (\forall (u2: T).((nf2 c u2) \to
((pc3 c u1 u2) \to (ex T (\lambda (u: T).(eq T u2 (lift (S i) O u))))))))))))
_ _) \Rightarrow True])) I (TLRef i) H5) in (False_ind (ex T (\lambda (u:
T).(eq T u2 (lift (S i) O u)))) H9))))))))))))))) c y u1 H0))) H))))).
-theorem ty3_inv_lref_nf2:
+lemma ty3_inv_lref_nf2:
\forall (g: G).(\forall (c: C).(\forall (u: T).(\forall (i: nat).((ty3 g c
(TLRef i) u) \to ((nf2 c (TLRef i)) \to ((nf2 c u) \to (ex T (\lambda (u0:
T).(eq T u (lift (S i) O u0))))))))))
(H: (ty3 g c (TLRef i) u)).(\lambda (H0: (nf2 c (TLRef i))).(\lambda (H1:
(nf2 c u)).(ty3_inv_lref_nf2_pc3 g c u i H H0 u H1 (pc3_refl c u)))))))).
-theorem ty3_inv_appls_lref_nf2:
+lemma ty3_inv_appls_lref_nf2:
\forall (g: G).(\forall (c: C).(\forall (vs: TList).(\forall (u1:
T).(\forall (i: nat).((ty3 g c (THeads (Flat Appl) vs (TLRef i)) u1) \to
((nf2 c (TLRef i)) \to ((nf2 c u1) \to (ex2 T (\lambda (u: T).(nf2 c (lift (S
(THeads (Flat Appl) t0 (lift (S i) O x)) (THead (Bind Abst) x0 x1) H13 t
Appl) u1 H4))))) H11))))) H8)))))))) H3))))))))))) vs))).
-theorem ty3_inv_lref_lref_nf2:
+lemma ty3_inv_lref_lref_nf2:
\forall (g: G).(\forall (c: C).(\forall (i: nat).(\forall (j: nat).((ty3 g c
(TLRef i) (TLRef j)) \to ((nf2 c (TLRef i)) \to ((nf2 c (TLRef j)) \to (lt i
j)))))))
TList).((pc3 c (THeads (Flat Appl) vs (lift (S i) O wi)) (THead (Bind Abst) w
u)) \to False)))))))).
-theorem ty3_nf2_inv_abst_premise_csort:
+lemma ty3_nf2_inv_abst_premise_csort:
\forall (w: T).(\forall (u: T).(\forall (m: nat).(ty3_nf2_inv_abst_premise
(CSort m) w u)))
\def
(lift (S i) O wi)) (THead (Bind Abst) w u))).(getl_gen_sort m i (CHead d
(Bind Abst) wi) H False))))))))).
-theorem ty3_nf2_inv_all:
+lemma ty3_nf2_inv_all:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (u: T).((ty3 g c t
u) \to ((nf2 c t) \to (or3 (ex3_2 T T (\lambda (w: T).(\lambda (u0: T).(eq T
t (THead (Bind Abst) w u0)))) (\lambda (w: T).(\lambda (_: T).(nf2 c w)))
(arity g c t x)).(\lambda (_: (arity g c u (asucc g x))).(arity_nf2_inv_all g
c t x H2 H0)))) H1)))))))).
-theorem ty3_nf2_inv_sort:
+lemma ty3_nf2_inv_sort:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (m: nat).((ty3 g c t
(TSort m)) \to ((nf2 c t) \to (or (ex2 nat (\lambda (n: nat).(eq T t (TSort
n))) (\lambda (n: nat).(eq nat m (next g n)))) (ex3_2 TList nat (\lambda (ws:
x0 x1 (refl_equal T (THeads (Flat Appl) x0 (TLRef x1))) H4 H5)) t H3)))))))
H2)) H1)))))))).
-theorem ty3_nf2_gen__ty3_nf2_inv_abst_aux:
+fact ty3_nf2_gen__ty3_nf2_inv_abst_aux:
\forall (c: C).(\forall (w1: T).(\forall (u1: T).((ty3_nf2_inv_abst_premise
c w1 u1) \to (\forall (t: T).(\forall (w2: T).(\forall (u2: T).((pc3 c (THead
(Flat Appl) t (THead (Bind Abst) w2 u2)) (THead (Bind Abst) w1 u1)) \to
vs (lift (S i) O wi)) (THead (Bind Abst) w2 u2) H2 t Appl) (THead (Bind Abst)
w1 u1) H0))))))))))))))).
-theorem ty3_nf2_inv_abst:
+lemma ty3_nf2_inv_abst:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (w: T).(\forall (u:
T).((ty3 g c t (THead (Bind Abst) w u)) \to ((nf2 c t) \to ((nf2 c w) \to
((ty3_nf2_inv_abst_premise c w u) \to (ex4_2 T T (\lambda (v: T).(\lambda (_:
include "basic_1/pc1/props.ma".
-theorem ty3_sred_wcpr0_pr0:
+lemma ty3_sred_wcpr0_pr0:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t: T).((ty3 g c1
t1 t) \to (\forall (c2: C).((wcpr0 c1 c2) \to (\forall (t2: T).((pr0 t1 t2)
\to (ty3 g c2 t2 t)))))))))
H9)) H8 H6)))]) in (H6 (refl_equal T (THead (Flat Cast) t3 t2)) (refl_equal T
t4))))))))))))))) c1 t1 t H))))).
-theorem ty3_sred_pr0:
+lemma ty3_sred_pr0:
\forall (t1: T).(\forall (t2: T).((pr0 t1 t2) \to (\forall (g: G).(\forall
(c: C).(\forall (t: T).((ty3 g c t1 t) \to (ty3 g c t2 t)))))))
\def
G).(\lambda (c: C).(\lambda (t: T).(\lambda (H0: (ty3 g c t1
t)).(ty3_sred_wcpr0_pr0 g c t1 t H0 c (wcpr0_refl c) t2 H))))))).
-theorem ty3_sred_pr1:
+lemma ty3_sred_pr1:
\forall (t1: T).(\forall (t2: T).((pr1 t1 t2) \to (\forall (g: G).(\forall
(c: C).(\forall (t: T).((ty3 g c t1 t) \to (ty3 g c t2 t)))))))
\def
G).(\lambda (c: C).(\lambda (t: T).(\lambda (H3: (ty3 g c t4 t)).(H2 g c t
(ty3_sred_pr0 t4 t3 H0 g c t H3)))))))))))) t1 t2 H))).
-theorem ty3_sred_pr2:
+lemma ty3_sred_pr2:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr2 c t1 t2) \to (\forall
(g: G).(\forall (t: T).((ty3 g c t1 t) \to (ty3 g c t2 t)))))))
\def
(ty3_sred_wcpr0_pr0 g c0 t3 t0 H3 c0 (wcpr0_refl c0) t4 H1) d u i H0 t
H2)))))))))))))) c t1 t2 H)))).
-theorem ty3_sred_pr3:
+lemma ty3_sred_pr3:
\forall (c: C).(\forall (t1: T).(\forall (t2: T).((pr3 c t1 t2) \to (\forall
(g: G).(\forall (t: T).((ty3 g c t1 t) \to (ty3 g c t2 t)))))))
\def
include "basic_1/ty3/pr3.ma".
-theorem ty3_cred_pr2:
+lemma ty3_cred_pr2:
\forall (g: G).(\forall (c: C).(\forall (v1: T).(\forall (v2: T).((pr2 c v1
v2) \to (\forall (b: B).(\forall (t1: T).(\forall (t2: T).((ty3 g (CHead c
(Bind b) v1) t1 t2) \to (ty3 g (CHead c (Bind b) v2) t1 t2)))))))))
(CHead c0 (Bind b) t) (csubst0_snd_bind b i u t2 t H2 c0)))))))))))))))) c v1
v2 H))))).
-theorem ty3_cred_pr3:
+lemma ty3_cred_pr3:
\forall (g: G).(\forall (c: C).(\forall (v1: T).(\forall (v2: T).((pr3 c v1
v2) \to (\forall (b: B).(\forall (t1: T).(\forall (t2: T).((ty3 g (CHead c
(Bind b) v1) t1 t2) \to (ty3 g (CHead c (Bind b) v2) t1 t2)))))))))
T).(\lambda (t4: T).(\lambda (H3: (ty3 g (CHead c (Bind b) t1) t0 t4)).(H2 b
t0 t4 (ty3_cred_pr2 g c t1 t2 H0 b t0 t4 H3)))))))))))) v1 v2 H))))).
-theorem ty3_gen_lift:
+lemma ty3_gen_lift:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (x: T).(\forall (h:
nat).(\forall (d: nat).((ty3 g c (lift h d t1) x) \to (\forall (e: C).((drop
h d c e) \to (ex2 T (\lambda (t2: T).(pc3 c (lift h d t2) x)) (\lambda (t2:
H19))))) H16)))) t3 H8))))) x0 H7)))))) (lift_gen_flat Cast t3 t2 x0 h x1
H5))))))))))))))) c y x H0))))) H))))))).
-theorem ty3_tred:
+lemma ty3_tred:
\forall (g: G).(\forall (c: C).(\forall (u: T).(\forall (t1: T).((ty3 g c u
t1) \to (\forall (t2: T).((pr3 c t1 t2) \to (ty3 g c u t2)))))))
\def
(ty3_sred_pr3 c u2 x H4 g t2 H0) in (let H_y0 \def (ty3_sred_pr3 c u1 x H3 g
t1 H) in (ty3_unique g c x t1 H_y0 t2 H_y)))))) H2)))))))))).
-theorem ty3_sred_back:
+lemma ty3_sred_back:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t0: T).((ty3 g c
t1 t0) \to (\forall (t2: T).((pr3 c t1 t2) \to (\forall (t: T).((ty3 g c t2
t) \to (ty3 g c t1 t)))))))))
include "basic_1/pc3/fwd.ma".
-theorem ty3_lift:
+lemma ty3_lift:
\forall (g: G).(\forall (e: C).(\forall (t1: T).(\forall (t2: T).((ty3 g e
t1 t2) \to (\forall (c: C).(\forall (d: nat).(\forall (h: nat).((drop h d c
e) \to (ty3 g c (lift h d t1) (lift h d t2))))))))))
(lift h d (THead (Flat Cast) t3 t0)) (lift_head (Flat Cast) t3 t0 h
d)))))))))))))) e t1 t2 H))))).
-theorem ty3_correct:
+lemma ty3_correct:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t2: T).((ty3 g c
t1 t2) \to (ex T (\lambda (t: T).(ty3 g c t2 t)))))))
\def
(THead (Flat Cast) t3 t2) (pc3_head_1 c0 t3 x0 (H3 x0 H7) (Flat Cast) t2) t4
H5))))) (ty3_gen_cast g c0 t0 t2 t4 H4)))))))))))) c u t1 H))))).
-theorem ty3_gen_abst_abst:
+lemma ty3_gen_abst_abst:
\forall (g: G).(\forall (c: C).(\forall (u: T).(\forall (t1: T).(\forall
(t2: T).((ty3 g c (THead (Bind Abst) u t1) (THead (Bind Abst) u t2)) \to (ex2
T (\lambda (w: T).(ty3 g c u w)) (\lambda (_: T).(ty3 g (CHead c (Bind Abst)
Abst c u t2 x H0)))) (ty3_correct g c (THead (Bind Abst) u t1) (THead (Bind
Abst) u t2) H))))))).
-theorem ty3_typecheck:
+lemma ty3_typecheck:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (v: T).((ty3 g c t
v) \to (ex T (\lambda (u: T).(ty3 g c (THead (Flat Cast) v t) u)))))))
\def
(THead (Flat Cast) x v) (ty3_cast g c t v H x H0)))) (ty3_correct g c t v
H)))))).
-theorem ty3_getl_subst0:
+lemma ty3_getl_subst0:
\forall (g: G).(\forall (c: C).(\forall (t: T).(\forall (u: T).((ty3 g c t
u) \to (\forall (v0: T).(\forall (t0: T).(\forall (i: nat).((subst0 i v0 t
t0) \to (\forall (b: B).(\forall (d: C).(\forall (v: T).((getl i c (CHead d
include "basic_1/sty0/fwd.ma".
-theorem ty3_sty0:
+lemma ty3_sty0:
\forall (g: G).(\forall (c: C).(\forall (u: T).(\forall (t1: T).((ty3 g c u
t1) \to (\forall (t2: T).((sty0 g c u t2) \to (ty3 g c u t2)))))))
\def
include "basic_1/getl/getl.ma".
-theorem ty3_gen_cabbr:
+lemma ty3_gen_cabbr:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t2: T).((ty3 g c
t1 t2) \to (\forall (e: C).(\forall (u: T).(\forall (d: nat).((getl d c
(CHead e (Bind Abbr) u)) \to (\forall (a0: C).((csubst1 d u c a0) \to
Cast) x1 x0)) (lift_flat Cast x1 x0 (S O) d)) (ty3_cast g a x2 x0 H15 x1
H10)))))))) H11))))))) H7)))))))))))))))))) c t1 t2 H))))).
-theorem ty3_gen_cvoid:
+lemma ty3_gen_cvoid:
\forall (g: G).(\forall (c: C).(\forall (t1: T).(\forall (t2: T).((ty3 g c
t1 t2) \to (\forall (e: C).(\forall (u: T).(\forall (d: nat).((getl d c
(CHead e (Bind Void) u)) \to (\forall (a: C).((drop (S O) d c a) \to (ex3_2 T
include "basic_1/wcpr0/defs.ma".
-let rec wcpr0_ind (P: (C \to (C \to Prop))) (f: (\forall (c: C).(P c c)))
-(f0: (\forall (c1: C).(\forall (c2: C).((wcpr0 c1 c2) \to ((P c1 c2) \to
+implied let rec wcpr0_ind (P: (C \to (C \to Prop))) (f: (\forall (c: C).(P c
+c))) (f0: (\forall (c1: C).(\forall (c2: C).((wcpr0 c1 c2) \to ((P c1 c2) \to
(\forall (u1: T).(\forall (u2: T).((pr0 u1 u2) \to (\forall (k: K).(P (CHead
c1 k u1) (CHead c2 k u2))))))))))) (c: C) (c0: C) (w: wcpr0 c c0) on w: P c
c0 \def match w with [(wcpr0_refl c1) \Rightarrow (f c1) | (wcpr0_comp c1 c2
w0 u1 u2 p k) \Rightarrow (f0 c1 c2 w0 ((wcpr0_ind P f f0) c1 c2 w0) u1 u2 p
k)].
-theorem wcpr0_gen_sort:
+lemma wcpr0_gen_sort:
\forall (x: C).(\forall (n: nat).((wcpr0 (CSort n) x) \to (eq C x (CSort
n))))
\def
| (CHead _ _ _) \Rightarrow True])) I (CSort n) H4) in (False_ind (eq C
(CHead c2 k u2) (CHead c1 k u1)) H5))))))))))) y x H0))) H))).
-theorem wcpr0_gen_head:
+lemma wcpr0_gen_head:
\forall (k: K).(\forall (c1: C).(\forall (x: C).(\forall (u1: T).((wcpr0
(CHead c1 k u1) x) \to (or (eq C x (CHead c1 k u1)) (ex3_2 C T (\lambda (c2:
C).(\lambda (u2: T).(eq C x (CHead c2 k u2)))) (\lambda (c2: C).(\lambda (_:
include "basic_1/getl/props.ma".
-theorem wcpr0_drop:
+lemma wcpr0_drop:
\forall (c1: C).(\forall (c2: C).((wcpr0 c1 c2) \to (\forall (h:
nat).(\forall (e1: C).(\forall (u1: T).(\forall (k: K).((drop h O c1 (CHead
e1 k u1)) \to (ex3_2 C T (\lambda (e2: C).(\lambda (u2: T).(drop h O c2
(Flat f) n c4 (CHead x0 k0 x1) H6 u2) H7 H8)))))) H5))))))))) k) h))))))))))
c1 c2 H))).
-theorem wcpr0_drop_back:
+lemma wcpr0_drop_back:
\forall (c1: C).(\forall (c2: C).((wcpr0 c2 c1) \to (\forall (h:
nat).(\forall (e1: C).(\forall (u1: T).(\forall (k: K).((drop h O c1 (CHead
e1 k u1)) \to (ex3_2 C T (\lambda (e2: C).(\lambda (u2: T).(drop h O c2
(Flat f) n c3 (CHead x0 k0 x1) H6 u1) H7 H8)))))) H5))))))))) k) h))))))))))
c2 c1 H))).
-theorem wcpr0_getl:
+lemma wcpr0_getl:
\forall (c1: C).(\forall (c2: C).((wcpr0 c1 c2) \to (\forall (h:
nat).(\forall (e1: C).(\forall (u1: T).(\forall (k: K).((getl h c1 (CHead e1
k u1)) \to (ex3_2 C T (\lambda (e2: C).(\lambda (u2: T).(getl h c2 (CHead e2
f) n c4 (CHead x0 k0 x1) H6 u2) H7 H8)))))) H5))))))))) k) h)))))))))) c1 c2
H))).
-theorem wcpr0_getl_back:
+lemma wcpr0_getl_back:
\forall (c1: C).(\forall (c2: C).((wcpr0 c2 c1) \to (\forall (h:
nat).(\forall (e1: C).(\forall (u1: T).(\forall (k: K).((getl h c1 (CHead e1
k u1)) \to (ex3_2 C T (\lambda (e2: C).(\lambda (u2: T).(getl h c2 (CHead e2
include "basic_1/clear/fwd.ma".
-theorem wf3_clear_conf:
+lemma wf3_clear_conf:
\forall (c1: C).(\forall (c: C).((clear c1 c) \to (\forall (g: G).(\forall
(c2: C).((wf3 g c1 c2) \to (wf3 g c c2))))))
\def
C).(\lambda (H2: (wf3 g (CHead e (Flat f) u) c2)).(let H_y \def
(wf3_gen_flat1 g e c2 u f H2) in (H1 g c2 H_y))))))))))) c1 c H))).
-theorem clear_wf3_trans:
+lemma clear_wf3_trans:
\forall (c1: C).(\forall (d1: C).((clear c1 d1) \to (\forall (g: G).(\forall
(d2: C).((wf3 g d1 d2) \to (ex2 C (\lambda (c2: C).(wf3 g c1 c2)) (\lambda
(c2: C).(clear c2 d2))))))))
include "basic_1/wf3/defs.ma".
-let rec wf3_ind (g: G) (P: (C \to (C \to Prop))) (f: (\forall (m: nat).(P
-(CSort m) (CSort m)))) (f0: (\forall (c1: C).(\forall (c2: C).((wf3 g c1 c2)
-\to ((P c1 c2) \to (\forall (u: T).(\forall (t: T).((ty3 g c1 u t) \to
+implied let rec wf3_ind (g: G) (P: (C \to (C \to Prop))) (f: (\forall (m:
+nat).(P (CSort m) (CSort m)))) (f0: (\forall (c1: C).(\forall (c2: C).((wf3 g
+c1 c2) \to ((P c1 c2) \to (\forall (u: T).(\forall (t: T).((ty3 g c1 u t) \to
(\forall (b: B).(P (CHead c1 (Bind b) u) (CHead c2 (Bind b) u))))))))))) (f1:
(\forall (c1: C).(\forall (c2: C).((wf3 g c1 c2) \to ((P c1 c2) \to (\forall
(u: T).(((\forall (t: T).((ty3 g c1 u t) \to False))) \to (\forall (b: B).(P
((wf3_ind g P f f0 f1 f2) c1 c2 w0) u f3 b) | (wf3_flat c1 c2 w0 u f3)
\Rightarrow (f2 c1 c2 w0 ((wf3_ind g P f f0 f1 f2) c1 c2 w0) u f3)].
-theorem wf3_gen_sort1:
+lemma wf3_gen_sort1:
\forall (g: G).(\forall (x: C).(\forall (m: nat).((wf3 g (CSort m) x) \to
(eq C x (CSort m)))))
\def
[(CSort _) \Rightarrow False | (CHead _ _ _) \Rightarrow True])) I (CSort m)
H3) in (False_ind (eq C c2 (CHead c1 (Flat f) u)) H4))))))))) y x H0))) H)))).
-theorem wf3_gen_bind1:
+lemma wf3_gen_bind1:
\forall (g: G).(\forall (c1: C).(\forall (x: C).(\forall (v: T).(\forall (b:
B).((wf3 g (CHead c1 (Bind b) v) x) \to (or (ex3_2 C T (\lambda (c2:
C).(\lambda (_: T).(eq C x (CHead c2 (Bind b) v)))) (\lambda (c2: C).(\lambda
(\lambda (_: C).(\forall (w: T).((ty3 g c1 v w) \to False))))) H4))))))))) y
x H0))) H)))))).
-theorem wf3_gen_flat1:
+lemma wf3_gen_flat1:
\forall (g: G).(\forall (c1: C).(\forall (x: C).(\forall (v: T).(\forall (f:
F).((wf3 g (CHead c1 (Flat f) v) x) \to (wf3 g c1 x))))))
\def
\def (eq_ind C c0 (\lambda (c: C).(wf3 g c c2)) H1 c1 H8) in H10))))) H5))
H4))))))))) y x H0))) H)))))).
-theorem wf3_gen_head2:
+lemma wf3_gen_head2:
\forall (g: G).(\forall (x: C).(\forall (c: C).(\forall (v: T).(\forall (k:
K).((wf3 g x (CHead c k v)) \to (ex B (\lambda (b: B).(eq K k (Bind b)))))))))
\def
include "basic_1/ty3/dec.ma".
-theorem wf3_getl_conf:
+lemma wf3_getl_conf:
\forall (b: B).(\forall (i: nat).(\forall (c1: C).(\forall (d1: C).(\forall
(v: T).((getl i c1 (CHead d1 (Bind b) v)) \to (\forall (g: G).(\forall (c2:
C).((wf3 g c1 c2) \to (\forall (w: T).((ty3 g d1 v w) \to (ex2 C (\lambda
v H5 g c2 H_y w H3))))) k H2 (getl_gen_S k c (CHead d1 (Bind b) v) t n
H1)))))))))))))) c1)))) i)).
-theorem getl_wf3_trans:
+lemma getl_wf3_trans:
\forall (i: nat).(\forall (c1: C).(\forall (d1: C).((getl i c1 d1) \to
(\forall (g: G).(\forall (d2: C).((wf3 g d1 d2) \to (ex2 C (\lambda (c2:
C).(wf3 g c1 c2)) (\lambda (c2: C).(getl i c2 d2)))))))))
include "basic_1/app/defs.ma".
-theorem wf3_total:
+lemma wf3_total:
\forall (g: G).(\forall (c1: C).(ex C (\lambda (c2: C).(wf3 g c1 c2))))
\def
\lambda (g: G).(\lambda (c1: C).(C_ind (\lambda (c: C).(ex C (\lambda (c2:
(f: F).(ex_intro C (\lambda (c2: C).(wf3 g (CHead c (Flat f) t) c2)) x
(wf3_flat g c x H1 t f))) k))) H0)))))) c1)).
-theorem ty3_shift1:
+lemma ty3_shift1:
\forall (g: G).(\forall (c: C).((wf3 g c c) \to (\forall (t1: T).(\forall
(t2: T).((ty3 g c t1 t2) \to (ty3 g (CSort (cbk c)) (app1 c t1) (app1 c
t2)))))))
(CSort (cbk c1)) (app1 c1 (THead (Flat f) u t1)) (app1 c1 (THead (Flat f) u
t2))) H10)))) H8)))))))))))))))) y c H0))) H))).
-theorem wf3_idem:
+lemma wf3_idem:
\forall (g: G).(\forall (c1: C).(\forall (c2: C).((wf3 g c1 c2) \to (wf3 g
c2 c2))))
\def
(c3: C).(\lambda (c4: C).(\lambda (_: (wf3 g c3 c4)).(\lambda (H1: (wf3 g c4
c4)).(\lambda (_: T).(\lambda (_: F).H1)))))) c1 c2 H)))).
-theorem wf3_ty3:
+lemma wf3_ty3:
\forall (g: G).(\forall (c1: C).(\forall (t: T).(\forall (u: T).((ty3 g c1 t
u) \to (ex2 C (\lambda (c2: C).(wf3 g c1 c2)) (\lambda (c2: C).(ty3 g c2 t
u)))))))
include "basic_1/wf3/getl.ma".
-theorem wf3_pr2_conf:
+lemma wf3_pr2_conf:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t2: T).((pr2 c1
t1 t2) \to (\forall (c2: C).((wf3 g c1 c2) \to (\forall (u: T).((ty3 g c1 t1
u) \to (pr2 c2 t1 t2)))))))))
(_: (wf3 g d x0)).(pr2_delta c2 x0 u i H8 t3 t4 H1 t H2)))) H7)))))
H5)))))))))))))))))) c1 t1 t2 H))))).
-theorem wf3_pr3_conf:
+lemma wf3_pr3_conf:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t2: T).((pr3 c1
t1 t2) \to (\forall (c2: C).((wf3 g c1 c2) \to (\forall (u: T).((ty3 g c1 t1
u) \to (pr3 c2 t1 t2)))))))))
t4 u)).(pr3_sing c2 t3 t4 (wf3_pr2_conf g c1 t4 t3 H0 c2 H3 u H4) t5 (H2 c2
H3 u (ty3_sred_pr2 c1 t4 t3 H0 g u H4))))))))))))) t1 t2 H))))).
-theorem wf3_pc3_conf:
+lemma wf3_pc3_conf:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t2: T).((pc3 c1
t1 t2) \to (\forall (c2: C).((wf3 g c1 c2) \to (\forall (u1: T).((ty3 g c1 t1
u1) \to (\forall (u2: T).((ty3 g c1 t2 u2) \to (pc3 c2 t1 t2)))))))))))
g c1 t1 x H4 c2 H0 u1 H1) t2 (wf3_pr3_conf g c1 t2 x H5 c2 H0 u2 H2)))))
H3)))))))))))).
-theorem wf3_ty3_conf:
+lemma wf3_ty3_conf:
\forall (g: G).(\forall (c1: C).(\forall (t1: T).(\forall (t2: T).((ty3 g c1
t1 t2) \to (\forall (c2: C).((wf3 g c1 c2) \to (ty3 g c2 t1 t2)))))))
\def
include "ground_1/blt/defs.ma".
-theorem lt_blt:
+lemma lt_blt:
\forall (x: nat).(\forall (y: nat).((lt y x) \to (eq bool (blt y x) true)))
\def
\lambda (x: nat).(nat_ind (\lambda (n: nat).(\forall (y: nat).((lt y n) \to
with [O \Rightarrow true | (S m) \Rightarrow (blt m n)]) true)))).(\lambda
(H1: (lt (S n0) (S n))).(H n0 (le_S_n (S n0) n H1))))) y)))) x).
-theorem le_bge:
+lemma le_bge:
\forall (x: nat).(\forall (y: nat).((le x y) \to (eq bool (blt y x) false)))
\def
\lambda (x: nat).(nat_ind (\lambda (n: nat).(\forall (y: nat).((le n y) \to
(eq bool (blt n0 (S n)) false)))).(\lambda (H1: (le (S n) (S n0))).(H n0
(le_S_n n n0 H1))))) y)))) x).
-theorem blt_lt:
+lemma blt_lt:
\forall (x: nat).(\forall (y: nat).((eq bool (blt y x) true) \to (lt y x)))
\def
\lambda (x: nat).(nat_ind (\lambda (n: nat).(\forall (y: nat).((eq bool (blt
\Rightarrow (blt m n)]) true) \to (lt n0 (S n))))).(\lambda (H1: (eq bool
(blt n0 n) true)).(lt_n_S n0 n (H n0 H1))))) y)))) x).
-theorem bge_le:
+lemma bge_le:
\forall (x: nat).(\forall (y: nat).((eq bool (blt y x) false) \to (le x y)))
\def
\lambda (x: nat).(nat_ind (\lambda (n: nat).(\forall (y: nat).((eq bool (blt
include "ground_1/preamble.ma".
-theorem nat_dec:
+lemma nat_dec:
\forall (n1: nat).(\forall (n2: nat).(or (eq nat n1 n2) ((eq nat n1 n2) \to
(\forall (P: Prop).P))))
\def
(eq nat (S n) n3) ((eq nat (S n) n3) \to (\forall (P0: Prop).P0)))) H0 n H3)
in (H4 (refl_equal nat n) P)))))))) (H n0)))) n2)))) n1).
-theorem simpl_plus_r:
+lemma simpl_plus_r:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((eq nat (plus m n)
(plus p n)) \to (eq nat m p))))
\def
nat).(eq nat n0 (plus n p))) (plus_sym p n) (plus m n) H) (plus n m)
(plus_sym n m)))))).
-theorem minus_Sx_Sy:
+lemma minus_Sx_Sy:
\forall (x: nat).(\forall (y: nat).(eq nat (minus (S x) (S y)) (minus x y)))
\def
\lambda (x: nat).(\lambda (y: nat).(refl_equal nat (minus x y))).
-theorem minus_plus_r:
+lemma minus_plus_r:
\forall (m: nat).(\forall (n: nat).(eq nat (minus (plus m n) n) m))
\def
\lambda (m: nat).(\lambda (n: nat).(eq_ind_r nat (plus n m) (\lambda (n0:
nat).(eq nat (minus n0 n) m)) (minus_plus n m) (plus m n) (plus_sym m n))).
-theorem plus_permute_2_in_3:
+lemma plus_permute_2_in_3:
\forall (x: nat).(\forall (y: nat).(\forall (z: nat).(eq nat (plus (plus x
y) z) (plus (plus x z) y))))
\def
(refl_equal nat (plus (plus x z) y)) (plus x (plus z y)) (plus_assoc_r x z
y)) (plus y z) (plus_sym y z)) (plus (plus x y) z) (plus_assoc_r x y z)))).
-theorem plus_permute_2_in_3_assoc:
+lemma plus_permute_2_in_3_assoc:
\forall (n: nat).(\forall (h: nat).(\forall (k: nat).(eq nat (plus (plus n
h) k) (plus n (plus k h)))))
\def
(refl_equal nat (plus (plus n k) h)) (plus n (plus k h)) (plus_assoc_l n k
h)) (plus (plus n h) k) (plus_permute_2_in_3 n h k)))).
-theorem plus_O:
+lemma plus_O:
\forall (x: nat).(\forall (y: nat).((eq nat (plus x y) O) \to (land (eq nat
x O) (eq nat y O))))
\def
True])) I O H1) in (False_ind (land (eq nat (S n) O) (eq nat y O)) H2)))]) in
(H1 (refl_equal nat O))))))) x).
-theorem minus_Sx_SO:
+lemma minus_Sx_SO:
\forall (x: nat).(eq nat (minus (S x) (S O)) x)
\def
\lambda (x: nat).(eq_ind nat x (\lambda (n: nat).(eq nat n x)) (refl_equal
nat x) (minus x O) (minus_n_O x)).
-theorem nat_dec_neg:
+lemma nat_dec_neg:
\forall (i: nat).(\forall (j: nat).(or (not (eq nat i j)) (eq nat i j)))
\def
\lambda (i: nat).(nat_ind (\lambda (n: nat).(\forall (j: nat).(or (not (eq
(eq nat (S n) (S n0))) (eq nat (S n) (S n0)) (f_equal nat nat S n n0 H1))) (H
n0)))) j)))) i).
-theorem neq_eq_e:
+lemma neq_eq_e:
\forall (i: nat).(\forall (j: nat).(\forall (P: Prop).((((not (eq nat i j))
\to P)) \to ((((eq nat i j) \to P)) \to P))))
\def
(eq nat i j)) \to P))).(\lambda (H0: (((eq nat i j) \to P))).(let o \def
(nat_dec_neg i j) in (or_ind (not (eq nat i j)) (eq nat i j) P H H0 o)))))).
-theorem le_false:
+lemma le_false:
\forall (m: nat).(\forall (n: nat).(\forall (P: Prop).((le m n) \to ((le (S
n) m) \to P))))
\def
(S n1))).(\lambda (H2: (le (S (S n1)) (S n))).(H n1 P (le_S_n n n1 H1)
(le_S_n (S n1) n H2))))))) n0)))) m).
-theorem le_Sx_x:
+lemma le_Sx_x:
\forall (x: nat).((le (S x) x) \to (\forall (P: Prop).P))
\def
\lambda (x: nat).(\lambda (H: (le (S x) x)).(\lambda (P: Prop).(let H0 \def
le_Sn_n in (False_ind P (H0 x H))))).
-theorem le_n_pred:
+lemma le_n_pred:
\forall (n: nat).(\forall (m: nat).((le n m) \to (le (pred n) (pred m))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (le n m)).(le_ind n (\lambda
nat).(\lambda (_: (le n m0)).(\lambda (H1: (le (pred n) (pred m0))).(le_trans
(pred n) (pred m0) m0 H1 (le_pred_n m0))))) m H))).
-theorem minus_le:
+lemma minus_le:
\forall (x: nat).(\forall (y: nat).(le (minus x y) x))
\def
\lambda (x: nat).(nat_ind (\lambda (n: nat).(\forall (y: nat).(le (minus n
nat).(\lambda (_: (le (match n0 with [O \Rightarrow (S n) | (S l) \Rightarrow
(minus n l)]) (S n))).(le_S (minus n n0) n (H n0)))) y)))) x).
-theorem le_plus_minus_sym:
+lemma le_plus_minus_sym:
\forall (n: nat).(\forall (m: nat).((le n m) \to (eq nat m (plus (minus m n)
n))))
\def
(plus n (minus m n)) (\lambda (n0: nat).(eq nat m n0)) (le_plus_minus n m H)
(plus (minus m n) n) (plus_sym (minus m n) n)))).
-theorem le_minus_minus:
+lemma le_minus_minus:
\forall (x: nat).(\forall (y: nat).((le x y) \to (\forall (z: nat).((le y z)
\to (le (minus y x) (minus z x))))))
\def
z (\lambda (n: nat).(le y n)) H0 (plus x (minus z x)) (le_plus_minus_r x z
(le_trans x y z H H0))) (plus x (minus y x)) (le_plus_minus_r x y H))))))).
-theorem le_minus_plus:
+lemma le_minus_plus:
\forall (z: nat).(\forall (x: nat).((le z x) \to (\forall (y: nat).(eq nat
(minus (plus x y) z) (plus (minus x z) y)))))
\def
n y) (S z0)) (plus (minus n (S z0)) y)))))).(\lambda (H1: (le (S z0) (S
n))).(\lambda (y: nat).(H n (le_S_n z0 n H1) y))))) x)))) z).
-theorem le_minus:
+lemma le_minus:
\forall (x: nat).(\forall (z: nat).(\forall (y: nat).((le (plus x y) z) \to
(le x (minus z y)))))
\def
y))) (le_minus_minus y (plus x y) (le_plus_r x y) z H) x (minus_plus_r x
y))))).
-theorem le_trans_plus_r:
+lemma le_trans_plus_r:
\forall (x: nat).(\forall (y: nat).(\forall (z: nat).((le (plus x y) z) \to
(le y z))))
\def
\lambda (x: nat).(\lambda (y: nat).(\lambda (z: nat).(\lambda (H: (le (plus
x y) z)).(le_trans y (plus x y) z (le_plus_r x y) H)))).
-theorem lt_x_O:
+lemma lt_x_O:
\forall (x: nat).((lt x O) \to (\forall (P: Prop).P))
\def
\lambda (x: nat).(\lambda (H: (le (S x) O)).(\lambda (P: Prop).(let H_y \def
ee with [O \Rightarrow True | (S _) \Rightarrow False])) I (S x) H_y) in
(False_ind P H0))))).
-theorem le_gen_S:
+lemma le_gen_S:
\forall (m: nat).(\forall (x: nat).((le (S m) x) \to (ex2 nat (\lambda (n:
nat).(eq nat x (S n))) (\lambda (n: nat).(le m n)))))
\def
nat).(le m n)) m0 (refl_equal nat (S m0)) (le_S_n m m0 (le_S (S m) m0 H2))))
x H1 H0))]) in (H0 (refl_equal nat x))))).
-theorem lt_x_plus_x_Sy:
+lemma lt_x_plus_x_Sy:
\forall (x: nat).(\forall (y: nat).(lt x (plus x (S y))))
\def
\lambda (x: nat).(\lambda (y: nat).(eq_ind_r nat (plus (S y) x) (\lambda (n:
nat).(lt x n)) (le_S_n (S x) (S (plus y x)) (le_n_S (S x) (S (plus y x))
(le_n_S x (plus y x) (le_plus_r y x)))) (plus x (S y)) (plus_sym x (S y)))).
-theorem simpl_lt_plus_r:
+lemma simpl_lt_plus_r:
\forall (p: nat).(\forall (n: nat).(\forall (m: nat).((lt (plus n p) (plus m
p)) \to (lt n m))))
\def
H1 \def (eq_ind nat (plus m p) (\lambda (n0: nat).(lt (plus p n) n0)) H0
(plus p m) (plus_sym m p)) in H1)))))).
-theorem minus_x_Sy:
+lemma minus_x_Sy:
\forall (x: nat).(\forall (y: nat).((lt y x) \to (eq nat (minus x y) (S
(minus x (S y))))))
\def
(H1: (lt (S n0) (S n))).(let H2 \def (le_S_n (S n0) n H1) in (H n0 H2)))))
y)))) x).
-theorem lt_plus_minus:
+lemma lt_plus_minus:
\forall (x: nat).(\forall (y: nat).((lt x y) \to (eq nat y (S (plus x (minus
y (S x)))))))
\def
\lambda (x: nat).(\lambda (y: nat).(\lambda (H: (lt x y)).(le_plus_minus (S
x) y H))).
-theorem lt_plus_minus_r:
+lemma lt_plus_minus_r:
\forall (x: nat).(\forall (y: nat).((lt x y) \to (eq nat y (S (plus (minus y
(S x)) x)))))
\def
(plus x (minus y (S x))) (\lambda (n: nat).(eq nat y (S n))) (lt_plus_minus x
y H) (plus (minus y (S x)) x) (plus_sym (minus y (S x)) x)))).
-theorem minus_x_SO:
+lemma minus_x_SO:
\forall (x: nat).((lt O x) \to (eq nat x (S (minus x (S O)))))
\def
\lambda (x: nat).(\lambda (H: (lt O x)).(eq_ind nat (minus x O) (\lambda (n:
nat).(eq nat x n)) (eq_ind nat x (\lambda (n: nat).(eq nat x n)) (refl_equal
nat x) (minus x O) (minus_n_O x)) (S (minus x (S O))) (minus_x_Sy x O H))).
-theorem le_x_pred_y:
+lemma le_x_pred_y:
\forall (y: nat).(\forall (x: nat).((lt x y) \to (le x (pred y))))
\def
\lambda (y: nat).(nat_ind (\lambda (n: nat).(\forall (x: nat).((lt x n) \to
x n) \to (le x (pred n)))))).(\lambda (x: nat).(\lambda (H0: (lt x (S
n))).(le_S_n x n H0))))) y).
-theorem lt_le_minus:
+lemma lt_le_minus:
\forall (x: nat).(\forall (y: nat).((lt x y) \to (le x (minus y (S O)))))
\def
\lambda (x: nat).(\lambda (y: nat).(\lambda (H: (lt x y)).(le_minus x y (S
O) (eq_ind_r nat (plus (S O) x) (\lambda (n: nat).(le n y)) H (plus x (S O))
(plus_sym x (S O)))))).
-theorem lt_le_e:
+lemma lt_le_e:
\forall (n: nat).(\forall (d: nat).(\forall (P: Prop).((((lt n d) \to P))
\to ((((le d n) \to P)) \to P))))
\def
d) \to P))).(\lambda (H0: (((le d n) \to P))).(let H1 \def (le_or_lt d n) in
(or_ind (le d n) (lt n d) P H0 H H1)))))).
-theorem lt_eq_e:
+lemma lt_eq_e:
\forall (x: nat).(\forall (y: nat).(\forall (P: Prop).((((lt x y) \to P))
\to ((((eq nat x y) \to P)) \to ((le x y) \to P)))))
\def
y) \to P))).(\lambda (H0: (((eq nat x y) \to P))).(\lambda (H1: (le x
y)).(or_ind (lt x y) (eq nat x y) P H H0 (le_lt_or_eq x y H1))))))).
-theorem lt_eq_gt_e:
+lemma lt_eq_gt_e:
\forall (x: nat).(\forall (y: nat).(\forall (P: Prop).((((lt x y) \to P))
\to ((((eq nat x y) \to P)) \to ((((lt y x) \to P)) \to P)))))
\def
\to P))).(lt_le_e x y P H (\lambda (H2: (le y x)).(lt_eq_e y x P H1 (\lambda
(H3: (eq nat y x)).(H0 (sym_eq nat y x H3))) H2)))))))).
-theorem lt_gen_xS:
+lemma lt_gen_xS:
\forall (x: nat).(\forall (n: nat).((lt x (S n)) \to (or (eq nat x O) (ex2
nat (\lambda (m: nat).(eq nat x (S m))) (\lambda (m: nat).(lt m n))))))
\def
(ex_intro2 nat (\lambda (m: nat).(eq nat (S n) (S m))) (\lambda (m: nat).(lt
m n0)) n (refl_equal nat (S n)) (le_S_n (S n) n0 H0))))))) x).
-theorem le_lt_false:
+lemma le_lt_false:
\forall (x: nat).(\forall (y: nat).((le x y) \to ((lt y x) \to (\forall (P:
Prop).P))))
\def
\lambda (x: nat).(\lambda (y: nat).(\lambda (H: (le x y)).(\lambda (H0: (lt
y x)).(\lambda (P: Prop).(False_ind P (le_not_lt x y H H0)))))).
-theorem lt_neq:
+lemma lt_neq:
\forall (x: nat).(\forall (y: nat).((lt x y) \to (not (eq nat x y))))
\def
\lambda (x: nat).(\lambda (y: nat).(\lambda (H: (lt x y)).(\lambda (H0: (eq
nat x y)).(let H1 \def (eq_ind nat x (\lambda (n: nat).(lt n y)) H y H0) in
(lt_n_n y H1))))).
-theorem arith0:
+lemma arith0:
\forall (h2: nat).(\forall (d2: nat).(\forall (n: nat).((le (plus d2 h2) n)
\to (\forall (h1: nat).(le (plus d2 h1) (minus (plus n h1) h2))))))
\def
d2)) (plus h2 (plus d2 h1)) (plus_assoc_l h2 d2 h1))) (plus d2 h1)
(minus_plus h2 (plus d2 h1))))))).
-theorem O_minus:
+lemma O_minus:
\forall (x: nat).(\forall (y: nat).((le x y) \to (eq nat (minus x y) O)))
\def
\lambda (x: nat).(nat_ind (\lambda (n: nat).(\forall (y: nat).((le n y) \to
x0) | (S l) \Rightarrow (minus x0 l)]) O)))).(\lambda (H1: (le (S x0) (S
n))).(H n (le_S_n x0 n H1))))) y)))) x).
-theorem minus_minus:
+lemma minus_minus:
\forall (z: nat).(\forall (x: nat).(\forall (y: nat).((le z x) \to ((le z y)
\to ((eq nat (minus x z) (minus y z)) \to (eq nat x y))))))
\def
nat (minus (S x0) (S z0)) (minus (S y0) (S z0)))).(f_equal nat nat S x0 y0
(IH x0 y0 (le_S_n z0 x0 H) (le_S_n z0 y0 H0) H1))))))) y)))) x)))) z).
-theorem plus_plus:
+lemma plus_plus:
\forall (z: nat).(\forall (x1: nat).(\forall (x2: nat).(\forall (y1:
nat).(\forall (y2: nat).((le x1 z) \to ((le x2 z) \to ((eq nat (plus (minus z
x1) y1) (plus (minus z x2) y2)) \to (eq nat (plus x1 y2) (plus x2 y1)))))))))
y1) (IH x2 x4 y1 y2 (le_S_n x2 z0 H) (le_S_n x4 z0 H0) H1))))))))) x3))))
x1)))) z).
-theorem le_S_minus:
+lemma le_S_minus:
\forall (d: nat).(\forall (h: nat).(\forall (n: nat).((le (plus d h) n) \to
(le d (S (minus n h))))))
\def
(le_plus_minus_sym h n (le_trans h (plus d h) n (le_plus_r d h) H))) in (le_S
d (minus n h) (le_minus d n h H))))))).
-theorem lt_x_pred_y:
+lemma lt_x_pred_y:
\forall (x: nat).(\forall (y: nat).((lt x (pred y)) \to (lt (S x) y)))
\def
\lambda (x: nat).(\lambda (y: nat).(nat_ind (\lambda (n: nat).((lt x (pred
include "ground_1/preamble.ma".
-theorem insert_eq:
+lemma insert_eq:
\forall (S: Type[0]).(\forall (x: S).(\forall (P: ((S \to Prop))).(\forall
(G: ((S \to Prop))).(((\forall (y: S).((P y) \to ((eq S y x) \to (G y)))))
\to ((P x) \to (G x))))))
(G: ((S \to Prop))).(\lambda (H: ((\forall (y: S).((P y) \to ((eq S y x) \to
(G y)))))).(\lambda (H0: (P x)).(H x H0 (refl_equal S x))))))).
-theorem unintro:
+lemma unintro:
\forall (A: Type[0]).(\forall (a: A).(\forall (P: ((A \to Prop))).(((\forall
(x: A).(P x))) \to (P a))))
\def
\lambda (A: Type[0]).(\lambda (a: A).(\lambda (P: ((A \to Prop))).(\lambda
(H: ((\forall (x: A).(P x)))).(H a)))).
-theorem xinduction:
+lemma xinduction:
\forall (A: Type[0]).(\forall (t: A).(\forall (P: ((A \to Prop))).(((\forall
(x: A).((eq A t x) \to (P x)))) \to (P t))))
\def
include "ground_1/plist/defs.ma".
-theorem papp_ss:
+lemma papp_ss:
\forall (is1: PList).(\forall (is2: PList).(eq PList (papp (Ss is1) (Ss
is2)) (Ss (papp is1 is2))))
\def
include "ground_1/types/defs.ma".
-theorem and3_rect:
+implied lemma and3_rect:
\forall (P0: Prop).(\forall (P1: Prop).(\forall (P2: Prop).(\forall (P:
Type[0]).(((P0 \to (P1 \to (P2 \to P)))) \to ((and3 P0 P1 P2) \to P)))))
\def
Type[0]).(\lambda (f: ((P0 \to (P1 \to (P2 \to P))))).(\lambda (a: (and3 P0
P1 P2)).(match a with [(and3_intro x x0 x1) \Rightarrow (f x x0 x1)])))))).
-theorem and3_ind:
+implied lemma and3_ind:
\forall (P0: Prop).(\forall (P1: Prop).(\forall (P2: Prop).(\forall (P:
Prop).(((P0 \to (P1 \to (P2 \to P)))) \to ((and3 P0 P1 P2) \to P)))))
\def
\lambda (P0: Prop).(\lambda (P1: Prop).(\lambda (P2: Prop).(\lambda (P:
Prop).(and3_rect P0 P1 P2 P)))).
-theorem and4_rect:
+implied lemma and4_rect:
\forall (P0: Prop).(\forall (P1: Prop).(\forall (P2: Prop).(\forall (P3:
Prop).(\forall (P: Type[0]).(((P0 \to (P1 \to (P2 \to (P3 \to P))))) \to
((and4 P0 P1 P2 P3) \to P))))))
P)))))).(\lambda (a: (and4 P0 P1 P2 P3)).(match a with [(and4_intro x x0 x1
x2) \Rightarrow (f x x0 x1 x2)]))))))).
-theorem and4_ind:
+implied lemma and4_ind:
\forall (P0: Prop).(\forall (P1: Prop).(\forall (P2: Prop).(\forall (P3:
Prop).(\forall (P: Prop).(((P0 \to (P1 \to (P2 \to (P3 \to P))))) \to ((and4
P0 P1 P2 P3) \to P))))))
\lambda (P0: Prop).(\lambda (P1: Prop).(\lambda (P2: Prop).(\lambda (P3:
Prop).(\lambda (P: Prop).(and4_rect P0 P1 P2 P3 P))))).
-theorem and5_rect:
+implied lemma and5_rect:
\forall (P0: Prop).(\forall (P1: Prop).(\forall (P2: Prop).(\forall (P3:
Prop).(\forall (P4: Prop).(\forall (P: Type[0]).(((P0 \to (P1 \to (P2 \to (P3
\to (P4 \to P)))))) \to ((and5 P0 P1 P2 P3 P4) \to P)))))))
(P2 \to (P3 \to (P4 \to P))))))).(\lambda (a: (and5 P0 P1 P2 P3 P4)).(match a
with [(and5_intro x x0 x1 x2 x3) \Rightarrow (f x x0 x1 x2 x3)])))))))).
-theorem and5_ind:
+implied lemma and5_ind:
\forall (P0: Prop).(\forall (P1: Prop).(\forall (P2: Prop).(\forall (P3:
Prop).(\forall (P4: Prop).(\forall (P: Prop).(((P0 \to (P1 \to (P2 \to (P3
\to (P4 \to P)))))) \to ((and5 P0 P1 P2 P3 P4) \to P)))))))
Prop).(\lambda (P4: Prop).(\lambda (P: Prop).(and5_rect P0 P1 P2 P3 P4
P)))))).
-theorem or3_ind:
+implied lemma or3_ind:
\forall (P0: Prop).(\forall (P1: Prop).(\forall (P2: Prop).(\forall (P:
Prop).(((P0 \to P)) \to (((P1 \to P)) \to (((P2 \to P)) \to ((or3 P0 P1 P2)
\to P)))))))
\Rightarrow (f x) | (or3_intro1 x) \Rightarrow (f0 x) | (or3_intro2 x)
\Rightarrow (f1 x)])))))))).
-theorem or4_ind:
+implied lemma or4_ind:
\forall (P0: Prop).(\forall (P1: Prop).(\forall (P2: Prop).(\forall (P3:
Prop).(\forall (P: Prop).(((P0 \to P)) \to (((P1 \to P)) \to (((P2 \to P))
\to (((P3 \to P)) \to ((or4 P0 P1 P2 P3) \to P)))))))))
(or4_intro1 x) \Rightarrow (f0 x) | (or4_intro2 x) \Rightarrow (f1 x) |
(or4_intro3 x) \Rightarrow (f2 x)])))))))))).
-theorem or5_ind:
+implied lemma or5_ind:
\forall (P0: Prop).(\forall (P1: Prop).(\forall (P2: Prop).(\forall (P3:
Prop).(\forall (P4: Prop).(\forall (P: Prop).(((P0 \to P)) \to (((P1 \to P))
\to (((P2 \to P)) \to (((P3 \to P)) \to (((P4 \to P)) \to ((or5 P0 P1 P2 P3
\Rightarrow (f0 x) | (or5_intro2 x) \Rightarrow (f1 x) | (or5_intro3 x)
\Rightarrow (f2 x) | (or5_intro4 x) \Rightarrow (f3 x)])))))))))))).
-theorem ex3_ind:
+implied lemma ex3_ind:
\forall (A0: Type[0]).(\forall (P0: ((A0 \to Prop))).(\forall (P1: ((A0 \to
Prop))).(\forall (P2: ((A0 \to Prop))).(\forall (P: Prop).(((\forall (x0:
A0).((P0 x0) \to ((P1 x0) \to ((P2 x0) \to P))))) \to ((ex3 A0 P0 P1 P2) \to
(e: (ex3 A0 P0 P1 P2)).(match e with [(ex3_intro x x0 x1 x2) \Rightarrow (f x
x0 x1 x2)]))))))).
-theorem ex4_ind:
+implied lemma ex4_ind:
\forall (A0: Type[0]).(\forall (P0: ((A0 \to Prop))).(\forall (P1: ((A0 \to
Prop))).(\forall (P2: ((A0 \to Prop))).(\forall (P3: ((A0 \to
Prop))).(\forall (P: Prop).(((\forall (x0: A0).((P0 x0) \to ((P1 x0) \to ((P2
P3)).(match e with [(ex4_intro x x0 x1 x2 x3) \Rightarrow (f x x0 x1 x2
x3)])))))))).
-theorem ex_2_ind:
+implied lemma ex_2_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (P0: ((A0 \to (A1 \to
Prop)))).(\forall (P: Prop).(((\forall (x0: A0).(\forall (x1: A1).((P0 x0 x1)
\to P)))) \to ((ex_2 A0 A1 P0) \to P)))))
A1).((P0 x0 x1) \to P))))).(\lambda (e: (ex_2 A0 A1 P0)).(match e with
[(ex_2_intro x x0 x1) \Rightarrow (f x x0 x1)])))))).
-theorem ex2_2_ind:
+implied lemma ex2_2_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (P0: ((A0 \to (A1 \to
Prop)))).(\forall (P1: ((A0 \to (A1 \to Prop)))).(\forall (P:
Prop).(((\forall (x0: A0).(\forall (x1: A1).((P0 x0 x1) \to ((P1 x0 x1) \to
P)))))).(\lambda (e: (ex2_2 A0 A1 P0 P1)).(match e with [(ex2_2_intro x x0 x1
x2) \Rightarrow (f x x0 x1 x2)]))))))).
-theorem ex3_2_ind:
+implied lemma ex3_2_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (P0: ((A0 \to (A1 \to
Prop)))).(\forall (P1: ((A0 \to (A1 \to Prop)))).(\forall (P2: ((A0 \to (A1
\to Prop)))).(\forall (P: Prop).(((\forall (x0: A0).(\forall (x1: A1).((P0 x0
(ex3_2 A0 A1 P0 P1 P2)).(match e with [(ex3_2_intro x x0 x1 x2 x3)
\Rightarrow (f x x0 x1 x2 x3)])))))))).
-theorem ex4_2_ind:
+implied lemma ex4_2_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (P0: ((A0 \to (A1 \to
Prop)))).(\forall (P1: ((A0 \to (A1 \to Prop)))).(\forall (P2: ((A0 \to (A1
\to Prop)))).(\forall (P3: ((A0 \to (A1 \to Prop)))).(\forall (P:
A1 P0 P1 P2 P3)).(match e with [(ex4_2_intro x x0 x1 x2 x3 x4) \Rightarrow (f
x x0 x1 x2 x3 x4)]))))))))).
-theorem ex_3_ind:
+implied lemma ex_3_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(P0: ((A0 \to (A1 \to (A2 \to Prop))))).(\forall (P: Prop).(((\forall (x0:
A0).(\forall (x1: A1).(\forall (x2: A2).((P0 x0 x1 x2) \to P))))) \to ((ex_3
P)))))).(\lambda (e: (ex_3 A0 A1 A2 P0)).(match e with [(ex_3_intro x x0 x1
x2) \Rightarrow (f x x0 x1 x2)]))))))).
-theorem ex2_3_ind:
+implied lemma ex2_3_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(P0: ((A0 \to (A1 \to (A2 \to Prop))))).(\forall (P1: ((A0 \to (A1 \to (A2
\to Prop))))).(\forall (P: Prop).(((\forall (x0: A0).(\forall (x1:
P))))))).(\lambda (e: (ex2_3 A0 A1 A2 P0 P1)).(match e with [(ex2_3_intro x
x0 x1 x2 x3) \Rightarrow (f x x0 x1 x2 x3)])))))))).
-theorem ex3_3_ind:
+implied lemma ex3_3_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(P0: ((A0 \to (A1 \to (A2 \to Prop))))).(\forall (P1: ((A0 \to (A1 \to (A2
\to Prop))))).(\forall (P2: ((A0 \to (A1 \to (A2 \to Prop))))).(\forall (P:
P)))))))).(\lambda (e: (ex3_3 A0 A1 A2 P0 P1 P2)).(match e with [(ex3_3_intro
x x0 x1 x2 x3 x4) \Rightarrow (f x x0 x1 x2 x3 x4)]))))))))).
-theorem ex4_3_ind:
+implied lemma ex4_3_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(P0: ((A0 \to (A1 \to (A2 \to Prop))))).(\forall (P1: ((A0 \to (A1 \to (A2
\to Prop))))).(\forall (P2: ((A0 \to (A1 \to (A2 \to Prop))))).(\forall (P3:
A0 A1 A2 P0 P1 P2 P3)).(match e with [(ex4_3_intro x x0 x1 x2 x3 x4 x5)
\Rightarrow (f x x0 x1 x2 x3 x4 x5)])))))))))).
-theorem ex5_3_ind:
+implied lemma ex5_3_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(P0: ((A0 \to (A1 \to (A2 \to Prop))))).(\forall (P1: ((A0 \to (A1 \to (A2
\to Prop))))).(\forall (P2: ((A0 \to (A1 \to (A2 \to Prop))))).(\forall (P3:
A1 A2 P0 P1 P2 P3 P4)).(match e with [(ex5_3_intro x x0 x1 x2 x3 x4 x5 x6)
\Rightarrow (f x x0 x1 x2 x3 x4 x5 x6)]))))))))))).
-theorem ex3_4_ind:
+implied lemma ex3_4_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(A3: Type[0]).(\forall (P0: ((A0 \to (A1 \to (A2 \to (A3 \to
Prop)))))).(\forall (P1: ((A0 \to (A1 \to (A2 \to (A3 \to Prop)))))).(\forall
[(ex3_4_intro x x0 x1 x2 x3 x4 x5) \Rightarrow (f x x0 x1 x2 x3 x4
x5)])))))))))).
-theorem ex4_4_ind:
+implied lemma ex4_4_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(A3: Type[0]).(\forall (P0: ((A0 \to (A1 \to (A2 \to (A3 \to
Prop)))))).(\forall (P1: ((A0 \to (A1 \to (A2 \to (A3 \to Prop)))))).(\forall
[(ex4_4_intro x x0 x1 x2 x3 x4 x5 x6) \Rightarrow (f x x0 x1 x2 x3 x4 x5
x6)]))))))))))).
-theorem ex4_5_ind:
+implied lemma ex4_5_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(A3: Type[0]).(\forall (A4: Type[0]).(\forall (P0: ((A0 \to (A1 \to (A2 \to
(A3 \to (A4 \to Prop))))))).(\forall (P1: ((A0 \to (A1 \to (A2 \to (A3 \to
e with [(ex4_5_intro x x0 x1 x2 x3 x4 x5 x6 x7) \Rightarrow (f x x0 x1 x2 x3
x4 x5 x6 x7)])))))))))))).
-theorem ex5_5_ind:
+implied lemma ex5_5_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(A3: Type[0]).(\forall (A4: Type[0]).(\forall (P0: ((A0 \to (A1 \to (A2 \to
(A3 \to (A4 \to Prop))))))).(\forall (P1: ((A0 \to (A1 \to (A2 \to (A3 \to
A3 A4 P0 P1 P2 P3 P4)).(match e with [(ex5_5_intro x x0 x1 x2 x3 x4 x5 x6 x7
x8) \Rightarrow (f x x0 x1 x2 x3 x4 x5 x6 x7 x8)]))))))))))))).
-theorem ex6_6_ind:
+implied lemma ex6_6_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(A3: Type[0]).(\forall (A4: Type[0]).(\forall (A5: Type[0]).(\forall (P0:
((A0 \to (A1 \to (A2 \to (A3 \to (A4 \to (A5 \to Prop)))))))).(\forall (P1:
P5)).(match e with [(ex6_6_intro x x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10)
\Rightarrow (f x x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10)]))))))))))))))).
-theorem ex6_7_ind:
+implied lemma ex6_7_ind:
\forall (A0: Type[0]).(\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall
(A3: Type[0]).(\forall (A4: Type[0]).(\forall (A5: Type[0]).(\forall (A6:
Type[0]).(\forall (P0: ((A0 \to (A1 \to (A2 \to (A3 \to (A4 \to (A5 \to (A6
include "ground_1/types/defs.ma".
-theorem ex2_sym:
+lemma ex2_sym:
\forall (A: Type[0]).(\forall (P: ((A \to Prop))).(\forall (Q: ((A \to
Prop))).((ex2 A (\lambda (x: A).(P x)) (\lambda (x: A).(Q x))) \to (ex2 A
(\lambda (x: A).(Q x)) (\lambda (x: A).(P x))))))
include "legacy_1/coq/defs.ma".
-theorem False_rect:
+implied lemma False_rect:
\forall (P: Type[0]).(False \to P)
\def
\lambda (P: Type[0]).(\lambda (f: False).(match f in False with [])).
-theorem False_ind:
+implied lemma False_ind:
\forall (P: Prop).(False \to P)
\def
\lambda (P: Prop).(False_rect P).
-theorem land_rect:
+implied lemma land_rect:
\forall (A: Prop).(\forall (B: Prop).(\forall (P: Type[0]).(((A \to (B \to
P))) \to ((land A B) \to P))))
\def
\to (B \to P)))).(\lambda (a: (land A B)).(match a with [(conj x x0)
\Rightarrow (f x x0)]))))).
-theorem land_ind:
+implied lemma land_ind:
\forall (A: Prop).(\forall (B: Prop).(\forall (P: Prop).(((A \to (B \to P)))
\to ((land A B) \to P))))
\def
\lambda (A: Prop).(\lambda (B: Prop).(\lambda (P: Prop).(land_rect A B P))).
-theorem or_ind:
+implied lemma or_ind:
\forall (A: Prop).(\forall (B: Prop).(\forall (P: Prop).(((A \to P)) \to
(((B \to P)) \to ((or A B) \to P)))))
\def
P))).(\lambda (f0: ((B \to P))).(\lambda (o: (or A B)).(match o with
[(or_introl x) \Rightarrow (f x) | (or_intror x) \Rightarrow (f0 x)])))))).
-theorem ex_ind:
+implied lemma ex_ind:
\forall (A: Type[0]).(\forall (P: ((A \to Prop))).(\forall (P0:
Prop).(((\forall (x: A).((P x) \to P0))) \to ((ex A P) \to P0))))
\def
Prop).(\lambda (f: ((\forall (x: A).((P x) \to P0)))).(\lambda (e: (ex A
P)).(match e with [(ex_intro x x0) \Rightarrow (f x x0)]))))).
-theorem ex2_ind:
+implied lemma ex2_ind:
\forall (A: Type[0]).(\forall (P: ((A \to Prop))).(\forall (Q: ((A \to
Prop))).(\forall (P0: Prop).(((\forall (x: A).((P x) \to ((Q x) \to P0))))
\to ((ex2 A P Q) \to P0)))))
\to P0))))).(\lambda (e: (ex2 A P Q)).(match e with [(ex_intro2 x x0 x1)
\Rightarrow (f x x0 x1)])))))).
-theorem eq_rect:
+implied lemma eq_rect:
\forall (A: Type[0]).(\forall (x: A).(\forall (P: ((A \to Type[0]))).((P x)
\to (\forall (y: A).((eq A x y) \to (P y))))))
\def
Type[0]))).(\lambda (f: (P x)).(\lambda (y: A).(\lambda (e: (eq A x
y)).(match e with [refl_equal \Rightarrow f])))))).
-theorem eq_ind:
+implied lemma eq_ind:
\forall (A: Type[0]).(\forall (x: A).(\forall (P: ((A \to Prop))).((P x) \to
(\forall (y: A).((eq A x y) \to (P y))))))
\def
\lambda (A: Type[0]).(\lambda (x: A).(\lambda (P: ((A \to Prop))).(eq_rect A
x P))).
-let rec le_ind (n: nat) (P: (nat \to Prop)) (f: P n) (f0: (\forall (m:
-nat).((le n m) \to ((P m) \to (P (S m)))))) (n0: nat) (l: le n n0) on l: P n0
-\def match l with [le_n \Rightarrow f | (le_S m l0) \Rightarrow (f0 m l0
+implied let rec le_ind (n: nat) (P: (nat \to Prop)) (f: P n) (f0: (\forall
+(m: nat).((le n m) \to ((P m) \to (P (S m)))))) (n0: nat) (l: le n n0) on l:
+P n0 \def match l with [le_n \Rightarrow f | (le_S m l0) \Rightarrow (f0 m l0
((le_ind n P f f0) m l0))].
-let rec Acc_ind (A: Type[0]) (R: (A \to (A \to Prop))) (P: (A \to Prop)) (f:
-(\forall (x: A).(((\forall (y: A).((R y x) \to (Acc A R y)))) \to (((\forall
-(y: A).((R y x) \to (P y)))) \to (P x))))) (a: A) (a0: Acc A R a) on a0: P a
-\def match a0 with [(Acc_intro x a1) \Rightarrow (f x a1 (\lambda (y:
-A).(\lambda (r0: (R y x)).((Acc_ind A R P f) y (a1 y r0)))))].
+implied let rec Acc_ind (A: Type[0]) (R: (A \to (A \to Prop))) (P: (A \to
+Prop)) (f: (\forall (x: A).(((\forall (y: A).((R y x) \to (Acc A R y)))) \to
+(((\forall (y: A).((R y x) \to (P y)))) \to (P x))))) (a: A) (a0: Acc A R a)
+on a0: P a \def match a0 with [(Acc_intro x a1) \Rightarrow (f x a1 (\lambda
+(y: A).(\lambda (r0: (R y x)).((Acc_ind A R P f) y (a1 y r0)))))].
include "legacy_1/coq/fwd.ma".
-theorem f_equal:
+lemma f_equal:
\forall (A: Type[0]).(\forall (B: Type[0]).(\forall (f: ((A \to
B))).(\forall (x: A).(\forall (y: A).((eq A x y) \to (eq B (f x) (f y)))))))
\def
B))).(\lambda (x: A).(\lambda (y: A).(\lambda (H: (eq A x y)).(eq_ind A x
(\lambda (a: A).(eq B (f x) (f a))) (refl_equal B (f x)) y H)))))).
-theorem f_equal2:
+lemma f_equal2:
\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall (B: Type[0]).(\forall
(f: ((A1 \to (A2 \to B)))).(\forall (x1: A1).(\forall (y1: A1).(\forall (x2:
A2).(\forall (y2: A2).((eq A1 x1 y1) \to ((eq A2 x2 y2) \to (eq B (f x1 x2)
y2)).(eq_ind A2 x2 (\lambda (a: A2).(eq B (f x1 x2) (f x1 a))) (refl_equal B
(f x1 x2)) y2 H0)) y1 H))))))))).
-theorem f_equal3:
+lemma f_equal3:
\forall (A1: Type[0]).(\forall (A2: Type[0]).(\forall (A3: Type[0]).(\forall
(B: Type[0]).(\forall (f: ((A1 \to (A2 \to (A3 \to B))))).(\forall (x1:
A1).(\forall (y1: A1).(\forall (x2: A2).(\forall (y2: A2).(\forall (x3:
x3 (\lambda (a: A3).(eq B (f x1 x2 x3) (f x1 x2 a))) (refl_equal B (f x1 x2
x3)) y3 H1)) y2 H0)) y1 H)))))))))))).
-theorem sym_eq:
+lemma sym_eq:
\forall (A: Type[0]).(\forall (x: A).(\forall (y: A).((eq A x y) \to (eq A y
x))))
\def
\lambda (A: Type[0]).(\lambda (x: A).(\lambda (y: A).(\lambda (H: (eq A x
y)).(eq_ind A x (\lambda (a: A).(eq A a x)) (refl_equal A x) y H)))).
-theorem eq_ind_r:
+lemma eq_ind_r:
\forall (A: Type[0]).(\forall (x: A).(\forall (P: ((A \to Prop))).((P x) \to
(\forall (y: A).((eq A y x) \to (P y))))))
\def
(H: (P x)).(\lambda (y: A).(\lambda (H0: (eq A y x)).(match (sym_eq A y x H0)
with [refl_equal \Rightarrow H])))))).
-theorem trans_eq:
+lemma trans_eq:
\forall (A: Type[0]).(\forall (x: A).(\forall (y: A).(\forall (z: A).((eq A
x y) \to ((eq A y z) \to (eq A x z))))))
\def
A).(\lambda (H: (eq A x y)).(\lambda (H0: (eq A y z)).(eq_ind A y (\lambda
(a: A).(eq A x a)) H z H0)))))).
-theorem sym_not_eq:
+lemma sym_not_eq:
\forall (A: Type[0]).(\forall (x: A).(\forall (y: A).((not (eq A x y)) \to
(not (eq A y x)))))
\def
A x y))).(\lambda (h2: (eq A y x)).(h1 (eq_ind A y (\lambda (a: A).(eq A a
y)) (refl_equal A y) x h2)))))).
-theorem nat_double_ind:
+lemma nat_double_ind:
\forall (R: ((nat \to (nat \to Prop)))).(((\forall (n: nat).(R O n))) \to
(((\forall (n: nat).(R (S n) O))) \to (((\forall (n: nat).(\forall (m:
nat).((R n m) \to (R (S n) (S m)))))) \to (\forall (n: nat).(\forall (m:
(\lambda (n1: nat).(R (S n0) n1)) (H0 n0) (\lambda (n1: nat).(\lambda (_: (R
(S n0) n1)).(H1 n0 n1 (H2 n1)))) m)))) n))))).
-theorem eq_add_S:
+lemma eq_add_S:
\forall (n: nat).(\forall (m: nat).((eq nat (S n) (S m)) \to (eq nat n m)))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (eq nat (S n) (S
m))).(f_equal nat nat pred (S n) (S m) H))).
-theorem O_S:
+lemma O_S:
\forall (n: nat).(not (eq nat O (S n)))
\def
\lambda (n: nat).(\lambda (H: (eq nat O (S n))).(eq_ind nat (S n) (\lambda
(n0: nat).(IsSucc n0)) I O (sym_eq nat O (S n) H))).
-theorem not_eq_S:
+lemma not_eq_S:
\forall (n: nat).(\forall (m: nat).((not (eq nat n m)) \to (not (eq nat (S
n) (S m)))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (not (eq nat n m))).(\lambda
(H0: (eq nat (S n) (S m))).(H (eq_add_S n m H0))))).
-theorem pred_Sn:
+lemma pred_Sn:
\forall (m: nat).(eq nat m (pred (S m)))
\def
\lambda (m: nat).(refl_equal nat (pred (S m))).
-theorem S_pred:
+lemma S_pred:
\forall (n: nat).(\forall (m: nat).((lt m n) \to (eq nat n (S (pred n)))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (lt m n)).(le_ind (S m)
m)))) (\lambda (m0: nat).(\lambda (_: (le (S m) m0)).(\lambda (_: (eq nat m0
(S (pred m0)))).(refl_equal nat (S (pred (S m0))))))) n H))).
-theorem le_trans:
+lemma le_trans:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((le n m) \to ((le m p)
\to (le n p)))))
\def
(\lambda (m0: nat).(\lambda (_: (le m m0)).(\lambda (IHle: (le n m0)).(le_S n
m0 IHle)))) p H0))))).
-theorem le_trans_S:
+lemma le_trans_S:
\forall (n: nat).(\forall (m: nat).((le (S n) m) \to (le n m)))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (le (S n) m)).(le_trans n (S
n) m (le_S n n (le_n n)) H))).
-theorem le_n_S:
+lemma le_n_S:
\forall (n: nat).(\forall (m: nat).((le n m) \to (le (S n) (S m))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (le n m)).(le_ind n (\lambda
(n0: nat).(le (S n) (S n0))) (le_n (S n)) (\lambda (m0: nat).(\lambda (_: (le
n m0)).(\lambda (IHle: (le (S n) (S m0))).(le_S (S n) (S m0) IHle)))) m H))).
-theorem le_O_n:
+lemma le_O_n:
\forall (n: nat).(le O n)
\def
\lambda (n: nat).(nat_ind (\lambda (n0: nat).(le O n0)) (le_n O) (\lambda
(n0: nat).(\lambda (IHn: (le O n0)).(le_S O n0 IHn))) n).
-theorem le_S_n:
+lemma le_S_n:
\forall (n: nat).(\forall (m: nat).((le (S n) (S m)) \to (le n m)))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (le (S n) (S m))).(le_ind (S
nat).(\lambda (H0: (le (S n) m0)).(\lambda (_: (le n (pred m0))).(le_trans_S
n m0 H0)))) (S m) H))).
-theorem le_Sn_O:
+lemma le_Sn_O:
\forall (n: nat).(not (le (S n) O))
\def
\lambda (n: nat).(\lambda (H: (le (S n) O)).(le_ind (S n) (\lambda (n0:
nat).(IsSucc n0)) I (\lambda (m: nat).(\lambda (_: (le (S n) m)).(\lambda (_:
(IsSucc m)).I))) O H)).
-theorem le_Sn_n:
+lemma le_Sn_n:
\forall (n: nat).(not (le (S n) n))
\def
\lambda (n: nat).(nat_ind (\lambda (n0: nat).(not (le (S n0) n0))) (le_Sn_O
O) (\lambda (n0: nat).(\lambda (IHn: (not (le (S n0) n0))).(\lambda (H: (le
(S (S n0)) (S n0))).(IHn (le_S_n (S n0) n0 H))))) n).
-theorem le_antisym:
+lemma le_antisym:
\forall (n: nat).(\forall (m: nat).((le n m) \to ((le m n) \to (eq nat n
m))))
\def
m0)) (let H2 \def (le_trans (S m0) n m0 H1 H) in ((let H3 \def (le_Sn_n m0)
in (\lambda (H4: (le (S m0) m0)).(H3 H4))) H2))))))) m h))).
-theorem le_n_O_eq:
+lemma le_n_O_eq:
\forall (n: nat).((le n O) \to (eq nat O n))
\def
\lambda (n: nat).(\lambda (H: (le n O)).(le_antisym O n (le_O_n n) H)).
-theorem le_elim_rel:
+lemma le_elim_rel:
\forall (P: ((nat \to (nat \to Prop)))).(((\forall (p: nat).(P O p))) \to
(((\forall (p: nat).(\forall (q: nat).((le p q) \to ((P p q) \to (P (S p) (S
q))))))) \to (\forall (n: nat).(\forall (m: nat).((le n m) \to (P n m))))))
(P (S n0) m0)).(H0 n0 m0 (le_trans_S n0 m0 H1) (IHn m0 (le_trans_S n0 m0
H1)))))) m Le))))) n)))).
-theorem lt_n_n:
+lemma lt_n_n:
\forall (n: nat).(not (lt n n))
\def
le_Sn_n.
-theorem lt_n_S:
+lemma lt_n_S:
\forall (n: nat).(\forall (m: nat).((lt n m) \to (lt (S n) (S m))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (lt n m)).(le_n_S (S n) m
H))).
-theorem lt_n_Sn:
+lemma lt_n_Sn:
\forall (n: nat).(lt n (S n))
\def
\lambda (n: nat).(le_n (S n)).
-theorem lt_S_n:
+lemma lt_S_n:
\forall (n: nat).(\forall (m: nat).((lt (S n) (S m)) \to (lt n m)))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (lt (S n) (S m))).(le_S_n (S
n) m H))).
-theorem lt_n_O:
+lemma lt_n_O:
\forall (n: nat).(not (lt n O))
\def
le_Sn_O.
-theorem lt_trans:
+lemma lt_trans:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((lt n m) \to ((lt m p)
\to (lt n p)))))
\def
(S n) m H) (\lambda (m0: nat).(\lambda (_: (le (S m) m0)).(\lambda (IHle: (lt
n m0)).(le_S (S n) m0 IHle)))) p H0))))).
-theorem lt_O_Sn:
+lemma lt_O_Sn:
\forall (n: nat).(lt O (S n))
\def
\lambda (n: nat).(le_n_S O n (le_O_n n)).
-theorem lt_le_S:
+lemma lt_le_S:
\forall (n: nat).(\forall (p: nat).((lt n p) \to (le (S n) p)))
\def
\lambda (n: nat).(\lambda (p: nat).(\lambda (H: (lt n p)).H)).
-theorem le_not_lt:
+lemma le_not_lt:
\forall (n: nat).(\forall (m: nat).((le n m) \to (not (lt m n))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (le n m)).(le_ind n (\lambda
m0)).(\lambda (IHle: (not (lt m0 n))).(\lambda (H1: (lt (S m0) n)).(IHle
(le_trans_S (S m0) n H1)))))) m H))).
-theorem le_lt_n_Sm:
+lemma le_lt_n_Sm:
\forall (n: nat).(\forall (m: nat).((le n m) \to (lt n (S m))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (le n m)).(le_n_S n m H))).
-theorem le_lt_trans:
+lemma le_lt_trans:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((le n m) \to ((lt m p)
\to (lt n p)))))
\def
(le_n_S n m H) (\lambda (m0: nat).(\lambda (_: (le (S m) m0)).(\lambda (IHle:
(lt n m0)).(le_S (S n) m0 IHle)))) p H0))))).
-theorem lt_le_trans:
+lemma lt_le_trans:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((lt n m) \to ((le m p)
\to (lt n p)))))
\def
(\lambda (m0: nat).(\lambda (_: (le m m0)).(\lambda (IHle: (lt n m0)).(le_S
(S n) m0 IHle)))) p H0))))).
-theorem lt_le_weak:
+lemma lt_le_weak:
\forall (n: nat).(\forall (m: nat).((lt n m) \to (le n m)))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (lt n m)).(le_trans_S n m
H))).
-theorem lt_n_Sm_le:
+lemma lt_n_Sm_le:
\forall (n: nat).(\forall (m: nat).((lt n (S m)) \to (le n m)))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (lt n (S m))).(le_S_n n m
H))).
-theorem le_lt_or_eq:
+lemma le_lt_or_eq:
\forall (n: nat).(\forall (m: nat).((le n m) \to (or (lt n m) (eq nat n m))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (le n m)).(le_ind n (\lambda
(or (lt n m0) (eq nat n m0))).(or_introl (lt n (S m0)) (eq nat n (S m0))
(le_n_S n m0 H0))))) m H))).
-theorem le_or_lt:
+lemma le_or_lt:
\forall (n: nat).(\forall (m: nat).(or (le n m) (lt m n)))
\def
\lambda (n: nat).(\lambda (m: nat).(nat_double_ind (\lambda (n0:
n0)) (le_n_S n0 m0 H0))) (\lambda (H0: (lt m0 n0)).(or_intror (le (S n0) (S
m0)) (lt (S m0) (S n0)) (le_n_S (S m0) n0 H0))) H)))) n m)).
-theorem plus_n_O:
+lemma plus_n_O:
\forall (n: nat).(eq nat n (plus n O))
\def
\lambda (n: nat).(nat_ind (\lambda (n0: nat).(eq nat n0 (plus n0 O)))
(refl_equal nat O) (\lambda (n0: nat).(\lambda (H: (eq nat n0 (plus n0
O))).(f_equal nat nat S n0 (plus n0 O) H))) n).
-theorem plus_n_Sm:
+lemma plus_n_Sm:
\forall (n: nat).(\forall (m: nat).(eq nat (S (plus n m)) (plus n (S m))))
\def
\lambda (m: nat).(\lambda (n: nat).(nat_ind (\lambda (n0: nat).(eq nat (S
nat).(\lambda (H: (eq nat (S (plus n0 n)) (plus n0 (S n)))).(f_equal nat nat
S (S (plus n0 n)) (plus n0 (S n)) H))) m)).
-theorem plus_sym:
+lemma plus_sym:
\forall (n: nat).(\forall (m: nat).(eq nat (plus n m) (plus m n)))
\def
\lambda (n: nat).(\lambda (m: nat).(nat_ind (\lambda (n0: nat).(eq nat (plus
(plus y m)) n0)) (f_equal nat nat S (plus y m) (plus m y) H) (plus m (S y))
(plus_n_Sm m y)))) n)).
-theorem plus_Snm_nSm:
+lemma plus_Snm_nSm:
\forall (n: nat).(\forall (m: nat).(eq nat (plus (S n) m) (plus n (S m))))
\def
\lambda (n: nat).(\lambda (m: nat).(eq_ind_r nat (plus m n) (\lambda (n0:
(n0: nat).(eq nat (S (plus m n)) n0)) (refl_equal nat (plus (S m) n)) (plus n
(S m)) (plus_sym n (S m))) (plus n m) (plus_sym n m))).
-theorem plus_assoc_l:
+lemma plus_assoc_l:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).(eq nat (plus n (plus m
p)) (plus (plus n m) p))))
\def
(plus (plus n0 m) p))).(f_equal nat nat S (plus n0 (plus m p)) (plus (plus n0
m) p) H))) n))).
-theorem plus_assoc_r:
+lemma plus_assoc_r:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).(eq nat (plus (plus n
m) p) (plus n (plus m p)))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (p: nat).(sym_eq nat (plus n
(plus m p)) (plus (plus n m) p) (plus_assoc_l n m p)))).
-theorem simpl_plus_l:
+lemma simpl_plus_l:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((eq nat (plus n m)
(plus n p)) \to (eq nat m p))))
\def
(plus n0 m) (plus n0 p) (f_equal nat nat (plus n0) (plus n0 m) (plus n0 p)
(eq_add_S (plus n0 m) (plus n0 p) H))))))))) n).
-theorem minus_n_O:
+lemma minus_n_O:
\forall (n: nat).(eq nat n (minus n O))
\def
\lambda (n: nat).(nat_ind (\lambda (n0: nat).(eq nat n0 (minus n0 O)))
(refl_equal nat O) (\lambda (n0: nat).(\lambda (_: (eq nat n0 (minus n0
O))).(refl_equal nat (S n0)))) n).
-theorem minus_n_n:
+lemma minus_n_n:
\forall (n: nat).(eq nat O (minus n n))
\def
\lambda (n: nat).(nat_ind (\lambda (n0: nat).(eq nat O (minus n0 n0)))
(refl_equal nat O) (\lambda (n0: nat).(\lambda (IHn: (eq nat O (minus n0
n0))).IHn)) n).
-theorem minus_Sn_m:
+lemma minus_Sn_m:
\forall (n: nat).(\forall (m: nat).((le m n) \to (eq nat (S (minus n m))
(minus (S n) m))))
\def
(le p q)).(\lambda (H0: (eq nat (S (minus q p)) (match p with [O \Rightarrow
(S q) | (S l) \Rightarrow (minus q l)]))).H0)))) m n Le))).
-theorem plus_minus:
+lemma plus_minus:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((eq nat n (plus m p))
\to (eq nat p (minus n m)))))
\def
nat p (minus m0 n0))))).(\lambda (H0: (eq nat (S m0) (S (plus n0 p)))).(H
(eq_add_S m0 (plus n0 p) H0)))))) m n))).
-theorem minus_plus:
+lemma minus_plus:
\forall (n: nat).(\forall (m: nat).(eq nat (minus (plus n m) n) m))
\def
\lambda (n: nat).(\lambda (m: nat).(sym_eq nat m (minus (plus n m) n)
(plus_minus (plus n m) n m (refl_equal nat (plus n m))))).
-theorem le_pred_n:
+lemma le_pred_n:
\forall (n: nat).(le (pred n) n)
\def
\lambda (n: nat).(nat_ind (\lambda (n0: nat).(le (pred n0) n0)) (le_n O)
(\lambda (n0: nat).(\lambda (_: (le (pred n0) n0)).(le_S (pred (S n0)) n0
(le_n n0)))) n).
-theorem le_plus_l:
+lemma le_plus_l:
\forall (n: nat).(\forall (m: nat).(le n (plus n m)))
\def
\lambda (n: nat).(nat_ind (\lambda (n0: nat).(\forall (m: nat).(le n0 (plus
((\forall (m: nat).(le n0 (plus n0 m))))).(\lambda (m: nat).(le_n_S n0 (plus
n0 m) (IHn m))))) n).
-theorem le_plus_r:
+lemma le_plus_r:
\forall (n: nat).(\forall (m: nat).(le m (plus n m)))
\def
\lambda (n: nat).(\lambda (m: nat).(nat_ind (\lambda (n0: nat).(le m (plus
n0 m))) (le_n m) (\lambda (n0: nat).(\lambda (H: (le m (plus n0 m))).(le_S m
(plus n0 m) H))) n)).
-theorem simpl_le_plus_l:
+lemma simpl_le_plus_l:
\forall (p: nat).(\forall (n: nat).(\forall (m: nat).((le (plus p n) (plus p
m)) \to (le n m))))
\def
(H: (le (S (plus p0 n)) (S (plus p0 m)))).(IHp n m (le_S_n (plus p0 n) (plus
p0 m) H))))))) p).
-theorem le_plus_trans:
+lemma le_plus_trans:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((le n m) \to (le n
(plus m p)))))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (p: nat).(\lambda (H: (le n
m)).(le_trans n m (plus m p) H (le_plus_l m p))))).
-theorem le_reg_l:
+lemma le_reg_l:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((le n m) \to (le (plus
p n) (plus p m)))))
\def
m))))).(\lambda (H: (le n m)).(le_n_S (plus p0 n) (plus p0 m) (IHp H)))))
p))).
-theorem le_plus_plus:
+lemma le_plus_plus:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).(\forall (q: nat).((le
n m) \to ((le p q) \to (le (plus n p) (plus m q)))))))
\def
nat).(\lambda (_: (le n m0)).(\lambda (H2: (le (plus n p) (plus m0 q))).(le_S
(plus n p) (plus m0 q) H2)))) m H)))))).
-theorem le_plus_minus:
+lemma le_plus_minus:
\forall (n: nat).(\forall (m: nat).((le n m) \to (eq nat m (plus n (minus m
n)))))
\def
(_: (le p q)).(\lambda (H0: (eq nat q (plus p (minus q p)))).(f_equal nat nat
S q (plus p (minus q p)) H0))))) n m Le))).
-theorem le_plus_minus_r:
+lemma le_plus_minus_r:
\forall (n: nat).(\forall (m: nat).((le n m) \to (eq nat (plus n (minus m
n)) m)))
\def
\lambda (n: nat).(\lambda (m: nat).(\lambda (H: (le n m)).(sym_eq nat m
(plus n (minus m n)) (le_plus_minus n m H)))).
-theorem simpl_lt_plus_l:
+lemma simpl_lt_plus_l:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((lt (plus p n) (plus p
m)) \to (lt n m))))
\def
m)))).(\lambda (H: (lt (S (plus p0 n)) (S (plus p0 m)))).(IHp (le_S_n (S
(plus p0 n)) (plus p0 m) H))))) p))).
-theorem lt_reg_l:
+lemma lt_reg_l:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((lt n m) \to (lt (plus
p n) (plus p m)))))
\def
m))))).(\lambda (H: (lt n m)).(lt_n_S (plus p0 n) (plus p0 m) (IHp H)))))
p))).
-theorem lt_reg_r:
+lemma lt_reg_r:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).((lt n m) \to (lt (plus
n p) (plus m p)))))
\def
(plus n0 n) (plus n0 m))).(lt_reg_l n m (S n0) H))) p) (plus m p) (plus_sym m
p)) (plus n p) (plus_sym n p))))).
-theorem le_lt_plus_plus:
+lemma le_lt_plus_plus:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).(\forall (q: nat).((le
n m) \to ((lt p q) \to (lt (plus n p) (plus m q)))))))
\def
(S p)) (\lambda (n0: nat).(le n0 (plus m q))) (le_plus_plus n m (S p) q H H0)
(plus (S n) p) (plus_Snm_nSm n p))))))).
-theorem lt_le_plus_plus:
+lemma lt_le_plus_plus:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).(\forall (q: nat).((lt
n m) \to ((le p q) \to (lt (plus n p) (plus m q)))))))
\def
nat).(\lambda (H: (le (S n) m)).(\lambda (H0: (le p q)).(le_plus_plus (S n) m
p q H H0)))))).
-theorem lt_plus_plus:
+lemma lt_plus_plus:
\forall (n: nat).(\forall (m: nat).(\forall (p: nat).(\forall (q: nat).((lt
n m) \to ((lt p q) \to (lt (plus n p) (plus m q)))))))
\def
nat).(\lambda (H: (lt n m)).(\lambda (H0: (lt p q)).(lt_le_plus_plus n m p q
H (lt_le_weak p q H0))))))).
-theorem well_founded_ltof:
+lemma well_founded_ltof:
\forall (A: Type[0]).(\forall (f: ((A \to nat))).(well_founded A (ltof A f)))
\def
\lambda (A: Type[0]).(\lambda (f: ((A \to nat))).(let H \def (\lambda (n:
(lt_n_Sm_le (f a) n0 ltSma)))))))))) n)) in (\lambda (a: A).(H (S (f a)) a
(le_n (S (f a))))))).
-theorem lt_wf:
+lemma lt_wf:
well_founded nat lt
\def
well_founded_ltof nat (\lambda (m: nat).m).
-theorem lt_wf_ind:
+lemma lt_wf_ind:
\forall (p: nat).(\forall (P: ((nat \to Prop))).(((\forall (n:
nat).(((\forall (m: nat).((lt m n) \to (P m)))) \to (P n)))) \to (P p)))
\def