]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/LAMBDA-TYPES/LambdaDelta-1/spare.ma
cicInspect: now we can choose not to count the Cic.Implicit constructors
[helm.git] / helm / software / matita / contribs / LAMBDA-TYPES / LambdaDelta-1 / spare.ma
index f40cf08fecd59a2717d8d4609aadfecd91ae2ffd..25afeec1d5afa2165fbb4ebebca73144253d13ab 100644 (file)
 
 include "LambdaDelta-1/theory.ma".
 
-definition cbk:
- C \to nat
-\def
- let rec cbk (c: C) on c: nat \def (match c with [(CSort m) \Rightarrow m | 
-(CHead c0 _ _) \Rightarrow (cbk c0)]) in cbk.
-
-definition app1:
- C \to (T \to T)
-\def
- let rec app1 (c: C) on c: (T \to T) \def (\lambda (t: T).(match c with 
-[(CSort _) \Rightarrow t | (CHead c0 k u) \Rightarrow (app1 c0 (THead k u 
-t))])) in app1.
-
-theorem 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
- \lambda (xs: TList).(TList_ind (\lambda (t: TList).(\forall (ts: 
-TList).(\forall (h: nat).(\forall (d: nat).((eq TList (lifts h d t) (lifts h 
-d ts)) \to (eq TList t ts)))))) (\lambda (ts: TList).(TList_ind (\lambda (t: 
-TList).(\forall (h: nat).(\forall (d: nat).((eq TList (lifts h d TNil) (lifts 
-h d t)) \to (eq TList TNil t))))) (\lambda (_: nat).(\lambda (_: 
-nat).(\lambda (H: (eq TList TNil TNil)).H))) (\lambda (t: T).(\lambda (t0: 
-TList).(\lambda (_: ((\forall (h: nat).(\forall (d: nat).((eq TList TNil 
-(lifts h d t0)) \to (eq TList TNil t0)))))).(\lambda (h: nat).(\lambda (d: 
-nat).(\lambda (H0: (eq TList TNil (TCons (lift h d t) (lifts h d t0)))).(let 
-H1 \def (eq_ind TList TNil (\lambda (ee: TList).(match ee in TList return 
-(\lambda (_: TList).Prop) with [TNil \Rightarrow True | (TCons _ _) 
-\Rightarrow False])) I (TCons (lift h d t) (lifts h d t0)) H0) in (False_ind 
-(eq TList TNil (TCons t t0)) H1)))))))) ts)) (\lambda (t: T).(\lambda (t0: 
-TList).(\lambda (H: ((\forall (ts: TList).(\forall (h: nat).(\forall (d: 
-nat).((eq TList (lifts h d t0) (lifts h d ts)) \to (eq TList t0 
-ts))))))).(\lambda (ts: TList).(TList_ind (\lambda (t1: TList).(\forall (h: 
-nat).(\forall (d: nat).((eq TList (lifts h d (TCons t t0)) (lifts h d t1)) 
-\to (eq TList (TCons t t0) t1))))) (\lambda (h: nat).(\lambda (d: 
-nat).(\lambda (H0: (eq TList (TCons (lift h d t) (lifts h d t0)) TNil)).(let 
-H1 \def (eq_ind TList (TCons (lift h d t) (lifts h d t0)) (\lambda (ee: 
-TList).(match ee in TList return (\lambda (_: TList).Prop) with [TNil 
-\Rightarrow False | (TCons _ _) \Rightarrow True])) I TNil H0) in (False_ind 
-(eq TList (TCons t t0) TNil) H1))))) (\lambda (t1: T).(\lambda (t2: 
-TList).(\lambda (_: ((\forall (h: nat).(\forall (d: nat).((eq TList (TCons 
-(lift h d t) (lifts h d t0)) (lifts h d t2)) \to (eq TList (TCons t t0) 
-t2)))))).(\lambda (h: nat).(\lambda (d: nat).(\lambda (H1: (eq TList (TCons 
-(lift h d t) (lifts h d t0)) (TCons (lift h d t1) (lifts h d t2)))).(let H2 
-\def (f_equal TList T (\lambda (e: TList).(match e in TList return (\lambda 
-(_: TList).T) with [TNil \Rightarrow ((let rec lref_map (f: ((nat \to nat))) 
-(d0: nat) (t3: T) on t3: T \def (match t3 with [(TSort n) \Rightarrow (TSort 
-n) | (TLRef i) \Rightarrow (TLRef (match (blt i d0) with [true \Rightarrow i 
-| false \Rightarrow (f i)])) | (THead k u t4) \Rightarrow (THead k (lref_map 
-f d0 u) (lref_map f (s k d0) t4))]) in lref_map) (\lambda (x: nat).(plus x 
-h)) d t) | (TCons t3 _) \Rightarrow t3])) (TCons (lift h d t) (lifts h d t0)) 
-(TCons (lift h d t1) (lifts h d t2)) H1) in ((let H3 \def (f_equal TList 
-TList (\lambda (e: TList).(match e in TList return (\lambda (_: TList).TList) 
-with [TNil \Rightarrow ((let rec lifts (h0: nat) (d0: nat) (ts0: TList) on 
-ts0: TList \def (match ts0 with [TNil \Rightarrow TNil | (TCons t3 ts1) 
-\Rightarrow (TCons (lift h0 d0 t3) (lifts h0 d0 ts1))]) in lifts) h d t0) | 
-(TCons _ t3) \Rightarrow t3])) (TCons (lift h d t) (lifts h d t0)) (TCons 
-(lift h d t1) (lifts h d t2)) H1) in (\lambda (H4: (eq T (lift h d t) (lift h 
-d t1))).(eq_ind T t (\lambda (t3: T).(eq TList (TCons t t0) (TCons t3 t2))) 
-(f_equal2 T TList TList TCons t t t0 t2 (refl_equal T t) (H t2 h d H3)) t1 
-(lift_inj t t1 h d H4)))) H2)))))))) ts))))) xs).
-
-theorem 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
- \lambda (c: C).(\lambda (t: T).(\lambda (ts: TList).(TList_ind (\lambda (t0: 
-TList).((nfs2 c (TApp t0 t)) \to (land (nfs2 c t0) (nf2 c t)))) (\lambda (H: 
-(land (nf2 c t) True)).(let H0 \def H in (and_ind (nf2 c t) True (land True 
-(nf2 c t)) (\lambda (H1: (nf2 c t)).(\lambda (_: True).(conj True (nf2 c t) I 
-H1))) H0))) (\lambda (t0: T).(\lambda (t1: TList).(\lambda (H: (((nfs2 c 
-(TApp t1 t)) \to (land (nfs2 c t1) (nf2 c t))))).(\lambda (H0: (land (nf2 c 
-t0) (nfs2 c (TApp t1 t)))).(let H1 \def H0 in (and_ind (nf2 c t0) (nfs2 c 
-(TApp t1 t)) (land (land (nf2 c t0) (nfs2 c t1)) (nf2 c t)) (\lambda (H2: 
-(nf2 c t0)).(\lambda (H3: (nfs2 c (TApp t1 t))).(let H_x \def (H H3) in (let 
-H4 \def H_x in (and_ind (nfs2 c t1) (nf2 c t) (land (land (nf2 c t0) (nfs2 c 
-t1)) (nf2 c t)) (\lambda (H5: (nfs2 c t1)).(\lambda (H6: (nf2 c t)).(conj 
-(land (nf2 c t0) (nfs2 c t1)) (nf2 c t) (conj (nf2 c t0) (nfs2 c t1) H2 H5) 
-H6))) H4))))) H1)))))) ts))).
-
-theorem 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
- \lambda (c: C).(\lambda (t1: T).(\lambda (t2: T).(\lambda (H: (pc3 c t1 
-t2)).(\lambda (H0: (nf2 c t2)).(let H1 \def H in (ex2_ind T (\lambda (t: 
-T).(pr3 c t1 t)) (\lambda (t: T).(pr3 c t2 t)) (pr3 c t1 t2) (\lambda (x: 
-T).(\lambda (H2: (pr3 c t1 x)).(\lambda (H3: (pr3 c t2 x)).(let H_y \def 
-(nf2_pr3_unfold c t2 x H3 H0) in (let H4 \def (eq_ind_r T x (\lambda (t: 
-T).(pr3 c t1 t)) H2 t2 H_y) in H4))))) H1)))))).
-
-theorem 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
- \lambda (c: C).(\lambda (t: T).(\lambda (t1: T).(\lambda (H: (pc3 c t 
-t1)).(\lambda (t2: T).(\lambda (H0: (pr3 c t t2)).(pc3_t t c t2 (pc3_pr3_x c 
-t2 t H0) t1 H)))))).
-
 axiom pc3_gen_appls_sort_abst:
  \forall (c: C).(\forall (vs: TList).(\forall (w: T).(\forall (u: T).(\forall 
 (n: nat).((pc3 c (THeads (Flat Appl) vs (TSort n)) (THead (Bind Abst) w u)) 
@@ -135,408 +36,3 @@ TList).(\forall (n: nat).((pc3 c (THeads (Flat Appl) vs (TLRef i)) (THeads
 (Flat Appl) ws (TSort n))) \to False))))))))
 .
 
