diff -r cb16f22bc660 -r 1eac050a36f4 Nominal/Nominal2_Abs.thy --- a/Nominal/Nominal2_Abs.thy Fri Nov 26 10:53:55 2010 +0000 +++ b/Nominal/Nominal2_Abs.thy Fri Nov 26 19:03:23 2010 +0000 @@ -525,11 +525,11 @@ lemma Abs_fresh_star: fixes x::"'a::fs" - shows "as \* Abs_set as x" - and "as \* Abs_res as x" - and "set bs \* Abs_lst bs x" + shows "as \ as' \ as \* Abs_set as' x" + and "as \ as' \ as \* Abs_res as' x" + and "bs \ set bs' \ bs \* Abs_lst bs' x" unfolding fresh_star_def - by(simp_all add: Abs_fresh_iff) + by(auto simp add: Abs_fresh_iff) section {* Infrastructure for building tuples of relations and functions *}