]> matita.cs.unibo.it Git - helm.git/blob - helm/coq-contribs/LAMBDA-TYPES/subst0_confluence.v
ocaml 3.09 transition
[helm.git] / helm / coq-contribs / LAMBDA-TYPES / subst0_confluence.v
1 (*#* #stop file *)
2
3 Require lift_gen.
4 Require subst0_gen.
5 Require subst0_defs.
6
7    Section subst0_confluence. (**********************************************)
8
9       Tactic Definition IH :=
10          Match Context With
11             | [ H1: (t2,u2:?; i2:?) (subst0 i2 u2 ?1 t2) -> ~?2=i2 -> ?;
12                 H2: (subst0 ?3 ?4 ?1 ?5); H3: ~?2=?3 |- ? ] ->
13                  LApply (H1 ?5 ?4 ?3); [ Clear H1; Intros H1 | XAuto ];
14                  LApply H1; [ Clear H1; Intros H1 | XAuto ];
15                  XElim H1; Intros
16             | [ H1: (t2,u2:?; i2:?) (subst0 i2 u2 ?1 t2) -> ~(s k ?2)=i2 -> ?;
17                 H2: (subst0 (s k ?3) ?4 ?1 ?5); H3: ~?2=?3 |- ? ] ->
18                  LApply (H1 ?5 ?4 (s k ?3)); [ Clear H1; Intros H1 | XAuto ];
19                  LApply H1; [ Clear H1; Intros H1 | Unfold not in H3; Unfold not; XEAuto ];
20                  XElim H1; Intros
21             | [ H1: (t2:T) (subst0 ?1 ?2 ?3 t2) -> ?; H2: (subst0 ?1 ?2 ?3 ?4) |- ? ] ->
22                  LApply (H1 ?4); [ Clear H1; Intros H1 | XAuto ];
23                  XElim H1; Intros H1; [ Try Rewrite H1 | XElim H1; Intros | Idtac | Idtac ].
24
25       Theorem subst0_confluence_neq : (t0,t1,u1:?; i1:?) (subst0 i1 u1 t0 t1) ->
26                                       (t2,u2:?; i2:?) (subst0 i2 u2 t0 t2) ->
27                                       ~i1=i2 ->
28                                       (EX t | (subst0 i2 u2 t1 t) &
29                                               (subst0 i1 u1 t2 t)).
30
31       Intros until 1; XElim H; Intros;
32       Subst0GenBase; Try Rewrite H in H0; Try Rewrite H1; Try Rewrite H3;
33       Try EqFalse; Repeat IH; XEAuto.
34       Qed.
35
36       Theorem subst0_confluence_eq : (t0,t1,u:?; i:?) (subst0 i u t0 t1) ->
37                                      (t2:?) (subst0 i u t0 t2) -> (OR
38                                      t1 = t2 |
39                                      (EX t | (subst0 i u t1 t) & (subst0 i u t2 t)) |
40                                      (subst0 i u t1 t2) |
41                                      (subst0 i u t2 t1)).
42       Intros until 1; XElim H; Intros;
43       Subst0GenBase; Try Rewrite H1; Try Rewrite H3;
44       Repeat IH; XEAuto.
45       Qed.
46
47    End subst0_confluence.
48
49       Tactic Definition Subst0Confluence :=
50          Match Context With
51             | [ H0: (subst0 ?1 ?2 ?3 ?4);
52                 H1: (subst0 ?1 ?2 ?3 ?5) |- ? ] ->
53               LApply (subst0_confluence_eq ?3 ?4 ?2 ?1); [ Clear H0; Intros H0 | XAuto ];
54               LApply (H0 ?5); [ Clear H0; Intros H0 | XAuto ];
55               XElim H0; [ Intros | Intros H0; XElim H0; Intros | Intros | Intros ]
56             | [ H0: (subst0 ?1 ?2 ?3 ?4);
57                 H1: (subst0 ?5 ?6 ?3 ?7) |- ? ] ->
58               LApply (subst0_confluence_neq ?3 ?4 ?2 ?1); [ Clear H0; Intros H0 | XAuto ];
59               LApply (H0 ?7 ?6 ?5); [ Clear H0 H1; Intros H0 | XAuto ];
60               LApply H0; [ Clear H0; Intros H0 | Simpl; XAuto ];
61               XElim H0; Intros.
62
63    Section subst0_confluence_lift. (*****************************************)
64
65       Theorem subst0_confluence_lift: (t0,t1,u:?; i:?) (subst0 i u t0 (lift (1) i t1)) ->
66                                       (t2:?) (subst0 i u t0 (lift (1) i t2)) ->
67                                       t1 = t2.
68       Intros; Subst0Confluence;
69       Try Subst0Gen; SymEqual; LiftGen; XEAuto.
70       Qed.
71
72    End subst0_confluence_lift.
73
74       Tactic Definition Subst0ConfluenceLift :=
75          Match Context With
76             | [ H0: (subst0 ?1 ?2 ?3 (lift (1) ?1 ?4));
77                 H1: (subst0 ?1 ?2 ?3 (lift (1) ?1 ?5)) |- ? ] ->
78               LApply (subst0_confluence_lift ?3 ?4 ?2 ?1); [ Clear H0; Intros H0 | XAuto ];
79               LApply (H0 ?5); [ Clear H0; Intros | XAuto ]
80             | _ -> Subst0Confluence.