]> matita.cs.unibo.it Git - helm.git/blob - helm/software/matita/contribs/LAMBDA-TYPES/Basic-1/subst/fwd.ma
Preparing for 0.5.9 release.
[helm.git] / helm / software / matita / contribs / LAMBDA-TYPES / Basic-1 / subst / fwd.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 (* This file was automatically generated: do not edit *********************)
16
17 include "Basic-1/subst/defs.ma".
18
19 theorem subst_sort:
20  \forall (v: T).(\forall (d: nat).(\forall (k: nat).(eq T (subst d v (TSort 
21 k)) (TSort k))))
22 \def
23  \lambda (_: T).(\lambda (_: nat).(\lambda (k: nat).(refl_equal T (TSort 
24 k)))).
25 (* COMMENTS
26 Initial nodes: 13
27 END *)
28
29 theorem subst_lref_lt:
30  \forall (v: T).(\forall (d: nat).(\forall (i: nat).((lt i d) \to (eq T 
31 (subst d v (TLRef i)) (TLRef i)))))
32 \def
33  \lambda (v: T).(\lambda (d: nat).(\lambda (i: nat).(\lambda (H: (lt i 
34 d)).(eq_ind_r bool true (\lambda (b: bool).(eq T (match b with [true 
35 \Rightarrow (TLRef i) | false \Rightarrow (match (blt d i) with [true 
36 \Rightarrow (TLRef (pred i)) | false \Rightarrow (lift d O v)])]) (TLRef i))) 
37 (refl_equal T (TLRef i)) (blt i d) (lt_blt d i H))))).
38 (* COMMENTS
39 Initial nodes: 73
40 END *)
41
42 theorem subst_lref_eq:
43  \forall (v: T).(\forall (i: nat).(eq T (subst i v (TLRef i)) (lift i O v)))
44 \def
45  \lambda (v: T).(\lambda (i: nat).(eq_ind_r bool false (\lambda (b: bool).(eq 
46 T (match b with [true \Rightarrow (TLRef i) | false \Rightarrow (match b with 
47 [true \Rightarrow (TLRef (pred i)) | false \Rightarrow (lift i O v)])]) (lift 
48 i O v))) (refl_equal T (lift i O v)) (blt i i) (le_bge i i (le_n i)))).
49 (* COMMENTS
50 Initial nodes: 71
51 END *)
52
53 theorem subst_lref_gt:
54  \forall (v: T).(\forall (d: nat).(\forall (i: nat).((lt d i) \to (eq T 
55 (subst d v (TLRef i)) (TLRef (pred i))))))
56 \def
57  \lambda (v: T).(\lambda (d: nat).(\lambda (i: nat).(\lambda (H: (lt d 
58 i)).(eq_ind_r bool false (\lambda (b: bool).(eq T (match b with [true 
59 \Rightarrow (TLRef i) | false \Rightarrow (match (blt d i) with [true 
60 \Rightarrow (TLRef (pred i)) | false \Rightarrow (lift d O v)])]) (TLRef 
61 (pred i)))) (eq_ind_r bool true (\lambda (b: bool).(eq T (match b with [true 
62 \Rightarrow (TLRef (pred i)) | false \Rightarrow (lift d O v)]) (TLRef (pred 
63 i)))) (refl_equal T (TLRef (pred i))) (blt d i) (lt_blt i d H)) (blt i d) 
64 (le_bge d i (lt_le_weak d i H)))))).
65 (* COMMENTS
66 Initial nodes: 130
67 END *)
68
69 theorem subst_head:
70  \forall (k: K).(\forall (w: T).(\forall (u: T).(\forall (t: T).(\forall (d: 
71 nat).(eq T (subst d w (THead k u t)) (THead k (subst d w u) (subst (s k d) w 
72 t)))))))
73 \def
74  \lambda (k: K).(\lambda (w: T).(\lambda (u: T).(\lambda (t: T).(\lambda (d: 
75 nat).(refl_equal T (THead k (subst d w u) (subst (s k d) w t))))))).
76 (* COMMENTS
77 Initial nodes: 37
78 END *)
79