src/FOL/ex/If.thy
changeset 62020 5d208fd2507d
parent 61489 b8d375aee0df
child 69590 e65314985426
--- a/src/FOL/ex/If.thy	Thu Dec 31 21:46:31 2015 +0100
+++ b/src/FOL/ex/If.thy	Fri Jan 01 10:49:00 2016 +0100
@@ -26,7 +26,7 @@
   apply (rule ifI)
   oops
 
-text\<open>Trying again from the beginning in order to use @{text blast}\<close>
+text\<open>Trying again from the beginning in order to use \<open>blast\<close>\<close>
 declare ifI [intro!]
 declare ifE [elim!]
 
@@ -45,7 +45,7 @@
 text \<open>An invalid formula. High-level rules permit a simpler diagnosis.\<close>
 lemma "if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,B,A))"
   apply auto
-    -- \<open>The next step will fail unless subgoals remain\<close>
+    \<comment> \<open>The next step will fail unless subgoals remain\<close>
   apply (tactic all_tac)
   oops
 
@@ -53,7 +53,7 @@
 lemma "if(if(P,Q,R), A, B) \<longleftrightarrow> if(P, if(Q,A,B), if(R,B,A))"
   unfolding if_def
   apply auto
-    -- \<open>The next step will fail unless subgoals remain\<close>
+    \<comment> \<open>The next step will fail unless subgoals remain\<close>
   apply (tactic all_tac)
   oops