-inductive tys3 (g: G) (c: C): TList \to (T \to Prop) \def
-| tys3_nil: \forall (u: T).(\forall (u0: T).((ty3 g c u u0) \to (tys3 g c 
-TNil u)))
-| tys3_cons: \forall (t: T).(\forall (u: T).((ty3 g c t u) \to (\forall (ts: 
-TList).((tys3 g c ts u) \to (tys3 g c (TCons t ts) u))))).
-
-theorem 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
- \lambda (g: G).(\lambda (c: C).(\lambda (u: T).(\lambda (H: (tys3 g c TNil 
-u)).(insert_eq TList TNil (\lambda (t: TList).(tys3 g c t u)) (\lambda (_: 
-TList).(ex T (\lambda (u0: T).(ty3 g c u u0)))) (\lambda (y: TList).(\lambda 
-(H0: (tys3 g c y u)).(tys3_ind g c (\lambda (t: TList).(\lambda (t0: T).((eq 
-TList t TNil) \to (ex T (\lambda (u0: T).(ty3 g c t0 u0)))))) (\lambda (u0: 
-T).(\lambda (u1: T).(\lambda (H1: (ty3 g c u0 u1)).(\lambda (_: (eq TList 
-TNil TNil)).(ex_intro T (\lambda (u2: T).(ty3 g c u0 u2)) u1 H1))))) (\lambda 
-(t: T).(\lambda (u0: T).(\lambda (_: (ty3 g c t u0)).(\lambda (ts: 
-TList).(\lambda (_: (tys3 g c ts u0)).(\lambda (_: (((eq TList ts TNil) \to 
-(ex T (\lambda (u1: T).(ty3 g c u0 u1)))))).(\lambda (H4: (eq TList (TCons t 
-ts) TNil)).(let H5 \def (eq_ind TList (TCons t ts) (\lambda (ee: 
-TList).(match ee in TList return (\lambda (_: TList).Prop) with [TNil 
-\Rightarrow False | (TCons _ _) \Rightarrow 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:
- \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)))))))
-\def
- \lambda (g: G).(\lambda (c: C).(\lambda (ts: TList).(\lambda (t: T).(\lambda 
-(u: T).(\lambda (H: (tys3 g c (TCons t ts) u)).(insert_eq TList (TCons t ts) 
-(\lambda (t0: TList).(tys3 g c t0 u)) (\lambda (_: TList).(land (ty3 g c t u) 
-(tys3 g c ts u))) (\lambda (y: TList).(\lambda (H0: (tys3 g c y u)).(tys3_ind 
-g c (\lambda (t0: TList).(\lambda (t1: T).((eq TList t0 (TCons t ts)) \to 
-(land (ty3 g c t t1) (tys3 g c ts t1))))) (\lambda (u0: T).(\lambda (u1: 
-T).(\lambda (_: (ty3 g c u0 u1)).(\lambda (H2: (eq TList TNil (TCons t 
-ts))).(let H3 \def (eq_ind TList TNil (\lambda (ee: TList).(match ee in TList 
-return (\lambda (_: TList).Prop) with [TNil \Rightarrow True | (TCons _ _) 
-\Rightarrow False])) I (TCons t ts) H2) in (False_ind (land (ty3 g c t u0) 
-(tys3 g c ts u0)) H3)))))) (\lambda (t0: T).(\lambda (u0: T).(\lambda (H1: 
-(ty3 g c t0 u0)).(\lambda (ts0: TList).(\lambda (H2: (tys3 g c ts0 
-u0)).(\lambda (H3: (((eq TList ts0 (TCons t ts)) \to (land (ty3 g c t u0) 
-(tys3 g c ts u0))))).(\lambda (H4: (eq TList (TCons t0 ts0) (TCons t 
-ts))).(let H5 \def (f_equal TList T (\lambda (e: TList).(match e in TList 
-return (\lambda (_: TList).T) with [TNil \Rightarrow t0 | (TCons t1 _) 
-\Rightarrow t1])) (TCons t0 ts0) (TCons t ts) H4) in ((let H6 \def (f_equal 
-TList TList (\lambda (e: TList).(match e in TList return (\lambda (_: 
-TList).TList) with [TNil \Rightarrow ts0 | (TCons _ t1) \Rightarrow t1])) 
-(TCons t0 ts0) (TCons t ts) H4) in (\lambda (H7: (eq T t0 t)).(let H8 \def 
-(eq_ind TList ts0 (\lambda (t1: TList).((eq TList t1 (TCons t ts)) \to (land 
-(ty3 g c t u0) (tys3 g c ts u0)))) H3 ts H6) in (let H9 \def (eq_ind TList 
-ts0 (\lambda (t1: TList).(tys3 g c t1 u0)) H2 ts H6) in (let H10 \def (eq_ind 
-T t0 (\lambda (t1: T).(ty3 g c t1 u0)) H1 t H7) in (conj (ty3 g c t u0) (tys3 
-g c ts u0) H10 H9)))))) H5))))))))) y u H0))) H)))))).
-
-theorem 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))) 
-(\lambda (u: T).(\lambda (t: T).(ty3 g c v (THead (Bind Abst) u t)))) 
-(\lambda (u: T).(\lambda (_: T).(ty3 g c w u))) (\lambda (u: T).(\lambda (t: 
-T).(nf2 c (THead (Bind Abst) u t))))))))))
-\def
- \lambda (g: G).(\lambda (c: C).(\lambda (w: T).(\lambda (v: T).(\lambda (x: 
-T).(\lambda (H: (ty3 g c (THead (Flat Appl) w v) x)).(ex3_2_ind T T (\lambda 
-(u: T).(\lambda (t: T).(pc3 c (THead (Flat Appl) w (THead (Bind Abst) u t)) 
-x))) (\lambda (u: T).(\lambda (t: T).(ty3 g c v (THead (Bind Abst) u t)))) 
-(\lambda (u: T).(\lambda (_: T).(ty3 g c w u))) (ex4_2 T T (\lambda (u: 
-T).(\lambda (t: T).(pc3 c (THead (Flat Appl) w (THead (Bind Abst) u t)) x))) 
-(\lambda (u: T).(\lambda (t: T).(ty3 g c v (THead (Bind Abst) u t)))) 
-(\lambda (u: T).(\lambda (_: T).(ty3 g c w u))) (\lambda (u: T).(\lambda (t: 
-T).(nf2 c (THead (Bind Abst) u t))))) (\lambda (x0: T).(\lambda (x1: 
-T).(\lambda (H0: (pc3 c (THead (Flat Appl) w (THead (Bind Abst) x0 x1)) 
-x)).(\lambda (H1: (ty3 g c v (THead (Bind Abst) x0 x1))).(\lambda (H2: (ty3 g 
-c w x0)).(let H_x \def (ty3_correct g c v (THead (Bind Abst) x0 x1) H1) in 
-(let H3 \def H_x in (ex_ind T (\lambda (t: T).(ty3 g c (THead (Bind Abst) x0 
-x1) t)) (ex4_2 T T (\lambda (u: T).(\lambda (t: T).(pc3 c (THead (Flat Appl) 
-w (THead (Bind Abst) u t)) x))) (\lambda (u: T).(\lambda (t: T).(ty3 g c v 
-(THead (Bind Abst) u t)))) (\lambda (u: T).(\lambda (_: T).(ty3 g c w u))) 
-(\lambda (u: T).(\lambda (t: T).(nf2 c (THead (Bind Abst) u t))))) (\lambda 
-(x2: T).(\lambda (H4: (ty3 g c (THead (Bind Abst) x0 x1) x2)).(let H_x0 \def 
-(ty3_correct g c w x0 H2) in (let H5 \def H_x0 in (ex_ind T (\lambda (t: 
-T).(ty3 g c x0 t)) (ex4_2 T T (\lambda (u: T).(\lambda (t: T).(pc3 c (THead 
-(Flat Appl) w (THead (Bind Abst) u t)) x))) (\lambda (u: T).(\lambda (t: 
-T).(ty3 g c v (THead (Bind Abst) u t)))) (\lambda (u: T).(\lambda (_: T).(ty3 
-g c w u))) (\lambda (u: T).(\lambda (t: T).(nf2 c (THead (Bind Abst) u t))))) 
-(\lambda (x3: T).(\lambda (H6: (ty3 g c x0 x3)).(let H7 \def (ty3_sn3 g c 
-(THead (Bind Abst) x0 x1) x2 H4) in (let H_x1 \def (nf2_sn3 c (THead (Bind 
-Abst) x0 x1) H7) in (let H8 \def H_x1 in (ex2_ind T (\lambda (u: T).(pr3 c 
-(THead (Bind Abst) x0 x1) u)) (\lambda (u: T).(nf2 c u)) (ex4_2 T T (\lambda 
-(u: T).(\lambda (t: T).(pc3 c (THead (Flat Appl) w (THead (Bind Abst) u t)) 
-x))) (\lambda (u: T).(\lambda (t: T).(ty3 g c v (THead (Bind Abst) u t)))) 
-(\lambda (u: T).(\lambda (_: T).(ty3 g c w u))) (\lambda (u: T).(\lambda (t: 
-T).(nf2 c (THead (Bind Abst) u t))))) (\lambda (x4: T).(\lambda (H9: (pr3 c 
-(THead (Bind Abst) x0 x1) x4)).(\lambda (H10: (nf2 c x4)).(let H11 \def 
-(pr3_gen_abst c x0 x1 x4 H9) in (ex3_2_ind T T (\lambda (u2: T).(\lambda (t2: 
-T).(eq T x4 (THead (Bind Abst) u2 t2)))) (\lambda (u2: T).(\lambda (_: 
-T).(pr3 c x0 u2))) (\lambda (_: T).(\lambda (t2: T).(\forall (b: B).(\forall 
-(u: T).(pr3 (CHead c (Bind b) u) x1 t2))))) (ex4_2 T T (\lambda (u: 
-T).(\lambda (t: T).(pc3 c (THead (Flat Appl) w (THead (Bind Abst) u t)) x))) 
-(\lambda (u: T).(\lambda (t: T).(ty3 g c v (THead (Bind Abst) u t)))) 
-(\lambda (u: T).(\lambda (_: T).(ty3 g c w u))) (\lambda (u: T).(\lambda (t: 
-T).(nf2 c (THead (Bind Abst) u t))))) (\lambda (x5: T).(\lambda (x6: 
-T).(\lambda (H12: (eq T x4 (THead (Bind Abst) x5 x6))).(\lambda (H13: (pr3 c 
-x0 x5)).(\lambda (H14: ((\forall (b: B).(\forall (u: T).(pr3 (CHead c (Bind 
-b) u) x1 x6))))).(let H15 \def (eq_ind T x4 (\lambda (t: T).(nf2 c t)) H10 
-(THead (Bind Abst) x5 x6) H12) in (let H16 \def (pr3_head_12 c x0 x5 H13 
-(Bind Abst) x1 x6 (H14 Abst x5)) in (ex4_2_intro T T (\lambda (u: T).(\lambda 
-(t: T).(pc3 c (THead (Flat Appl) w (THead (Bind Abst) u t)) x))) (\lambda (u: 
-T).(\lambda (t: T).(ty3 g c v (THead (Bind Abst) u t)))) (\lambda (u: 
-T).(\lambda (_: T).(ty3 g c w u))) (\lambda (u: T).(\lambda (t: T).(nf2 c 
-(THead (Bind Abst) u t)))) x5 x6 (pc3_pr3_conf c (THead (Flat Appl) w (THead 
-(Bind Abst) x0 x1)) x H0 (THead (Flat Appl) w (THead (Bind Abst) x5 x6)) 
-(pr3_thin_dx c (THead (Bind Abst) x0 x1) (THead (Bind Abst) x5 x6) H16 w 
-Appl)) (ty3_conv g c (THead (Bind Abst) x5 x6) x2 (ty3_sred_pr3 c (THead 
-(Bind Abst) x0 x1) (THead (Bind Abst) x5 x6) H16 g x2 H4) v (THead (Bind 
-Abst) x0 x1) H1 (pc3_pr3_r c (THead (Bind Abst) x0 x1) (THead (Bind Abst) x5 
-x6) H16)) (ty3_conv g c x5 x3 (ty3_sred_pr3 c x0 x5 H13 g x3 H6) w x0 H2 
-(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:
- \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))))))))))))
-\def
- \lambda (g: G).(\lambda (c: C).(\lambda (u1: T).(\lambda (i: nat).(\lambda 
-(H: (ty3 g c (TLRef i) u1)).(insert_eq T (TLRef i) (\lambda (t: T).(ty3 g c t 
-u1)) (\lambda (t: T).((nf2 c t) \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))))))))) (\lambda 
-(y: T).(\lambda (H0: (ty3 g c y u1)).(ty3_ind g (\lambda (c0: C).(\lambda (t: 
-T).(\lambda (t0: T).((eq T t (TLRef i)) \to ((nf2 c0 t) \to (\forall (u2: 
-T).((nf2 c0 u2) \to ((pc3 c0 t0 u2) \to (ex T (\lambda (u: T).(eq T u2 (lift 
-(S i) O u)))))))))))) (\lambda (c0: C).(\lambda (t2: T).(\lambda (t: 
-T).(\lambda (_: (ty3 g c0 t2 t)).(\lambda (_: (((eq T t2 (TLRef i)) \to ((nf2 
-c0 t2) \to (\forall (u2: T).((nf2 c0 u2) \to ((pc3 c0 t u2) \to (ex T 
-(\lambda (u: T).(eq T u2 (lift (S i) O u))))))))))).(\lambda (u: T).(\lambda 
-(t1: T).(\lambda (H3: (ty3 g c0 u t1)).(\lambda (H4: (((eq T u (TLRef i)) \to 
-((nf2 c0 u) \to (\forall (u2: T).((nf2 c0 u2) \to ((pc3 c0 t1 u2) \to (ex T 
-(\lambda (u0: T).(eq T u2 (lift (S i) O u0))))))))))).(\lambda (H5: (pc3 c0 
-t1 t2)).(\lambda (H6: (eq T u (TLRef i))).(\lambda (H7: (nf2 c0 u)).(\lambda 
-(u2: T).(\lambda (H8: (nf2 c0 u2)).(\lambda (H9: (pc3 c0 t2 u2)).(let H10 
-\def (eq_ind T u (\lambda (t0: T).(nf2 c0 t0)) H7 (TLRef i) H6) in (let H11 
-\def (eq_ind T u (\lambda (t0: T).((eq T t0 (TLRef i)) \to ((nf2 c0 t0) \to 
-(\forall (u3: T).((nf2 c0 u3) \to ((pc3 c0 t1 u3) \to (ex T (\lambda (u0: 
-T).(eq T u3 (lift (S i) O u0)))))))))) H4 (TLRef i) H6) in (let H12 \def 
-(eq_ind T u (\lambda (t0: T).(ty3 g c0 t0 t1)) H3 (TLRef i) H6) in (let H_y 
-\def (H11 (refl_equal T (TLRef i)) H10 u2 H8) in (H_y (pc3_t t2 c0 t1 H5 u2 
-H9))))))))))))))))))))) (\lambda (c0: C).(\lambda (m: nat).(\lambda (H1: (eq 
-T (TSort m) (TLRef i))).(\lambda (_: (nf2 c0 (TSort m))).(\lambda (u2: 
-T).(\lambda (_: (nf2 c0 u2)).(\lambda (_: (pc3 c0 (TSort (next g m)) 
-u2)).(let H5 \def (eq_ind T (TSort m) (\lambda (ee: T).(match ee in T return 
-(\lambda (_: T).Prop) with [(TSort _) \Rightarrow True | (TLRef _) 
-\Rightarrow False | (THead _ _ _) \Rightarrow False])) I (TLRef i) H1) in 
-(False_ind (ex T (\lambda (u: T).(eq T u2 (lift (S i) O u)))) H5))))))))) 
-(\lambda (n: nat).(\lambda (c0: C).(\lambda (d: C).(\lambda (u: T).(\lambda 
-(H1: (getl n c0 (CHead d (Bind Abbr) u))).(\lambda (t: T).(\lambda (_: (ty3 g 
-d u t)).(\lambda (_: (((eq T u (TLRef i)) \to ((nf2 d u) \to (\forall (u2: 
-T).((nf2 d u2) \to ((pc3 d t u2) \to (ex T (\lambda (u0: T).(eq T u2 (lift (S 
-i) O u0))))))))))).(\lambda (H4: (eq T (TLRef n) (TLRef i))).(\lambda (H5: 
-(nf2 c0 (TLRef n))).(\lambda (u2: T).(\lambda (_: (nf2 c0 u2)).(\lambda (H7: 
-(pc3 c0 (lift (S n) O t) u2)).(let H8 \def (f_equal T nat (\lambda (e: 
-T).(match e in T return (\lambda (_: T).nat) with [(TSort _) \Rightarrow n | 
-(TLRef n0) \Rightarrow n0 | (THead _ _ _) \Rightarrow n])) (TLRef n) (TLRef 
-i) H4) in (let H9 \def (eq_ind nat n (\lambda (n0: nat).(pc3 c0 (lift (S n0) 
-O t) u2)) H7 i H8) in (let H10 \def (eq_ind nat n (\lambda (n0: nat).(nf2 c0 
-(TLRef n0))) H5 i H8) in (let H11 \def (eq_ind nat n (\lambda (n0: nat).(getl 
-n0 c0 (CHead d (Bind Abbr) u))) H1 i H8) in (nf2_gen_lref c0 d u i H11 H10 
-(ex T (\lambda (u0: T).(eq T u2 (lift (S i) O u0)))))))))))))))))))))) 
-(\lambda (n: nat).(\lambda (c0: C).(\lambda (d: C).(\lambda (u: T).(\lambda 
-(H1: (getl n c0 (CHead d (Bind Abst) u))).(\lambda (t: T).(\lambda (_: (ty3 g 
-d u t)).(\lambda (_: (((eq T u (TLRef i)) \to ((nf2 d u) \to (\forall (u2: 
-T).((nf2 d u2) \to ((pc3 d t u2) \to (ex T (\lambda (u0: T).(eq T u2 (lift (S 
-i) O u0))))))))))).(\lambda (H4: (eq T (TLRef n) (TLRef i))).(\lambda (H5: 
-(nf2 c0 (TLRef n))).(\lambda (u2: T).(\lambda (H6: (nf2 c0 u2)).(\lambda (H7: 
-(pc3 c0 (lift (S n) O u) u2)).(let H8 \def (f_equal T nat (\lambda (e: 
-T).(match e in T return (\lambda (_: T).nat) with [(TSort _) \Rightarrow n | 
-(TLRef n0) \Rightarrow n0 | (THead _ _ _) \Rightarrow n])) (TLRef n) (TLRef 
-i) H4) in (let H9 \def (eq_ind nat n (\lambda (n0: nat).(pc3 c0 (lift (S n0) 
-O u) u2)) H7 i H8) in (let H10 \def (eq_ind nat n (\lambda (n0: nat).(nf2 c0 
-(TLRef n0))) H5 i H8) in (let H11 \def (eq_ind nat n (\lambda (n0: nat).(getl 
-n0 c0 (CHead d (Bind Abst) u))) H1 i H8) in (let H_y \def (pc3_nf2_unfold c0 
-(lift (S i) O u) u2 H9 H6) in (let H12 \def (pr3_gen_lift c0 u u2 (S i) O H_y 
-d (getl_drop Abst c0 d u i H11)) in (ex2_ind T (\lambda (t2: T).(eq T u2 
-(lift (S i) O t2))) (\lambda (t2: T).(pr3 d u t2)) (ex T (\lambda (u0: T).(eq 
-T u2 (lift (S i) O u0)))) (\lambda (x: T).(\lambda (H13: (eq T u2 (lift (S i) 
-O x))).(\lambda (_: (pr3 d u x)).(eq_ind_r T (lift (S i) O x) (\lambda (t0: 
-T).(ex T (\lambda (u0: T).(eq T t0 (lift (S i) O u0))))) (ex_intro T (\lambda 
-(u0: T).(eq T (lift (S i) O x) (lift (S i) O u0))) x (refl_equal T (lift (S 
-i) O x))) u2 H13)))) H12)))))))))))))))))))) (\lambda (c0: C).(\lambda (u: 
-T).(\lambda (t: T).(\lambda (_: (ty3 g c0 u t)).(\lambda (_: (((eq T u (TLRef 
-i)) \to ((nf2 c0 u) \to (\forall (u2: T).((nf2 c0 u2) \to ((pc3 c0 t u2) \to 
-(ex T (\lambda (u0: T).(eq T u2 (lift (S i) O u0))))))))))).(\lambda (b: 
-B).(\lambda (t1: T).(\lambda (t2: T).(\lambda (_: (ty3 g (CHead c0 (Bind b) 
-u) t1 t2)).(\lambda (_: (((eq T t1 (TLRef i)) \to ((nf2 (CHead c0 (Bind b) u) 
-t1) \to (\forall (u2: T).((nf2 (CHead c0 (Bind b) u) u2) \to ((pc3 (CHead c0 
-(Bind b) u) t2 u2) \to (ex T (\lambda (u0: T).(eq T u2 (lift (S i) O 
-u0))))))))))).(\lambda (H5: (eq T (THead (Bind b) u t1) (TLRef i))).(\lambda 
-(_: (nf2 c0 (THead (Bind b) u t1))).(\lambda (u2: T).(\lambda (_: (nf2 c0 
-u2)).(\lambda (_: (pc3 c0 (THead (Bind b) u t2) u2)).(let H9 \def (eq_ind T 
-(THead (Bind b) u t1) (\lambda (ee: T).(match ee in T return (\lambda (_: 
-T).Prop) with [(TSort _) \Rightarrow False | (TLRef _) \Rightarrow False | 
-(THead _ _ _) \Rightarrow True])) I (TLRef i) H5) in (False_ind (ex T 
-(\lambda (u0: T).(eq T u2 (lift (S i) O u0)))) H9))))))))))))))))) (\lambda 
-(c0: C).(\lambda (w: T).(\lambda (u: T).(\lambda (_: (ty3 g c0 w u)).(\lambda 
-(_: (((eq T w (TLRef i)) \to ((nf2 c0 w) \to (\forall (u2: T).((nf2 c0 u2) 
-\to ((pc3 c0 u u2) \to (ex T (\lambda (u0: T).(eq T u2 (lift (S i) O 
-u0))))))))))).(\lambda (v: T).(\lambda (t: T).(\lambda (_: (ty3 g c0 v (THead 
-(Bind Abst) u t))).(\lambda (_: (((eq T v (TLRef i)) \to ((nf2 c0 v) \to 
-(\forall (u2: T).((nf2 c0 u2) \to ((pc3 c0 (THead (Bind Abst) u t) u2) \to 
-(ex T (\lambda (u0: T).(eq T u2 (lift (S i) O u0))))))))))).(\lambda (H5: (eq 
-T (THead (Flat Appl) w v) (TLRef i))).(\lambda (_: (nf2 c0 (THead (Flat Appl) 
-w v))).(\lambda (u2: T).(\lambda (_: (nf2 c0 u2)).(\lambda (_: (pc3 c0 (THead 
-(Flat Appl) w (THead (Bind Abst) u t)) u2)).(let H9 \def (eq_ind T (THead 
-(Flat Appl) w v) (\lambda (ee: T).(match ee in T return (\lambda (_: T).Prop) 
-with [(TSort _) \Rightarrow False | (TLRef _) \Rightarrow False | (THead _ _ 
-_) \Rightarrow True])) I (TLRef i) H5) in (False_ind (ex T (\lambda (u0: 
-T).(eq T u2 (lift (S i) O u0)))) H9)))))))))))))))) (\lambda (c0: C).(\lambda 
-(t1: T).(\lambda (t2: T).(\lambda (_: (ty3 g c0 t1 t2)).(\lambda (_: (((eq T 
-t1 (TLRef i)) \to ((nf2 c0 t1) \to (\forall (u2: T).((nf2 c0 u2) \to ((pc3 c0 
-t2 u2) \to (ex T (\lambda (u: T).(eq T u2 (lift (S i) O u))))))))))).(\lambda 
-(t0: T).(\lambda (_: (ty3 g c0 t2 t0)).(\lambda (_: (((eq T t2 (TLRef i)) \to 
-((nf2 c0 t2) \to (\forall (u2: T).((nf2 c0 u2) \to ((pc3 c0 t0 u2) \to (ex T 
-(\lambda (u: T).(eq T u2 (lift (S i) O u))))))))))).(\lambda (H5: (eq T 
-(THead (Flat Cast) t2 t1) (TLRef i))).(\lambda (_: (nf2 c0 (THead (Flat Cast) 
-t2 t1))).(\lambda (u2: T).(\lambda (_: (nf2 c0 u2)).(\lambda (_: (pc3 c0 
-(THead (Flat Cast) t0 t2) u2)).(let H9 \def (eq_ind T (THead (Flat Cast) t2 
-t1) (\lambda (ee: T).(match ee in T return (\lambda (_: T).Prop) with [(TSort 
-_) \Rightarrow False | (TLRef _) \Rightarrow False | (THead _ _ _) 
-\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:
- \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))))))))))
-\def
- \lambda (g: G).(\lambda (c: C).(\lambda (u: T).(\lambda (i: nat).(\lambda 
-(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:
- \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 
-i) O u))) (\lambda (u: T).(pc3 c (THeads (Flat Appl) vs (lift (S i) O u)) 
-u1))))))))))
-\def
- \lambda (g: G).(\lambda (c: C).(\lambda (vs: TList).(TList_ind (\lambda (t: 
-TList).(\forall (u1: T).(\forall (i: nat).((ty3 g c (THeads (Flat Appl) t 
-(TLRef i)) u1) \to ((nf2 c (TLRef i)) \to ((nf2 c u1) \to (ex2 T (\lambda (u: 
-T).(nf2 c (lift (S i) O u))) (\lambda (u: T).(pc3 c (THeads (Flat Appl) t 
-(lift (S i) O u)) u1))))))))) (\lambda (u1: T).(\lambda (i: nat).(\lambda (H: 
-(ty3 g c (TLRef i) u1)).(\lambda (H0: (nf2 c (TLRef i))).(\lambda (H1: (nf2 c 
-u1)).(let H_x \def (ty3_inv_lref_nf2 g c u1 i H H0 H1) in (let H2 \def H_x in 
-(ex_ind T (\lambda (u0: T).(eq T u1 (lift (S i) O u0))) (ex2 T (\lambda (u: 
-T).(nf2 c (lift (S i) O u))) (\lambda (u: T).(pc3 c (lift (S i) O u) u1))) 
-(\lambda (x: T).(\lambda (H3: (eq T u1 (lift (S i) O x))).(let H4 \def 
-(eq_ind T u1 (\lambda (t: T).(nf2 c t)) H1 (lift (S i) O x) H3) in (eq_ind_r 
-T (lift (S i) O x) (\lambda (t: T).(ex2 T (\lambda (u: T).(nf2 c (lift (S i) 
-O u))) (\lambda (u: T).(pc3 c (lift (S i) O u) t)))) (ex_intro2 T (\lambda 
-(u: T).(nf2 c (lift (S i) O u))) (\lambda (u: T).(pc3 c (lift (S i) O u) 
-(lift (S i) O x))) x H4 (pc3_refl c (lift (S i) O x))) u1 H3)))) H2)))))))) 
-(\lambda (t: T).(\lambda (t0: TList).(\lambda (H: ((\forall (u1: T).(\forall 
-(i: nat).((ty3 g c (THeads (Flat Appl) t0 (TLRef i)) u1) \to ((nf2 c (TLRef 
-i)) \to ((nf2 c u1) \to (ex2 T (\lambda (u: T).(nf2 c (lift (S i) O u))) 
-(\lambda (u: T).(pc3 c (THeads (Flat Appl) t0 (lift (S i) O u)) 
-u1)))))))))).(\lambda (u1: T).(\lambda (i: nat).(\lambda (H0: (ty3 g c (THead 
-(Flat Appl) t (THeads (Flat Appl) t0 (TLRef i))) u1)).(\lambda (H1: (nf2 c 
-(TLRef i))).(\lambda (_: (nf2 c u1)).(let H_x \def (ty3_gen_appl_nf2 g c t 
-(THeads (Flat Appl) t0 (TLRef i)) u1 H0) in (let H3 \def H_x in (ex4_2_ind T 
-T (\lambda (u: T).(\lambda (t1: T).(pc3 c (THead (Flat Appl) t (THead (Bind 
-Abst) u t1)) u1))) (\lambda (u: T).(\lambda (t1: T).(ty3 g c (THeads (Flat 
-Appl) t0 (TLRef i)) (THead (Bind Abst) u t1)))) (\lambda (u: T).(\lambda (_: 
-T).(ty3 g c t u))) (\lambda (u: T).(\lambda (t1: T).(nf2 c (THead (Bind Abst) 
-u t1)))) (ex2 T (\lambda (u: T).(nf2 c (lift (S i) O u))) (\lambda (u: 
-T).(pc3 c (THead (Flat Appl) t (THeads (Flat Appl) t0 (lift (S i) O u))) 
-u1))) (\lambda (x0: T).(\lambda (x1: T).(\lambda (H4: (pc3 c (THead (Flat 
-Appl) t (THead (Bind Abst) x0 x1)) u1)).(\lambda (H5: (ty3 g c (THeads (Flat 
-Appl) t0 (TLRef i)) (THead (Bind Abst) x0 x1))).(\lambda (_: (ty3 g c t 
-x0)).(\lambda (H7: (nf2 c (THead (Bind Abst) x0 x1))).(let H8 \def 
-(nf2_gen_abst c x0 x1 H7) in (and_ind (nf2 c x0) (nf2 (CHead c (Bind Abst) 
-x0) x1) (ex2 T (\lambda (u: T).(nf2 c (lift (S i) O u))) (\lambda (u: T).(pc3 
-c (THead (Flat Appl) t (THeads (Flat Appl) t0 (lift (S i) O u))) u1))) 
-(\lambda (H9: (nf2 c x0)).(\lambda (H10: (nf2 (CHead c (Bind Abst) x0) 
-x1)).(let H_y \def (H (THead (Bind Abst) x0 x1) i H5 H1) in (let H11 \def 
-(H_y (nf2_abst_shift c x0 H9 x1 H10)) in (ex2_ind T (\lambda (u: T).(nf2 c 
-(lift (S i) O u))) (\lambda (u: T).(pc3 c (THeads (Flat Appl) t0 (lift (S i) 
-O u)) (THead (Bind Abst) x0 x1))) (ex2 T (\lambda (u: T).(nf2 c (lift (S i) O 
-u))) (\lambda (u: T).(pc3 c (THead (Flat Appl) t (THeads (Flat Appl) t0 (lift 
-(S i) O u))) u1))) (\lambda (x: T).(\lambda (H12: (nf2 c (lift (S i) O 
-x))).(\lambda (H13: (pc3 c (THeads (Flat Appl) t0 (lift (S i) O x)) (THead 
-(Bind Abst) x0 x1))).(ex_intro2 T (\lambda (u: T).(nf2 c (lift (S i) O u))) 
-(\lambda (u: T).(pc3 c (THead (Flat Appl) t (THeads (Flat Appl) t0 (lift (S 
-i) O u))) u1)) x H12 (pc3_t (THead (Flat Appl) t (THead (Bind Abst) x0 x1)) c 
-(THead (Flat Appl) t (THeads (Flat Appl) t0 (lift (S i) O x))) (pc3_thin_dx c 
-(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:
- \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)))))))
-\def
- \lambda (g: G).(\lambda (c: C).(\lambda (i: nat).(\lambda (j: nat).(\lambda 
-(H: (ty3 g c (TLRef i) (TLRef j))).(\lambda (H0: (nf2 c (TLRef i))).(\lambda 
-(H1: (nf2 c (TLRef j))).(let H_x \def (ty3_inv_lref_nf2 g c (TLRef j) i H H0 
-H1) in (let H2 \def H_x in (ex_ind T (\lambda (u0: T).(eq T (TLRef j) (lift 
-(S i) O u0))) (lt i j) (\lambda (x: T).(\lambda (H3: (eq T (TLRef j) (lift (S 
-i) O x))).(let H_x0 \def (lift_gen_lref x O (S i) j H3) in (let H4 \def H_x0 
-in (or_ind (land (lt j O) (eq T x (TLRef j))) (land (le (plus O (S i)) j) (eq 
-T x (TLRef (minus j (S i))))) (lt i j) (\lambda (H5: (land (lt j O) (eq T x 
-(TLRef j)))).(and_ind (lt j O) (eq T x (TLRef j)) (lt i j) (\lambda (H6: (lt 
-j O)).(\lambda (_: (eq T x (TLRef j))).(lt_x_O j H6 (lt i j)))) H5)) (\lambda 
-(H5: (land (le (plus O (S i)) j) (eq T x (TLRef (minus j (S i)))))).(and_ind 
-(le (plus O (S i)) j) (eq T x (TLRef (minus j (S i)))) (lt i j) (\lambda (H6: 
-(le (plus O (S i)) j)).(\lambda (_: (eq T x (TLRef (minus j (S i))))).H6)) 
-H5)) H4))))) H2))))))))).
-
-theorem 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)))))))
-\def
- \lambda (g: G).(\lambda (c: C).(\lambda (H: (wf3 g c c)).(insert_eq C c 
-(\lambda (c0: C).(wf3 g c0 c)) (\lambda (c0: C).(\forall (t1: T).(\forall 
-(t2: T).((ty3 g c0 t1 t2) \to (ty3 g (CSort (cbk c0)) (app1 c0 t1) (app1 c0 
-t2)))))) (\lambda (y: C).(\lambda (H0: (wf3 g y c)).(wf3_ind g (\lambda (c0: 
-C).(\lambda (c1: C).((eq C c0 c1) \to (\forall (t1: T).(\forall (t2: T).((ty3 
-g c0 t1 t2) \to (ty3 g (CSort (cbk c0)) (app1 c0 t1) (app1 c0 t2)))))))) 
-(\lambda (m: nat).(\lambda (_: (eq C (CSort m) (CSort m))).(\lambda (t1: 
-T).(\lambda (t2: T).(\lambda (H2: (ty3 g (CSort m) t1 t2)).H2))))) (\lambda 
-(c1: C).(\lambda (c2: C).(\lambda (H1: (wf3 g c1 c2)).(\lambda (H2: (((eq C 
-c1 c2) \to (\forall (t1: T).(\forall (t2: T).((ty3 g c1 t1 t2) \to (ty3 g 
-(CSort (cbk c1)) (app1 c1 t1) (app1 c1 t2)))))))).(\lambda (u: T).(\lambda 
-(t: T).(\lambda (H3: (ty3 g c1 u t)).(\lambda (b: B).(\lambda (H4: (eq C 
-(CHead c1 (Bind b) u) (CHead c2 (Bind b) u))).(\lambda (t1: T).(\lambda (t2: 
-T).(\lambda (H5: (ty3 g (CHead c1 (Bind b) u) t1 t2)).(let H6 \def (f_equal C 
-C (\lambda (e: C).(match e in C return (\lambda (_: C).C) with [(CSort _) 
-\Rightarrow c1 | (CHead c0 _ _) \Rightarrow c0])) (CHead c1 (Bind b) u) 
-(CHead c2 (Bind b) u) H4) in (let H7 \def (eq_ind_r C c2 (\lambda (c0: 
-C).((eq C c1 c0) \to (\forall (t3: T).(\forall (t4: T).((ty3 g c1 t3 t4) \to 
-(ty3 g (CSort (cbk c1)) (app1 c1 t3) (app1 c1 t4))))))) H2 c1 H6) in (let H8 
-\def (eq_ind_r C c2 (\lambda (c0: C).(wf3 g c1 c0)) H1 c1 H6) in (ex_ind T 
-(\lambda (t0: T).(ty3 g (CHead c1 (Bind b) u) t2 t0)) (ty3 g (CSort (cbk c1)) 
-(app1 c1 (THead (Bind b) u t1)) (app1 c1 (THead (Bind b) u t2))) (\lambda (x: 
-T).(\lambda (_: (ty3 g (CHead c1 (Bind b) u) t2 x)).(H7 (refl_equal C c1) 
-(THead (Bind b) u t1) (THead (Bind b) u t2) (ty3_bind g c1 u t H3 b t1 t2 
-H5)))) (ty3_correct g (CHead c1 (Bind b) u) t1 t2 H5))))))))))))))))) 
-(\lambda (c1: C).(\lambda (c2: C).(\lambda (H1: (wf3 g c1 c2)).(\lambda (H2: 
-(((eq C c1 c2) \to (\forall (t1: T).(\forall (t2: T).((ty3 g c1 t1 t2) \to 
-(ty3 g (CSort (cbk c1)) (app1 c1 t1) (app1 c1 t2)))))))).(\lambda (u: 
-T).(\lambda (H3: ((\forall (t: T).((ty3 g c1 u t) \to False)))).(\lambda (b: 
-B).(\lambda (H4: (eq C (CHead c1 (Bind b) u) (CHead c2 (Bind Void) (TSort 
-O)))).(\lambda (t1: T).(\lambda (t2: T).(\lambda (H5: (ty3 g (CHead c1 (Bind 
-b) u) t1 t2)).(let H6 \def (f_equal C C (\lambda (e: C).(match e in C return 
-(\lambda (_: C).C) with [(CSort _) \Rightarrow c1 | (CHead c0 _ _) 
-\Rightarrow c0])) (CHead c1 (Bind b) u) (CHead c2 (Bind Void) (TSort O)) H4) 
-in ((let H7 \def (f_equal C B (\lambda (e: C).(match e in C return (\lambda 
-(_: C).B) with [(CSort _) \Rightarrow b | (CHead _ k _) \Rightarrow (match k 
-in K return (\lambda (_: K).B) with [(Bind b0) \Rightarrow b0 | (Flat _) 
-\Rightarrow b])])) (CHead c1 (Bind b) u) (CHead c2 (Bind Void) (TSort O)) H4) 
-in ((let H8 \def (f_equal C T (\lambda (e: C).(match e in C return (\lambda 
-(_: C).T) with [(CSort _) \Rightarrow u | (CHead _ _ t) \Rightarrow t])) 
-(CHead c1 (Bind b) u) (CHead c2 (Bind Void) (TSort O)) H4) in (\lambda (H9: 
-(eq B b Void)).(\lambda (H10: (eq C c1 c2)).(let H11 \def (eq_ind B b 
-(\lambda (b0: B).(ty3 g (CHead c1 (Bind b0) u) t1 t2)) H5 Void H9) in 
-(eq_ind_r B Void (\lambda (b0: B).(ty3 g (CSort (cbk (CHead c1 (Bind b0) u))) 
-(app1 (CHead c1 (Bind b0) u) t1) (app1 (CHead c1 (Bind b0) u) t2))) (let H12 
-\def (eq_ind T u (\lambda (t: T).(ty3 g (CHead c1 (Bind Void) t) t1 t2)) H11 
-(TSort O) H8) in (let H13 \def (eq_ind T u (\lambda (t: T).(\forall (t0: 
-T).((ty3 g c1 t t0) \to False))) H3 (TSort O) H8) in (eq_ind_r T (TSort O) 
-(\lambda (t: T).(ty3 g (CSort (cbk (CHead c1 (Bind Void) t))) (app1 (CHead c1 
-(Bind Void) t) t1) (app1 (CHead c1 (Bind Void) t) t2))) (let H14 \def 
-(eq_ind_r C c2 (\lambda (c0: C).((eq C c1 c0) \to (\forall (t3: T).(\forall 
-(t4: T).((ty3 g c1 t3 t4) \to (ty3 g (CSort (cbk c1)) (app1 c1 t3) (app1 c1 
-t4))))))) H2 c1 H10) in (let H15 \def (eq_ind_r C c2 (\lambda (c0: C).(wf3 g 
-c1 c0)) H1 c1 H10) in (ex_ind T (\lambda (t: T).(ty3 g (CHead c1 (Bind Void) 
-(TSort O)) t2 t)) (ty3 g (CSort (cbk c1)) (app1 c1 (THead (Bind Void) (TSort 
-O) t1)) (app1 c1 (THead (Bind Void) (TSort O) t2))) (\lambda (x: T).(\lambda 
-(_: (ty3 g (CHead c1 (Bind Void) (TSort O)) t2 x)).(H14 (refl_equal C c1) 
-(THead (Bind Void) (TSort O) t1) (THead (Bind Void) (TSort O) t2) (ty3_bind g 
-c1 (TSort O) (TSort (next g O)) (ty3_sort g c1 O) Void t1 t2 H12)))) 
-(ty3_correct g (CHead c1 (Bind Void) (TSort O)) t1 t2 H12)))) u H8))) b 
-H9))))) H7)) H6))))))))))))) (\lambda (c1: C).(\lambda (c2: C).(\lambda (H1: 
-(wf3 g c1 c2)).(\lambda (H2: (((eq C c1 c2) \to (\forall (t1: T).(\forall 
-(t2: T).((ty3 g c1 t1 t2) \to (ty3 g (CSort (cbk c1)) (app1 c1 t1) (app1 c1 
-t2)))))))).(\lambda (u: T).(\lambda (f: F).(\lambda (H3: (eq C (CHead c1 
-(Flat f) u) c2)).(\lambda (t1: T).(\lambda (t2: T).(\lambda (_: (ty3 g (CHead 
-c1 (Flat f) u) t1 t2)).(let H5 \def (f_equal C C (\lambda (e: C).e) (CHead c1 
-(Flat f) u) c2 H3) in (let H6 \def (eq_ind_r C c2 (\lambda (c0: C).((eq C c1 
-c0) \to (\forall (t3: T).(\forall (t4: T).((ty3 g c1 t3 t4) \to (ty3 g (CSort 
-(cbk c1)) (app1 c1 t3) (app1 c1 t4))))))) H2 (CHead c1 (Flat f) u) H5) in 
-(let H7 \def (eq_ind_r C c2 (\lambda (c0: C).(wf3 g c1 c0)) H1 (CHead c1 
-(Flat f) u) H5) in (let H_x \def (wf3_gen_head2 g c1 c1 u (Flat f) H7) in 
-(let H8 \def H_x in (ex_ind B (\lambda (b: B).(eq K (Flat f) (Bind b))) (ty3 
-g (CSort (cbk c1)) (app1 c1 (THead (Flat f) u t1)) (app1 c1 (THead (Flat f) u 
-t2))) (\lambda (x: B).(\lambda (H9: (eq K (Flat f) (Bind x))).(let H10 \def 
-(eq_ind K (Flat f) (\lambda (ee: K).(match ee in K return (\lambda (_: 
-K).Prop) with [(Bind _) \Rightarrow False | (Flat _) \Rightarrow True])) I 
-(Bind x) H9) in (False_ind (ty3 g (CSort (cbk c1)) (app1 c1 (THead (Flat f) u 
-t1)) (app1 c1 (THead (Flat f) u t2))) H10)))) H8)))))))))))))))) y c H0))) 
-H))).
-