diff -r e4e8cbd9d780 -r abc140f21caa src/HOL/Isar_Examples/Peirce.thy --- a/src/HOL/Isar_Examples/Peirce.thy Thu Feb 20 23:16:33 2014 +0100 +++ b/src/HOL/Isar_Examples/Peirce.thy Thu Feb 20 23:46:40 2014 +0100 @@ -19,18 +19,18 @@ there is negation elimination; it holds in intuitionistic logic as well.} *} -theorem "((A --> B) --> A) --> A" +theorem "((A \ B) \ A) \ A" proof - assume "(A --> B) --> A" + assume "(A \ B) \ A" show A proof (rule classical) - assume "~ A" - have "A --> B" + assume "\ A" + have "A \ B" proof assume A - with `~ A` show B by contradiction + with `\ A` show B by contradiction qed - with `(A --> B) --> A` show A .. + with `(A \ B) \ A` show A .. qed qed @@ -48,19 +48,19 @@ contrast, strong assumptions (as introduced by \isacommand{assume}) are solved immediately. *} -theorem "((A --> B) --> A) --> A" +theorem "((A \ B) \ A) \ A" proof - assume "(A --> B) --> A" + assume "(A \ B) \ A" show A proof (rule classical) - presume "A --> B" - with `(A --> B) --> A` show A .. + presume "A \ B" + with `(A \ B) \ A` show A .. next - assume "~ A" - show "A --> B" + assume "\ A" + show "A \ B" proof assume A - with `~ A` show B by contradiction + with `\ A` show B by contradiction qed qed qed