diff -r e43e71a75838 -r 76560ce8dead src/HOL/Isar_Examples/Drinker.thy --- a/src/HOL/Isar_Examples/Drinker.thy Tue Jul 21 19:04:36 2015 +0200 +++ b/src/HOL/Isar_Examples/Drinker.thy Wed Jul 22 14:55:42 2015 +0200 @@ -33,17 +33,17 @@ theorem Drinker's_Principle: "\x. drunk x \ (\x. drunk x)" proof cases - fix a assume "\x. drunk x" - then have "drunk a \ (\x. drunk x)" .. + assume "\x. drunk x" + then have "drunk a \ (\x. drunk x)" for a .. then show ?thesis .. next assume "\ (\x. drunk x)" then have "\x. \ drunk x" by (rule de_Morgan) - then obtain a where a: "\ drunk a" .. + then obtain a where "\ drunk a" .. have "drunk a \ (\x. drunk x)" proof assume "drunk a" - with a show "\x. drunk x" by contradiction + with \\ drunk a\ show "\x. drunk x" by contradiction qed then show ?thesis .. qed