src/HOL/HOLCF/IOA/meta_theory/TL.thy
changeset 62002 f1599e98c4d0
parent 62001 1f2788fb0b8b
child 62004 8c6226d88ced
--- a/src/HOL/HOLCF/IOA/meta_theory/TL.thy	Wed Dec 30 21:56:12 2015 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/TL.thy	Wed Dec 30 21:57:52 2015 +0100
@@ -2,7 +2,7 @@
     Author:     Olaf Müller
 *)
 
-section {* A General Temporal Logic *}
+section \<open>A General Temporal Logic\<close>
 
 theory TL
 imports Pred Sequence
@@ -147,7 +147,7 @@
 "s~=UU & s~=nil --> tsuffix s2 (TL$s) --> tsuffix s2 s"
 apply (unfold tsuffix_def suffix_def)
 apply auto
-apply (tactic {* Seq_case_simp_tac @{context} "s" 1 *})
+apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>)
 apply (rule_tac x = "a\<leadsto>s1" in exI)
 apply auto
 done