]> matita.cs.unibo.it Git - helm.git/blob - helm/coq-contribs/LAMBDA-TYPES/contexts_defs.v
some reorganization and some corrections
[helm.git] / helm / coq-contribs / LAMBDA-TYPES / contexts_defs.v
1 (*#* #stop file *)
2
3 Require Export terms_defs.
4
5       Inductive Set C := CSort : nat -> C
6                       |  CTail : C -> K -> T -> C.
7
8       Hint f3CKT : ltlc := Resolve (f_equal3 C K T).
9
10       Tactic Definition CGenBase :=
11          Match Context With
12             | [ H: (CTail ? ? ?) = (CTail ? ? ?) |- ? ] -> Inversion H; Clear H
13             | _                                         -> TGenBase.
14
15       Definition r: K -> nat -> nat := [k;i] Cases k of
16          | (Bind _) => i
17          | (Flat _) => (S i)
18       end.
19
20       Fixpoint app [c:C] : nat -> T -> T := [j;t]Cases j c of
21          | (0)   _             => t
22          | _     (CSort _)     => t
23          | (S i) (CTail c k u) => (app c (r k i) (TTail k u t))
24       end.
25
26    Section r_props. (********************************************************)
27
28       Theorem r_S: (k:?; i:?) (r k (S i)) = (S (r k i)).
29       XElim k; XAuto.
30       Qed.
31
32       Theorem r_plus_sym: (k:?; i,j:?) (r k (plus i j)) = (plus i (r k j)).
33       XElim k; Intros; Simpl; XAuto.
34       Qed.
35
36       Theorem r_minus: (i,n:?) (lt n i) ->
37                        (k:?) (minus (r k i) (S n)) = (r k (minus i (S n))).
38       XElim k; Intros; Simpl; XEAuto.
39       Qed.
40
41       Theorem r_dis: (k:?; P:Prop)
42                      (((i:?) (r k i) = i) -> P) ->
43                      (((i:?) (r k i) = (S i)) -> P) -> P.
44       XElim k; XAuto.
45       Qed.
46
47    End r_props.
48
49       Tactic Definition RRw :=
50          Repeat (Rewrite r_S Orelse Rewrite r_plus_sym).
51
52    Section r_arith. (********************************************************)
53
54       Theorem r_arith0: (k:?; i:?) (minus (r k (S i)) (1)) = (r k i).
55       Intros; RRw; Rewrite minus_Sx_SO; XAuto.
56       Qed.
57
58       Theorem r_arith1: (k:?; i,j:?) (minus (r k (S i)) (S j)) = (minus (r k i) j).
59       Intros; RRw; XAuto.
60       Qed.
61
62    End r_arith.
63
64    Section app_props. (******************************************************)
65
66       Theorem app_csort : (t:?; i,n:?) (app (CSort n) i t) = t.
67       XElim i; Intros; Simpl; XAuto.
68       Qed.
69
70       Theorem app_O : (c:?; t:?) (app c (0) t) = t.
71       XElim c; XAuto.
72       Qed.
73
74    End app_props.
75
76       Hints Resolve app_csort app_O : ltlc.