src/HOL/IMP/Def_Init_Small.thy
changeset 80914 d97fdabd9e2b
parent 67613 ce654b0e6d69
--- a/src/HOL/IMP/Def_Init_Small.thy	Fri Sep 20 19:07:10 2024 +0200
+++ b/src/HOL/IMP/Def_Init_Small.thy	Fri Sep 20 19:51:08 2024 +0200
@@ -7,7 +7,7 @@
 subsection "Initialization-Sensitive Small Step Semantics"
 
 inductive
-  small_step :: "(com \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix "\<rightarrow>" 55)
+  small_step :: "(com \<times> state) \<Rightarrow> (com \<times> state) \<Rightarrow> bool" (infix \<open>\<rightarrow>\<close> 55)
 where
 Assign:  "aval a s = Some i \<Longrightarrow> (x ::= a, s) \<rightarrow> (SKIP, s(x := Some i))" |
 
@@ -21,7 +21,7 @@
 
 lemmas small_step_induct = small_step.induct[split_format(complete)]
 
-abbreviation small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix "\<rightarrow>*" 55)
+abbreviation small_steps :: "com * state \<Rightarrow> com * state \<Rightarrow> bool" (infix \<open>\<rightarrow>*\<close> 55)
 where "x \<rightarrow>* y == star small_step x y"