]> matita.cs.unibo.it Git - helm.git/blob - helm/coq-contribs/LAMBDA-TYPES/pr0_subst1.v
ocaml 3.09 transition
[helm.git] / helm / coq-contribs / LAMBDA-TYPES / pr0_subst1.v
1 (*#* #stop file *)
2
3 Require subst1_defs.
4 Require pr0_defs.
5 Require pr0_subst0.
6
7    Section pr0_subst1_props. (***********************************************)
8
9       Theorem pr0_delta1: (u1,u2:?) (pr0 u1 u2) -> (t1,t2:?) (pr0 t1 t2) ->
10                           (w:?) (subst1 (0) u2 t2 w) ->
11                           (pr0 (TTail (Bind Abbr) u1 t1) (TTail (Bind Abbr) u2 w)).
12       Intros until 3; XElim H1; Clear w; XEAuto.
13       Qed.
14
15       Theorem pr0_subst1_back: (u2,t1,t2:?; i:?) (subst1 i u2 t1 t2) ->
16                                (u1:?) (pr0 u1 u2) ->
17                                (EX t | (subst1 i u1 t1 t) & (pr0 t t2)).
18       Intros until 1; XElim H; Intros;
19       Try Pr0Subst0; XEAuto.
20       Qed.
21
22       Theorem pr0_subst1_fwd: (u2,t1,t2:?; i:?) (subst1 i u2 t1 t2) ->
23                               (u1:?) (pr0 u2 u1) ->
24                               (EX t | (subst1 i u1 t1 t) & (pr0 t2 t)).
25       Intros until 1; XElim H; Intros;
26       Try Pr0Subst0; XEAuto.
27       Qed.
28
29       Theorem pr0_subst1: (t1,t2:?) (pr0 t1 t2) ->
30                           (v1,w1:?; i:?) (subst1 i v1 t1 w1) ->
31                           (v2:?) (pr0 v1 v2) ->
32                           (EX w2 | (pr0 w1 w2) & (subst1 i v2 t2 w2)).
33       Intros until 2; XElim H0; Intros;
34       Try Pr0Subst0; XEAuto.
35       Qed.
36
37    End pr0_subst1_props.
38
39       Hints Resolve pr0_delta1 : ltlc.
40
41       Tactic Definition Pr0Subst1 :=
42          Match Context With
43             | [ H1: (pr0 ?1 ?2); H2: (subst1 ?3 ?4 ?1 ?5);
44                 H3: (pr0 ?4 ?6) |- ? ] ->
45                LApply (pr0_subst1 ?1 ?2); [ Clear H1; Intros H1 | XAuto ];
46                LApply (H1 ?4 ?5 ?3); [ Clear H1 H2; Intros H1 | XAuto ];
47                LApply (H1 ?6); [ Clear H1; Intros H1 | XAuto ];
48                XElim H1; Intros
49             | [ H1: (pr0 ?1 ?2); H2: (subst1 ?3 ?4 ?1 ?5) |- ? ] ->
50                LApply (pr0_subst1 ?1 ?2); [ Clear H1; Intros H1 | XAuto ];
51                LApply (H1 ?4 ?5 ?3); [ Clear H1 H2; Intros H1 | XAuto ];
52                LApply (H1 ?4); [ Clear H1; Intros H1 | XAuto ];
53                XElim H1; Intros
54             | [ H1: (subst1 ?0 ?1 ?2 ?3); H2: (pr0 ?4 ?1) |- ? ] ->
55                LApply (pr0_subst1_back ?1 ?2 ?3 ?0); [ Clear H1; Intros H1 | XAuto ];
56                LApply (H1 ?4); [ Clear H1 H2; Intros H1 | XAuto ];
57                XElim H1; Intros
58             | [ H1: (subst1 ?0 ?1 ?2 ?3); H2: (pr0 ?1 ?4) |- ? ] ->
59                LApply (pr0_subst1_fwd ?1 ?2 ?3 ?0); [ Clear H1; Intros H1 | XAuto ];
60                LApply (H1 ?4); [ Clear H1 H2; Intros H1 | XAuto ];
61                XElim H1; Intros
62             | _ -> Pr0Subst0.
63