src/HOL/Imperative_HOL/Ref.thy
changeset 81706 7beb0cf38292
parent 80914 d97fdabd9e2b
--- a/src/HOL/Imperative_HOL/Ref.thy	Thu Jan 02 08:37:52 2025 +0100
+++ b/src/HOL/Imperative_HOL/Ref.thy	Thu Jan 02 08:37:55 2025 +0100
@@ -284,7 +284,7 @@
 code_printing constant Ref.update \<rightharpoonup> (SML) "(fn/ ()/ =>/ _/ :=/ _)"
 code_printing constant "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" \<rightharpoonup> (SML) infixl 6 "="
 
-code_reserved Eval Unsynchronized
+code_reserved (Eval) Unsynchronized
 
 
 text \<open>OCaml\<close>
@@ -296,7 +296,7 @@
 code_printing constant Ref.update \<rightharpoonup> (OCaml) "(fun/ ()/ ->/ _/ :=/ _)"
 code_printing constant "HOL.equal :: 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" \<rightharpoonup> (OCaml) infixl 4 "="
 
-code_reserved OCaml ref
+code_reserved (OCaml) ref
 
 
 text \<open>Haskell\<close>