src/HOL/Library/Open_State_Syntax.thy
changeset 81187 c66e24eae281
parent 80914 d97fdabd9e2b
--- a/src/HOL/Library/Open_State_Syntax.thy	Fri Oct 18 15:42:31 2024 +0200
+++ b/src/HOL/Library/Open_State_Syntax.thy	Fri Oct 18 15:45:38 2024 +0200
@@ -119,22 +119,23 @@
 
 nonterminal sdo_binds and sdo_bind
 
-definition "sdo_syntax = ()"
-
 syntax
-  "_sdo_block" :: "sdo_binds \<Rightarrow> 'a" (\<open>exec {//(2  _)//}\<close> [12] 62)
-  "_sdo_bind"  :: "[pttrn, 'a] \<Rightarrow> sdo_bind" (\<open>(_ \<leftarrow>/ _)\<close> 13)
-  "_sdo_let" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" (\<open>(2let _ =/ _)\<close> [1000, 13] 13)
-  "_sdo_then" :: "'a \<Rightarrow> sdo_bind" (\<open>_\<close> [14] 13)
-  "_sdo_final" :: "'a \<Rightarrow> sdo_binds" (\<open>_\<close>)
-  "_sdo_cons" :: "[sdo_bind, sdo_binds] \<Rightarrow> sdo_binds" (\<open>_;//_\<close> [13, 12] 12)
+  "_sdo_block" :: "sdo_binds \<Rightarrow> 'a"
+    (\<open>(\<open>open_block notation=\<open>mixfix exec block\<close>\<close>exec {//(2  _)//})\<close> [12] 62)
+  "_sdo_bind"  :: "[pttrn, 'a] \<Rightarrow> sdo_bind"
+    (\<open>(\<open>indent=2 notation=\<open>infix exec bind\<close>\<close>_ \<leftarrow>/ _)\<close> 13)
+  "_sdo_let" :: "[pttrn, 'a] \<Rightarrow> sdo_bind"
+    (\<open>(\<open>indent=2 notation=\<open>infix exec let\<close>\<close>let _ =/ _)\<close> [1000, 13] 13)
+  "_sdo_then" :: "'a \<Rightarrow> sdo_bind"  (\<open>_\<close> [14] 13)
+  "_sdo_final" :: "'a \<Rightarrow> sdo_binds"  (\<open>_\<close>)
+  "_sdo_cons" :: "[sdo_bind, sdo_binds] \<Rightarrow> sdo_binds"
+    (\<open>(\<open>open_block notation=\<open>infix exec next\<close>\<close>_;//_)\<close> [13, 12] 12)
 
 syntax (ASCII)
-  "_sdo_bind" :: "[pttrn, 'a] \<Rightarrow> sdo_bind" (\<open>(_ <-/ _)\<close> 13)
+  "_sdo_bind" :: "[pttrn, 'a] \<Rightarrow> sdo_bind"
+    (\<open>(\<open>indent=2 notation=\<open>infix exec bind\<close>\<close>_ <-/ _)\<close> 13)
 
 syntax_consts
-  "_sdo_block" "_sdo_cons" == sdo_syntax and
-  "_sdo_bind" == scomp and
   "_sdo_let" == Let
 
 translations