diff -r 7c12f5476d1b -r b3deb964ad26 Quot/Nominal/Abs.thy --- a/Quot/Nominal/Abs.thy Tue Feb 02 11:23:17 2010 +0100 +++ b/Quot/Nominal/Abs.thy Tue Feb 02 11:56:37 2010 +0100 @@ -129,6 +129,26 @@ apply(clarsimp) done +lemma alpha_gen_atom_eqvt: + assumes a: "\x. pi \ (f x) = f (pi \ x)" + and b: "\pia. ({atom a}, t) \gen \x1 x2. R x1 x2 \ R (pi \ x1) (pi \ x2) f pia ({atom b}, s)" + shows "\pia. ({atom (pi \ a)}, pi \ t) \gen R f pia ({atom (pi \ b)}, 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)+ + apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1]) + apply(simp add: a[symmetric] atom_eqvt eqvts) + apply(rule conjI) + apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1]) + apply(simp add: a[symmetric] eqvts atom_eqvt) + apply(subst permute_eqvt[symmetric]) + apply(simp) + done + + fun alpha_abs where