diff -r ca6ca6fc28af -r 25d11b449e92 Nominal/Ex/TypeSchemes1.thy --- a/Nominal/Ex/TypeSchemes1.thy Tue Aug 07 18:53:50 2012 +0100 +++ b/Nominal/Ex/TypeSchemes1.thy Tue Aug 07 18:54:52 2012 +0100 @@ -80,8 +80,7 @@ where "\ = lookup \ X" | "\ T2> = (\) \ (\)" - unfolding eqvt_def subst_graph_def - apply (rule, perm_simp, rule) + apply(simp add: eqvt_def subst_graph_aux_def) apply(rule TrueI) apply(case_tac x) apply(rule_tac y="b" in ty.exhaust) @@ -119,8 +118,7 @@ substs :: "(name \ ty) list \ tys \ tys" ("_<_>" [100,60] 120) where "fset (map_fset atom Xs) \* \ \ \ = All [Xs].(\)" - unfolding eqvt_def substs_graph_def - apply (rule, perm_simp, rule) + apply(simp add: eqvt_def substs_graph_aux_def) apply auto[2] apply (rule_tac y="b" and c="a" in tys.strong_exhaust) apply auto[1] @@ -279,10 +277,7 @@ where "atom ` (fset Xs) \* T \ T \\ All [Xs].T' \ (\\. T = \ \ atom ` Dom \ = atom ` fset Xs \ supp T')" -unfolding generalises_graph_def -unfolding eqvt_def -apply(perm_simp) -apply(simp) +apply(simp add: generalises_graph_aux_def eqvt_def) apply(rule TrueI) apply(case_tac x) apply(simp) @@ -409,8 +404,7 @@ where "ftv (Var X) = {|X|}" | "ftv (T1 \ T2) = (ftv T1) |\| (ftv T2)" - unfolding eqvt_def ftv_graph_def - apply (rule, perm_simp, rule) + apply(simp add: eqvt_def ftv_graph_aux_def) apply(rule TrueI) apply(rule_tac y="x" in ty.exhaust) apply(blast)