(****************************************************************************) (* *) (* 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_gen. Require subst0_lift. Require subst0_subst0. Require subst0_confluence. Require pr0_defs. Require pr0_lift. Section pr0_subst0. (*****************************************************) Tactic Definition IH := ( Match Context With | [ H1: (u1:?) (pr0 u1 ?1) -> ?; H2: (pr0 ?2 ?1) |- ? ] -> LApply (H1 ?2); [ Clear H1; Intros H1 | XAuto ]; XElim H1; Intros | [ H1: (u1:?) (pr0 ?1 u1) -> ?; H2: (pr0 ?1 ?2) |- ? ] -> LApply (H1 ?2); [ Clear H1; Intros H1 | XAuto ]; XElim H1; Intros | [ H1: (v1,w1:?; i:?) (subst0 i v1 ?1 w1) -> (v2:T) (pr0 v1 v2) -> ?; H2: (subst0 ?2 ?3 ?1 ?4); H3: (pr0 ?3 ?5) |- ? ] -> LApply (H1 ?3 ?4 ?2); [ Clear H1; Intros H1 | XAuto ]; LApply (H1 ?5); [ Clear H1; Intros H1 | XAuto ]; XElim H1; [ Intros | Intros H1; XElim H1; Intros ] ). Lemma pr0_subst0_back: (u2,t1,t2:?; i:?) (subst0 i u2 t1 t2) -> (u1:?) (pr0 u1 u2) -> (EX t | (subst0 i u1 t1 t) & (pr0 t t2)). Proof. Intros until 1; XElim H; Clear i u2 t1 t2; Intros; Repeat IH; XDEAuto 4. Qed. Lemma pr0_subst0_fwd: (u2,t1,t2:?; i:?) (subst0 i u2 t1 t2) -> (u1:?) (pr0 u2 u1) -> (EX t | (subst0 i u1 t1 t) & (pr0 t2 t)). Proof. Intros until 1; XElim H; Clear i u2 t1 t2; Intros; Repeat IH; XDEAuto 4. Qed. Hints Resolve sym_eq sym_not_eq ex2_sym pr0_subst0_fwd subst0_trans : ld. Theorem pr0_subst0: (t1,t2:?) (pr0 t1 t2) -> (v1,w1:?; i:?) (subst0 i v1 t1 w1) -> (v2:?) (pr0 v1 v2) -> (pr0 w1 t2) \/ (EX w2 | (pr0 w1 w2) & (subst0 i v2 t2 w2)). Proof. Intros until 1; XElim H; Clear t1 t2; Intros. (* case 1: pr0_refl *) XDEAuto 4. (* case 2: pr0_cong *) Subst0Gen; Subst; Repeat IH; XDEAuto 5. (* case 3: pr0_beta *) Repeat Subst0Gen; Subst; Repeat IH; XDEAuto 5. (* case 4: pr0_upsilon *) Repeat Subst0Gen; Subst; Repeat IH; XDEAuto 7. (* case 5: pr0_delta *) Subst0Gen; Subst; Repeat IH. (* case 5.1: *) XDEAuto 3. (* case 5.2: *) Subst0Subst0; Arith9'In H8 i; XDEAuto 5. (* case 5.3: *) XDEAuto 3. (* case 5.4: *) Subst0Confluence; XDEAuto 5. (* case 5.5: *) XDEAuto 3. (* case 5.6: *) Subst0Subst0; Arith9'In H9 i; XDEAuto 5. (* case 5.7: *) Subst0Confluence; XDEAuto 5. (* case 5.8: *) Subst0Subst0; Arith9'In H10 i; Subst0Confluence; XDEAuto 6. (* case 6: pr0_zeta *) Repeat Subst0Gen; Subst; Try (Simpl in H5; Rewrite <- minus_n_O in H5); Try (Simpl in H6; Rewrite <- minus_n_O in H6); Try Rewrite <- minus_n_O in H7; Try IH; XDEAuto 4. (* case 7: pr0_tau *) Subst0Gen; Subst; Try IH; XDEAuto 4. Qed. End pr0_subst0. Tactic Definition Pr0Subst0 := ( Match Context With | [ H1: (pr0 ?1 ?2); H2: (subst0 ?3 ?4 ?1 ?5); H3: (pr0 ?4 ?6) |- ? ] -> LApply (pr0_subst0 ?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 | Intros H1; XElim H1; Intros ] | [ H1: (pr0 ?1 ?2); H2: (subst0 ?3 ?4 ?1 ?5) |- ? ] -> LApply (pr0_subst0 ?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 | Intros H1; XElim H1; Intros ] | [ _: (subst0 ?0 ?1 ?2 ?3); _: (pr0 ?4 ?1) |- ? ] -> LApply (pr0_subst0_back ?1 ?2 ?3 ?0); [ Intros H_x | XAuto ]; LApply (H_x ?4); [ Clear H_x; Intros H_x | XAuto ]; XElim H_x; Intros | [ H1: (subst0 ?0 ?1 ?2 ?3); H2: (pr0 ?1 ?4) |- ? ] -> LApply (pr0_subst0_fwd ?1 ?2 ?3 ?0); [ Clear H1; Intros H1 | XAuto ]; LApply (H1 ?4); [ Clear H1 H2; Intros H1 | XAuto ]; XElim H1; Intros ).