(****************************************************************************) (* *) (* 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_confluence. Require subst1_defs. Require subst1_gen. Section subst1_confluence. (**********************************************) Hints Resolve sym_not_eq : ld. Theorem subst1_confluence_neq: (t0,t1,u1:?; i1:?) (subst1 i1 u1 t0 t1) -> (t2,u2:?; i2:?) (subst1 i2 u2 t0 t2) -> ~i1=i2 -> (EX t | (subst1 i2 u2 t1 t) & (subst1 i1 u1 t2 t) ). Proof. Intros until 1; XElim H; Clear t1; Intros. (* case 1; subst1_refl *) XDEAuto 2. (* case 2; subst1_single *) XElim H0; Intros; Try Subst0Confluence; XDEAuto 4. Qed. Theorem subst1_confluence_eq: (t0,t1,u:?; i:?) (subst1 i u t0 t1) -> (t2:?) (subst1 i u t0 t2) -> (EX t | (subst1 i u t1 t) & (subst1 i u t2 t) ). Proof. Intros until 1; XElim H; Intros. (* case 1; subst1_refl *) XDEAuto 2. (* case 2; subst1_single *) XElim H0; Intros; Try Subst0Confluence; Try Rewrite H0; XDEAuto 4. Qed. Hints Resolve sym_eq : ld. Theorem subst1_confluence_lift: (t0,t1,u:?; i:?) (subst1 i u t0 (lift (1) i t1)) -> (t2:?) (subst1 i u t0 (lift (1) i t2)) -> t1 = t2. Proof. Intros until 1; InsertEq H '(lift (1) i t1); XElim H; Clear y; Intros. (* case 1: subst1_refl *) Rewrite H in H0; Clear H t0. Subst1Gen; SymEqual; LiftGen; XDEAuto 2. (* case 2: subst1_single *) Rewrite H0 in H; Clear H0 t2. InsertEq H1 '(lift (1) i t3); XElim H0; Clear y; Intros. (* case 2.1: subst1_refl *) Rewrite H0 in H; Clear H0 t0; Subst0Gen. (* case 2.2: subst1_single *) Rewrite H1 in H0; Clear H1 t2; Subst0ConfluenceLift; XAuto. Qed. End subst1_confluence. Tactic Definition Subst1Confluence := ( Match Context With | [ H0: (subst1 ?1 ?2 ?3 (lift (1) ?1 ?4)); H1: (subst1 ?1 ?2 ?3 (lift (1) ?1 ?5)) |- ? ] -> LApply (subst1_confluence_lift ?3 ?4 ?2 ?1); [ Clear H0; Intros H0 | XAuto ]; LApply (H0 ?5); [ Clear H0; Intros | XAuto ] | [ H0: (subst1 ?1 ?2 ?3 ?4); H1: (subst1 ?1 ?2 ?3 ?5) |- ? ] -> LApply (subst1_confluence_eq ?3 ?4 ?2 ?1); [ Clear H0; Intros H0 | XAuto ]; LApply (H0 ?5); [ Clear H0; Intros H0 | XAuto ]; XElim H0; Intros | [ H0: (subst0 ?1 ?2 ?3 ?4); H1: (subst1 ?1 ?2 ?3 ?5) |- ? ] -> LApply (subst1_confluence_eq ?3 ?4 ?2 ?1); [ Clear H0; Intros H0 | XAuto ]; LApply (H0 ?5); [ Clear H0; Intros H0 | XAuto ]; XElim H0; Intros | [ H0: (subst1 ?1 ?2 ?3 ?4); H1: (subst1 ?5 ?6 ?3 ?7) |- ? ] -> LApply (subst1_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 | XAuto ]; XElim H0; Intros | [ H0: (subst0 ?1 ?2 ?3 ?4); H1: (subst1 ?5 ?6 ?3 ?7) |- ? ] -> LApply (subst1_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 | XAuto ]; XElim H0; Intros ).