(****************************************************************************) (* *) (* The formal system \lambda\delta *) (* *) (* Author: Ferruccio Guidi *) (* *) (* This file is distributed under the terms of the *) (* GNU General Public License Version 2 *) (* *) (****************************************************************************) Require subst1_defs. Require pr0_defs. Require pr0_subst0. Section pr0_subst1_props. (***********************************************) Lemma pr0_delta1: (u1,u2:?) (pr0 u1 u2) -> (t1,t2:?) (pr0 t1 t2) -> (w:?) (subst1 (0) u2 t2 w) -> (pr0 (THead (Bind Abbr) u1 t1) (THead (Bind Abbr) u2 w)). Proof. Intros until 3; XElim H1; Clear w; XDEAuto 2. Qed. Lemma pr0_subst1_back: (u2,t1,t2:?; i:?) (subst1 i u2 t1 t2) -> (u1:?) (pr0 u1 u2) -> (EX t | (subst1 i u1 t1 t) & (pr0 t t2)). Proof. Intros until 1; XElim H; Intros; Try Pr0Subst0; XDEAuto 3. Qed. Lemma pr0_subst1_fwd: (u2,t1,t2:?; i:?) (subst1 i u2 t1 t2) -> (u1:?) (pr0 u2 u1) -> (EX t | (subst1 i u1 t1 t) & (pr0 t2 t)). Proof. Intros until 1; XElim H; Intros; Try Pr0Subst0; XDEAuto 3. Qed. Theorem pr0_subst1: (t1,t2:?) (pr0 t1 t2) -> (v1,w1:?; i:?) (subst1 i v1 t1 w1) -> (v2:?) (pr0 v1 v2) -> (EX w2 | (pr0 w1 w2) & (subst1 i v2 t2 w2)). Proof. Intros until 2; XElim H0; Intros; Try Pr0Subst0; XDEAuto 3. Qed. End pr0_subst1_props. Hints Resolve pr0_delta1 : ld. Tactic Definition Pr0Subst1 := ( Match Context With | [ H1: (pr0 ?1 ?2); H2: (subst1 ?3 ?4 ?1 ?5); H3: (pr0 ?4 ?6) |- ? ] -> LApply (pr0_subst1 ?1 ?2); [ Clear H1; Intros H1 | XAuto ]; LApply (H1 ?4 ?5 ?3); [ Clear H1 H2; Intros H1 | XAuto ]; LApply (H1 ?6); [ Clear H1; Intros H1 | XAuto ]; XElim H1; Intros | [ H1: (pr0 ?1 ?2); H2: (subst1 ?3 ?4 ?1 ?5) |- ? ] -> LApply (pr0_subst1 ?1 ?2); [ Clear H1; Intros H1 | XAuto ]; LApply (H1 ?4 ?5 ?3); [ Clear H1 H2; Intros H1 | XAuto ]; LApply (H1 ?4); [ Clear H1; Intros H1 | XAuto ]; XElim H1; Intros | [ H1: (subst1 ?0 ?1 ?2 ?3); H2: (pr0 ?4 ?1) |- ? ] -> LApply (pr0_subst1_back ?1 ?2 ?3 ?0); [ Clear H1; Intros H1 | XAuto ]; LApply (H1 ?4); [ Clear H1 H2; Intros H1 | XAuto ]; XElim H1; Intros | [ H1: (subst1 ?0 ?1 ?2 ?3); H2: (pr0 ?1 ?4) |- ? ] -> LApply (pr0_subst1_fwd ?1 ?2 ?3 ?0); [ Clear H1; Intros H1 | XAuto ]; LApply (H1 ?4); [ Clear H1 H2; Intros H1 | XAuto ]; XElim H1; Intros | [ |- ? ] -> Pr0Subst0 ).