(****************************************************************************) (* *) (* The formal system \lambda\delta *) (* *) (* Author: Ferruccio Guidi *) (* *) (* This file is distributed under the terms of the *) (* GNU General Public License Version 2 *) (* *) (****************************************************************************) Require lift_gen. Require subst0_gen. Require subst0_defs. Section subst0_confluence. (**********************************************) Tactic Definition IH := ( Match Context With | [ H1: (t2,u2:?; i2:?) (subst0 i2 u2 ?1 t2) -> ~?2=i2 -> ?; H2: (subst0 ?3 ?4 ?1 ?5); H3: ~?2=?3 |- ? ] -> LApply (H1 ?5 ?4 ?3); [ Clear H1; Intros H1 | XAuto ]; LApply H1; [ Clear H1; Intros H1 | XAuto ]; XElim H1; Intros | [ H1: (t2,u2:?; i2:?) (subst0 i2 u2 ?1 t2) -> ~(s k ?2)=i2 -> ?; H2: (subst0 (s k ?3) ?4 ?1 ?5); H3: ~?2=?3 |- ? ] -> LApply (H1 ?5 ?4 (s k ?3)); [ Clear H1; Intros H1 | XAuto ]; LApply H1; [ Clear H1; Intros H1 | Unfold not in H3; Unfold not; XDEAuto 2 ]; XElim H1; Intros; Unfold not | [ H1: (t2:T) (subst0 ?1 ?2 ?3 t2) -> ?; H2: (subst0 ?1 ?2 ?3 ?4) |- ? ] -> LApply (H1 ?4); [ Clear H1; Intros H1 | XAuto ]; XElim H1; Intros H1; [ Try Rewrite H1 | XElim H1; Intros | Idtac | Idtac ] ). Hints Resolve s_inj : ld. Theorem subst0_confluence_neq: (t0,t1,u1:?; i1:?) (subst0 i1 u1 t0 t1) -> (t2,u2:?; i2:?) (subst0 i2 u2 t0 t2) -> ~i1=i2 -> (EX t | (subst0 i2 u2 t1 t) & (subst0 i1 u1 t2 t)). Proof. Intros until 1; XElim H; Intros; Subst0GenBase; Try Rewrite H in H0; Try Rewrite H1; Try Rewrite H3; Try EqFalse; Repeat IH; XDEAuto 4. Qed. Hints Resolve sym_eq : ld. Theorem subst0_confluence_eq: (t0,t1,u:?; i:?) (subst0 i u t0 t1) -> (t2:?) (subst0 i u t0 t2) -> (OR t1 = t2 | (EX t | (subst0 i u t1 t) & (subst0 i u t2 t)) | (subst0 i u t1 t2) | (subst0 i u t2 t1)). Proof. Intros until 1; XElim H; Intros; Subst0GenBase; Try Rewrite H1; Try Rewrite H3; Repeat IH; XDEAuto 5. Qed. End subst0_confluence. Tactic Definition Subst0Confluence := ( Match Context With | [ H0: (subst0 ?1 ?2 ?3 ?4); H1: (subst0 ?1 ?2 ?3 ?5) |- ? ] -> LApply (subst0_confluence_eq ?3 ?4 ?2 ?1); [ Clear H0; Intros H0 | XAuto ]; LApply (H0 ?5); [ Clear H0; Intros H0 | XAuto ]; XElim H0; [ Intros | Intros H0; XElim H0; Intros | Intros | Intros ] | [ H0: (subst0 ?1 ?2 ?3 ?4); H1: (subst0 ?5 ?6 ?3 ?7) |- ? ] -> LApply (subst0_confluence_neq ?3 ?4 ?2 ?1); [ Clear H0; Intros H0 | XAuto ]; LApply (H0 ?7 ?6 ?5); [ Clear H0 H1; Intros H0 | XAuto ]; LApply H0; [ Clear H0; Intros H0 | Simpl; XAuto ]; XElim H0; Intros ). Section subst0_confluence_lift. (*****************************************) Theorem subst0_confluence_lift: (t0,t1,u:?; i:?) (subst0 i u t0 (lift (1) i t1)) -> (t2:?) (subst0 i u t0 (lift (1) i t2)) -> t1 = t2. Proof. Intros; Subst0Confluence; Try Subst0Gen; SymEqual; LiftGen; XDEAuto 2. Qed. End subst0_confluence_lift. Tactic Definition Subst0ConfluenceLift := ( Match Context With | [ H0: (subst0 ?1 ?2 ?3 (lift (1) ?1 ?4)); H1: (subst0 ?1 ?2 ?3 (lift (1) ?1 ?5)) |- ? ] -> LApply (subst0_confluence_lift ?3 ?4 ?2 ?1); [ Clear H0; Intros H0 | XAuto ]; LApply (H0 ?5); [ Clear H0; Intros | XAuto ] | [ |- ? ] -> Subst0Confluence ).