(****************************************************************************) (* *) (* The formal system \lambda\delta *) (* *) (* Author: Ferruccio Guidi *) (* *) (* This file is distributed under the terms of the *) (* GNU General Public License Version 2 *) (* *) (****************************************************************************) Require subst0_subst0. Require subst1_defs. Section subst1_subst1. (**************************************************) Theorem subst1_subst1: (t1,t2,u2:?; j:?) (subst1 j u2 t1 t2) -> (u1,u:?; i:?) (subst1 i u u1 u2) -> (EX t | (subst1 j u1 t1 t) & (subst1 (S (plus i j)) u t t2)). Proof. Intros until 1; XElim H; Clear t2. (* case 1: subst1_refl on first premise *) XDEAuto 2. (* case 2: subst1_single on first premise *) Intros until 2; InsertEq H0 u2; XElim H0; Clear y; Intros. (* case 2.1: subst1_refl on second premise *) Rewrite H0; Clear H0 u1; XDEAuto 3. (* case 2.2: subst1_single on second premise *) Rewrite H1 in H0; Clear H1 t0; Subst0Subst0; XDEAuto 4. Qed. Theorem subst1_subst1_back: (t1,t2,u2:?; j:?) (subst1 j u2 t1 t2) -> (u1,u:?; i:?) (subst1 i u u2 u1) -> (EX t | (subst1 j u1 t1 t) & (subst1 (S (plus i j)) u t2 t)). Proof. Intros until 1; XElim H; Clear t2. (* case 1: subst1_refl on first premise *) XDEAuto 2. (* case 2: subst1_single on first premise *) Intros until 2; XElim H0; Clear u1; Intros; Try Subst0Subst0; XDEAuto 4. Qed. Hints Resolve subst0_trans : ld. Theorem subst1_trans: (t2,t1,v:?; i:?) (subst1 i v t1 t2) -> (t3:?) (subst1 i v t2 t3) -> (subst1 i v t1 t3). Proof. Intros until 1; XElim H; Clear t2. (* case 1: subst1_refl on first premise *) XDEAuto 2. (* case 2: subst1_single on first premise *) Intros until 2; XElim H0; Clear t3; XDEAuto 3. Qed. End subst1_subst1. Tactic Definition Subst1Subst1 := ( Match Context With | [ H1: (subst1 ?0 ?1 ?2 ?3); H2: (subst1 ?4 ?5 ?6 ?1) |- ? ] -> LApply (subst1_subst1 ?2 ?3 ?1 ?0); [ Intros H_x | XAuto ]; LApply (H_x ?6 ?5 ?4); [ Clear H_x; Intros H_x | XAuto ]; XElim H_x; Intros | [ H1: (subst1 ?0 ?1 ?2 ?3); H2: (subst0 ?4 ?5 ?6 ?1) |- ? ] -> LApply (subst1_subst1 ?2 ?3 ?1 ?0); [ Intros H_x | XAuto ]; LApply (H_x ?6 ?5 ?4); [ Clear H_x; Intros H_x | XAuto ]; XElim H_x; Intros | [ H1: (subst1 ?0 ?1 ?2 ?3); H2: (subst1 ?4 ?5 ?1 ?6) |- ? ] -> LApply (subst1_subst1_back ?2 ?3 ?1 ?0); [ Intros H_x | XAuto ]; LApply (H_x ?6 ?5 ?4); [ Clear H_x; Intros H_x | XAuto ]; XElim H_x; Intros ).