diff -r cbcd4997dac5 -r 22a084c9316b Nominal/Abs.thy --- a/Nominal/Abs.thy Tue Mar 02 08:58:28 2010 +0100 +++ b/Nominal/Abs.thy Tue Mar 02 11:04:49 2010 +0100 @@ -114,14 +114,13 @@ done lemma alpha_gen_compose_eqvt: - assumes b: "\pia. (g d, t) \gen (\x1 x2. R x1 x2 \ R (pi \ x1) (pi \ x2)) f pia (g e, s)" + fixes pia + assumes b: "(g d, t) \gen (\x1 x2. R x1 x2 \ R (pi \ x1) (pi \ x2)) f pia (g e, s)" and c: "\y. pi \ (g y) = g (pi \ y)" and a: "\x. pi \ (f x) = f (pi \ x)" - shows "\pia. (g (pi \ d), pi \ t) \gen R f pia (g (pi \ e), pi \ s)" + shows "(g (pi \ d), pi \ t) \gen R f (pi \ pia) (g (pi \ e), pi \ s)" using b apply - - apply(erule exE) - apply(rule_tac x="pi \ pia" in exI) apply(simp add: alpha_gen.simps) apply(erule conjE)+ apply(rule conjI) @@ -419,6 +418,7 @@ apply(simp add: supp_swap) done +(* thm supp_perm lemma perm_induct_test: @@ -502,6 +502,6 @@ apply(simp) oops - +*) end