diff -r 1c9931e5039a -r 2ff84b1f551f Nominal/Test.thy --- a/Nominal/Test.thy Thu Mar 18 08:03:42 2010 +0100 +++ b/Nominal/Test.thy Thu Mar 18 08:32:55 2010 +0100 @@ -43,7 +43,7 @@ apply(simp only: infinite_Un) apply(simp only: Collect_disj_eq) apply(simp only: supp_def[symmetric]) -apply(rule_tac t="supp (Lm name lm_raw)" and s="supp (Abs {atom name} lm_raw)" in subst) +apply(rule_tac t="supp (Lm name lm)" and s="supp (Abs {atom name} lm)" in subst) apply(simp (no_asm) only: supp_def) apply(simp only: lm_perm) apply(simp only: permute_ABS) @@ -68,8 +68,8 @@ lemma assumes a0: "finite (supp c)" and a1: "\name c. P c (Vr name)" - and a2: "\lm_raw1 lm_raw2 c. \\d. P d lm_raw1; \d. P d lm_raw2\ \ P c (Ap lm_raw1 lm_raw2)" - and a3: "\name lm_raw c. \\d. P d lm_raw\ \ P c (Lm name lm_raw)" + and a2: "\lm1 lm2 c. \\d. P d lm1; \d. P d lm2\ \ P c (Ap lm1 lm2)" + and a3: "\name lm c. \\d. P d lm\ \ P c (Lm name lm)" shows "P c lm" proof - have "\p c. P c (p \ lm)" @@ -80,11 +80,11 @@ apply(rule a2) apply(assumption) apply(assumption) - apply(subgoal_tac "\new::name. (atom new) \ (c, Lm (p \ name) (p \ lm_raw)) + apply(subgoal_tac "\new::name. (atom new) \ (c, Lm (p \ name) (p \ lm)) \ sort_of (atom new) = sort_of (atom name)") apply(erule exE) - apply(rule_tac t="(p \ Lm name lm_raw)" and - s="((((atom (p \ name)) \ (atom new)) + p) \ Lm name lm_raw)" in subst) + apply(rule_tac t="(p \ Lm name lm)" and + s="((((atom (p \ name)) \ (atom new)) + p) \ Lm name lm)" in subst) apply(simp) apply(subst lm_perm) apply(subst (2) lm_perm) @@ -150,7 +150,7 @@ apply(simp only: infinite_Un) apply(simp only: Collect_disj_eq) (* LAM case *) -apply(rule_tac t="supp (LAM name lam_raw)" and s="supp (Abs {atom name} lam_raw)" in subst) +apply(rule_tac t="supp (LAM name lam)" and s="supp (Abs {atom name} lam)" in subst) apply(simp (no_asm) only: supp_def) apply(simp only: lam_bp_perm) apply(simp only: permute_ABS) @@ -162,7 +162,7 @@ apply(simp only: eqvts) apply(simp only: supp_Abs) (* LET case *) -apply(rule_tac t="supp (LET bp_raw lam_raw)" and s="supp (Abs (bi bp_raw) lam_raw) \ fv_bi bp_raw" in subst) +apply(rule_tac t="supp (LET bp lam)" and s="supp (Abs (bi bp) lam) \ fv_bi bp" in subst) apply(simp (no_asm) only: supp_def) apply(simp only: lam_bp_perm) apply(simp only: permute_ABS) @@ -239,7 +239,7 @@ apply(simp only: infinite_Un) apply(simp only: Collect_disj_eq) (* LAM case *) -apply(rule_tac t="supp (LAM' name lam'_raw)" and s="supp (Abs {atom name} lam'_raw)" in subst) +apply(rule_tac t="supp (LAM' name lam')" and s="supp (Abs {atom name} lam')" in subst) apply(simp (no_asm) only: supp_def) apply(simp only: lam'_bp'_perm) apply(simp only: permute_ABS) @@ -251,8 +251,8 @@ apply(simp only: eqvts) apply(simp only: supp_Abs) (* LET case *) -apply(rule_tac t="supp (LET' bp'_raw lam'_raw)" and - s="supp (Abs (bi' bp'_raw) (bp'_raw, lam'_raw))" in subst) +apply(rule_tac t="supp (LET' bp' lam')" and + s="supp (Abs (bi' bp') (bp', lam'))" in subst) apply(simp (no_asm) only: supp_def) apply(simp only: lam'_bp'_perm) apply(simp only: permute_ABS) @@ -325,7 +325,7 @@ apply(simp only: infinite_Un) apply(simp only: Collect_disj_eq) (* LAM case *) -apply(rule_tac t="supp (Lam name trm'_raw)" and s="supp (Abs {atom name} trm'_raw)" in subst) +apply(rule_tac t="supp (Lam name trm')" and s="supp (Abs {atom name} trm')" in subst) apply(simp (no_asm) only: supp_def) apply(simp only: trm'_pat'_perm) apply(simp only: permute_ABS) @@ -337,8 +337,8 @@ apply(simp only: eqvts) apply(simp only: supp_Abs) (* LET case *) -apply(rule_tac t="supp (Let pat'_raw trm'_raw1 trm'_raw2)" - and s="supp (Abs (f pat'_raw) trm'_raw2) \ supp trm'_raw1 \ fv_f pat'_raw" in subst) +apply(rule_tac t="supp (Let pat' trm'1 trm'2)" + and s="supp (Abs (f pat') trm'2) \ supp trm'1 \ fv_f pat'" in subst) apply(simp (no_asm) only: supp_def) apply(simp only: trm'_pat'_perm) apply(simp only: permute_ABS) @@ -460,7 +460,7 @@ apply(simp only: infinite_Un) apply(simp only: Collect_disj_eq) (* All case *) -apply(rule_tac t="supp (All fun t_raw)" and s="supp (Abs (atom ` fun) t_raw)" in subst) +apply(rule_tac t="supp (All fun t)" and s="supp (Abs (atom ` fun) t)" in subst) apply(simp (no_asm) only: supp_def) apply(simp only: t_tyS_perm) apply(simp only: permute_ABS)