]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_1A/fsubst0/fwd.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_1A / fsubst0 / fwd.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 (* This file was automatically generated: do not edit *********************)
16
17 include "basic_1A/fsubst0/defs.ma".
18
19 implied lemma fsubst0_ind:
20  \forall (i: nat).(\forall (v: T).(\forall (c1: C).(\forall (t1: T).(\forall 
21 (P: ((C \to (T \to Prop)))).(((\forall (t2: T).((subst0 i v t1 t2) \to (P c1 
22 t2)))) \to (((\forall (c2: C).((csubst0 i v c1 c2) \to (P c2 t1)))) \to 
23 (((\forall (t2: T).((subst0 i v t1 t2) \to (\forall (c2: C).((csubst0 i v c1 
24 c2) \to (P c2 t2)))))) \to (\forall (c: C).(\forall (t: T).((fsubst0 i v c1 
25 t1 c t) \to (P c t)))))))))))
26 \def
27  \lambda (i: nat).(\lambda (v: T).(\lambda (c1: C).(\lambda (t1: T).(\lambda 
28 (P: ((C \to (T \to Prop)))).(\lambda (f: ((\forall (t2: T).((subst0 i v t1 
29 t2) \to (P c1 t2))))).(\lambda (f0: ((\forall (c2: C).((csubst0 i v c1 c2) 
30 \to (P c2 t1))))).(\lambda (f1: ((\forall (t2: T).((subst0 i v t1 t2) \to 
31 (\forall (c2: C).((csubst0 i v c1 c2) \to (P c2 t2))))))).(\lambda (c: 
32 C).(\lambda (t: T).(\lambda (f2: (fsubst0 i v c1 t1 c t)).(match f2 with 
33 [(fsubst0_snd x x0) \Rightarrow (f x x0) | (fsubst0_fst x x0) \Rightarrow (f0 
34 x x0) | (fsubst0_both x x0 x1 x2) \Rightarrow (f1 x x0 x1 x2)]))))))))))).
35
36 lemma fsubst0_gen_base:
37  \forall (c1: C).(\forall (c2: C).(\forall (t1: T).(\forall (t2: T).(\forall 
38 (v: T).(\forall (i: nat).((fsubst0 i v c1 t1 c2 t2) \to (or3 (land (eq C c1 
39 c2) (subst0 i v t1 t2)) (land (eq T t1 t2) (csubst0 i v c1 c2)) (land (subst0 
40 i v t1 t2) (csubst0 i v c1 c2)))))))))
41 \def
42  \lambda (c1: C).(\lambda (c2: C).(\lambda (t1: T).(\lambda (t2: T).(\lambda 
43 (v: T).(\lambda (i: nat).(\lambda (H: (fsubst0 i v c1 t1 c2 t2)).(fsubst0_ind 
44 i v c1 t1 (\lambda (c: C).(\lambda (t: T).(or3 (land (eq C c1 c) (subst0 i v 
45 t1 t)) (land (eq T t1 t) (csubst0 i v c1 c)) (land (subst0 i v t1 t) (csubst0 
46 i v c1 c))))) (\lambda (t0: T).(\lambda (H0: (subst0 i v t1 t0)).(or3_intro0 
47 (land (eq C c1 c1) (subst0 i v t1 t0)) (land (eq T t1 t0) (csubst0 i v c1 
48 c1)) (land (subst0 i v t1 t0) (csubst0 i v c1 c1)) (conj (eq C c1 c1) (subst0 
49 i v t1 t0) (refl_equal C c1) H0)))) (\lambda (c0: C).(\lambda (H0: (csubst0 i 
50 v c1 c0)).(or3_intro1 (land (eq C c1 c0) (subst0 i v t1 t1)) (land (eq T t1 
51 t1) (csubst0 i v c1 c0)) (land (subst0 i v t1 t1) (csubst0 i v c1 c0)) (conj 
52 (eq T t1 t1) (csubst0 i v c1 c0) (refl_equal T t1) H0)))) (\lambda (t0: 
53 T).(\lambda (H0: (subst0 i v t1 t0)).(\lambda (c0: C).(\lambda (H1: (csubst0 
54 i v c1 c0)).(or3_intro2 (land (eq C c1 c0) (subst0 i v t1 t0)) (land (eq T t1 
55 t0) (csubst0 i v c1 c0)) (land (subst0 i v t1 t0) (csubst0 i v c1 c0)) (conj 
56 (subst0 i v t1 t0) (csubst0 i v c1 c0) H0 H1)))))) c2 t2 H))))))).
57