]> matita.cs.unibo.it Git - helm.git/blob - helm/coq-contribs/LAMBDA-TYPES/subst0_subst0.v
ocaml 3.09 transition
[helm.git] / helm / coq-contribs / LAMBDA-TYPES / subst0_subst0.v
1 (*#* #stop file *)
2
3 Require subst0_defs.
4 Require subst0_gen.
5 Require subst0_lift.
6
7    Section subst0_subst0. (**************************************************)
8
9       Tactic Definition IH :=
10          Match Context With
11             | [ H1: (u1,u:?; i:?) (subst0 i u u1 ?1) -> ?;
12                 H2: (subst0 ?2 ?3 ?4 ?1) |- ? ] ->
13                LApply (H1 ?4 ?3 ?2); [ Clear H1; Intros H1 | XAuto ];
14                XElim H1; Intros
15             | [ H1: (u1,u:?; i:?) (subst0 i u ?1 u1) -> ?;
16                 H2: (subst0 ?2 ?3 ?1 ?4) |- ? ] ->
17                LApply (H1 ?4 ?3 ?2); [ Clear H1; Intros H1 | XAuto ];
18                XElim H1; Intros.
19
20       Theorem subst0_subst0: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) ->
21                              (u1,u:?; i:?) (subst0 i u u1 u2) ->
22                              (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t t2)).
23       Intros until 1; XElim H; Intros.
24 (* case 1 : subst0_lref *)
25       Arith5 i0 i; XEAuto.
26 (* case 2 : subst0_fst *)
27       IH; XEAuto.
28 (* case 3 : subst0_snd *)
29       IH; SRwBackIn H2; XEAuto.
30 (* case 4 : subst0_both *)
31       Repeat IH; SRwBackIn H4; XEAuto.
32       Qed.
33
34       Theorem subst0_subst0_back: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) ->
35                                   (u1,u:?; i:?) (subst0 i u u2 u1) ->
36                                   (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t2 t)).
37       Intros until 1; XElim H; Intros.
38 (* case 1 : subst0_lref *)
39       Arith5 i0 i; XEAuto.
40 (* case 2 : subst0_fst *)
41       IH; XEAuto.
42 (* case 3 : subst0_snd *)
43       IH; SRwBackIn H2; XEAuto.
44 (* case 4 : subst0_both *)
45       Repeat IH; SRwBackIn H4; XEAuto.
46       Qed.
47
48       Theorem subst0_trans: (t2,t1,v:?; i:?) (subst0 i v t1 t2) ->
49                             (t3:?) (subst0 i v t2 t3) ->
50                             (subst0 i v t1 t3).
51       Intros until 1; XElim H; Intros;
52       Subst0Gen; Try Rewrite H1; Try Rewrite H3; XAuto.
53       Qed.
54
55    End subst0_subst0.
56
57       Hints Resolve subst0_trans : ltlc.
58
59       Tactic Definition Subst0Subst0 :=
60          Match Context With
61             | [ H1: (subst0 ?0 ?1 ?2 ?3); H2: (subst0 ?4 ?5 ?6 ?1) |- ? ] ->
62                LApply (subst0_subst0 ?2 ?3 ?1 ?0); [ Intros H_x | XAuto ];
63                LApply (H_x ?6 ?5 ?4); [ Clear H_x; Intros H_x | XAuto ];
64                XElim H_x; Intros
65             | [ H1: (subst0 ?0 ?1 ?2 ?3); H2: (subst0 ?4 ?5 ?1 ?6) |- ? ] ->
66                LApply (subst0_subst0_back ?2 ?3 ?1 ?0); [ Intros H_x | XAuto ];
67                LApply (H_x ?6 ?5 ?4); [ Clear H_x; Intros H_x | XAuto ];
68                XElim H_x; Intros.