(****************************************************************************) (* *) (* 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_defs. Require subst0_gen. Require subst0_lift. Section subst0_subst0. (**************************************************) Tactic Definition IH := ( Match Context With | [ H1: (u1,u:?; i:?) (subst0 i u u1 ?1) -> ?; H2: (subst0 ?2 ?3 ?4 ?1) |- ? ] -> LApply (H1 ?4 ?3 ?2); [ Clear H1; Intros H1 | XAuto ]; XElim H1; Intros | [ H1: (u1,u:?; i:?) (subst0 i u ?1 u1) -> ?; H2: (subst0 ?2 ?3 ?1 ?4) |- ? ] -> LApply (H1 ?4 ?3 ?2); [ Clear H1; Intros H1 | XAuto ]; XElim H1; Intros ). Theorem subst0_subst0: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) -> (u1,u:?; i:?) (subst0 i u u1 u2) -> (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t t2)). Proof. Intros until 1; XElim H; Intros. (* case 1 : subst0_lref *) Arith5 i0 i; XDEAuto 3. (* case 2 : subst0_fst *) IH; XDEAuto 4. (* case 3 : subst0_snd *) IH; SRwBackIn H2; XDEAuto 4. (* case 4 : subst0_both *) Repeat IH; SRwBackIn H4; XDEAuto 4. Qed. Theorem subst0_subst0_back: (t1,t2,u2:?; j:?) (subst0 j u2 t1 t2) -> (u1,u:?; i:?) (subst0 i u u2 u1) -> (EX t | (subst0 j u1 t1 t) & (subst0 (S (plus i j)) u t2 t)). Proof. Intros until 1; XElim H; Intros. (* case 1 : subst0_lref *) Arith5 i0 i; XDEAuto 3. (* case 2 : subst0_fst *) IH; XDEAuto 4. (* case 3 : subst0_snd *) IH; SRwBackIn H2; XDEAuto 4. (* case 4 : subst0_both *) Repeat IH; SRwBackIn H4; XDEAuto 4. Qed. Theorem subst0_trans: (t2,t1,v:?; i:?) (subst0 i v t1 t2) -> (t3:?) (subst0 i v t2 t3) -> (subst0 i v t1 t3). Proof. Intros until 1; XElim H; Intros; Subst0Gen; Try Rewrite H1; Try Rewrite H3; XAuto. Qed. End subst0_subst0. Tactic Definition Subst0Subst0 := ( Match Context With | [ H1: (subst0 ?0 ?1 ?2 ?3); H2: (subst0 ?4 ?5 ?6 ?1) |- ? ] -> LApply (subst0_subst0 ?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: (subst0 ?0 ?1 ?2 ?3); H2: (subst0 ?4 ?5 ?1 ?6) |- ? ] -> LApply (subst0_subst0_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 ).