diff -r 22a084c9316b -r c145c380e195 Nominal/Abs.thy --- a/Nominal/Abs.thy Tue Mar 02 11:04:49 2010 +0100 +++ b/Nominal/Abs.thy Tue Mar 02 12:28:07 2010 +0100 @@ -75,12 +75,11 @@ done lemma alpha_gen_compose_sym: - assumes b: "\pi. (aa, t) \gen (\x1 x2. R x1 x2 \ R x2 x1) f pi (ab, s)" + fixes pi + assumes b: "(aa, t) \gen (\x1 x2. R x1 x2 \ R x2 x1) f pi (ab, s)" and a: "\pi t s. (R t s \ R (pi \ t) (pi \ s))" - shows "\pi. (ab, s) \gen R f pi (aa, t)" + shows "(ab, s) \gen R f (- pi) (aa, t)" using b apply - - apply(erule exE) - apply(rule_tac x="- pi" in exI) apply(simp add: alpha_gen.simps) apply(erule conjE)+ apply(rule conjI) @@ -92,16 +91,14 @@ done lemma alpha_gen_compose_trans: - assumes b: "\pi\perm. (aa, t) \gen (\x1 x2. R x1 x2 \ (\x. R x2 x \ R x1 x)) f pi (ab, ta)" - and c: "\pi\perm. (ab, ta) \gen R f pi (ac, sa)" + fixes pi pia + assumes b: "(aa, t) \gen (\x1 x2. R x1 x2 \ (\x. R x2 x \ R x1 x)) f pi (ab, ta)" + and c: "(ab, ta) \gen R f pia (ac, sa)" and a: "\pi t s. (R t s \ R (pi \ t) (pi \ s))" - shows "\pi\perm. (aa, t) \gen R f pi (ac, sa)" + shows "(aa, t) \gen R f (pia + pi) (ac, sa)" using b c apply - apply(simp add: alpha_gen.simps) apply(erule conjE)+ - apply(erule exE)+ - apply(erule conjE)+ - apply(rule_tac x="pia + pi" in exI) apply(simp add: fresh_star_plus) apply(drule_tac x="- pia \ sa" in spec) apply(drule mp)