]> matita.cs.unibo.it Git - helm.git/blob - helm/coq-contribs/LAMBDA-TYPES/pr0_subst0.v
ocaml 3.09 transition
[helm.git] / helm / coq-contribs / LAMBDA-TYPES / pr0_subst0.v
1 Require subst0_gen.
2 Require subst0_lift.
3 Require subst0_subst0.
4 Require subst0_confluence.
5 Require pr0_defs.
6 Require pr0_lift.
7
8 (*#* #stop file *)
9
10    Section pr0_subst0. (*****************************************************)
11
12       Tactic Definition IH :=
13          Match Context With
14             | [ H1: (u1:?) (pr0 u1 ?1) -> ?; H2: (pr0 ?2 ?1) |- ? ] ->
15                LApply (H1 ?2); [ Clear H1; Intros H1 | XAuto ];
16                XElim H1; Intros
17             | [ H1: (u1:?) (pr0 ?1 u1) -> ?; H2: (pr0 ?1 ?2) |- ? ] ->
18                LApply (H1 ?2); [ Clear H1; Intros H1 | XAuto ];
19                XElim H1; Intros
20             | [ H1: (v1,w1:?; i:?) (subst0 i v1 ?1 w1) -> (v2:T) (pr0 v1 v2) -> ?;
21                 H2: (subst0 ?2 ?3 ?1 ?4); H3: (pr0 ?3 ?5) |- ? ] ->
22                LApply (H1 ?3 ?4 ?2); [ Clear H1; Intros H1 | XAuto ];
23                LApply (H1 ?5); [ Clear H1; Intros H1 | XAuto ];
24                XElim H1; [ Intros | Intros H1; XElim H1; Intros ].
25
26       Theorem pr0_subst0_back: (u2,t1,t2:?; i:?) (subst0 i u2 t1 t2) ->
27                                (u1:?) (pr0 u1 u2) ->
28                                (EX t | (subst0 i u1 t1 t) & (pr0 t t2)).
29       Intros until 1; XElim H; Intros;
30       Repeat IH; XEAuto.
31       Qed.
32
33       Theorem pr0_subst0_fwd: (u2,t1,t2:?; i:?) (subst0 i u2 t1 t2) ->
34                               (u1:?) (pr0 u2 u1) ->
35                               (EX t | (subst0 i u1 t1 t) & (pr0 t2 t)).
36       Intros until 1; XElim H; Intros;
37       Repeat IH; XEAuto.
38       Qed.
39
40       Hints Resolve pr0_subst0_fwd : ltlc.
41
42 (*#* #start file *)
43
44 (*#* #caption "confluence with strict substitution" *)
45 (*#* #cap #cap t1, t2 #alpha v1 in W1, v2 in W2, w1 in U1, w2 in U2 *)
46
47       Theorem pr0_subst0: (t1,t2:?) (pr0 t1 t2) ->
48                           (v1,w1:?; i:?) (subst0 i v1 t1 w1) ->
49                           (v2:?) (pr0 v1 v2) ->
50                           (pr0 w1 t2) \/
51                           (EX w2 | (pr0 w1 w2) & (subst0 i v2 t2 w2)).
52
53 (*#* #stop file *)
54
55       Intros until 1; XElim H; Clear t1 t2; Intros.
56 (* case 1: pr0_refl *)
57       XEAuto.
58 (* case 2: pr0_cong *)
59       Subst0Gen; Rewrite H3; Repeat IH; XEAuto.
60 (* case 3: pr0_beta *)
61       Repeat Subst0Gen; Rewrite H3; Try Rewrite H5; Try Rewrite H6;
62       Repeat IH; XEAuto.
63 (* case 4: pr0_upsilon *)
64       Repeat Subst0Gen; Rewrite H6; Try Rewrite H8; Try Rewrite H9;
65       Repeat IH; XDEAuto 7.
66 (* case 5: pr0_delta *)
67       Subst0Gen; Rewrite H4; Repeat IH;
68       [ XEAuto | Idtac | XEAuto | Idtac | XEAuto | Idtac | Idtac | Idtac ].
69       Subst0Subst0; Arith9'In H9 i; XEAuto.
70       Subst0Confluence; XEAuto.
71       Subst0Subst0; Arith9'In H10 i; XEAuto.
72       Subst0Confluence; XEAuto.
73       Subst0Subst0; Arith9'In H11 i; Subst0Confluence; XDEAuto 6.
74 (* case 6: pr0_zeta *)
75       Repeat Subst0Gen; Rewrite H2; Try Rewrite H4; Try Rewrite H5;
76       Try (Simpl in H5; Rewrite <- minus_n_O in H5);
77       Try (Simpl in H6; Rewrite <- minus_n_O in H6);
78       Try IH; XEAuto.
79 (* case 7: pr0_epsilon *)
80       Subst0Gen; Rewrite H1; Try IH; XEAuto.
81       Qed.
82
83    End pr0_subst0.
84
85       Tactic Definition Pr0Subst0 :=
86          Match Context With
87             | [ H1: (pr0 ?1 ?2); H2: (subst0 ?3 ?4 ?1 ?5);
88                 H3: (pr0 ?4 ?6) |- ? ] ->
89                LApply (pr0_subst0 ?1 ?2); [ Clear H1; Intros H1 | XAuto ];
90                LApply (H1 ?4 ?5 ?3); [ Clear H1 H2; Intros H1 | XAuto ];
91                LApply (H1 ?6); [ Clear H1; Intros H1 | XAuto ];
92                XElim H1; [ Intros | Intros H1; XElim H1; Intros ]
93             | [ H1: (pr0 ?1 ?2); H2: (subst0 ?3 ?4 ?1 ?5) |- ? ] ->
94                LApply (pr0_subst0 ?1 ?2); [ Clear H1; Intros H1 | XAuto ];
95                LApply (H1 ?4 ?5 ?3); [ Clear H1 H2; Intros H1 | XAuto ];
96                LApply (H1 ?4); [ Clear H1; Intros H1 | XAuto ];
97                XElim H1; [ Intros | Intros H1; XElim H1; Intros ]
98             | [ _: (subst0 ?0 ?1 ?2 ?3); _: (pr0 ?4 ?1) |- ? ] ->
99                LApply (pr0_subst0_back ?1 ?2 ?3 ?0); [ Intros H_x | XAuto ];
100                LApply (H_x ?4); [ Clear H_x; Intros H_x | XAuto ];
101                XElim H_x; Intros
102             | [ H1: (subst0 ?0 ?1 ?2 ?3); H2: (pr0 ?1 ?4) |- ? ] ->
103                LApply (pr0_subst0_fwd ?1 ?2 ?3 ?0); [ Clear H1; Intros H1 | XAuto ];
104                LApply (H1 ?4); [ Clear H1 H2; Intros H1 | XAuto ];
105                XElim H1; Intros.