--- a/src/Doc/Tutorial/Protocol/Message.thy Mon Sep 30 13:00:42 2024 +0200
+++ b/src/Doc/Tutorial/Protocol/Message.thy Mon Sep 30 20:30:59 2024 +0200
@@ -79,13 +79,10 @@
(*<*)
text\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
-nonterminal mtuple_args
syntax
- "" :: "'a \<Rightarrow> mtuple_args" ("_")
- "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" ("_,/ _")
- "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" ("(2\<lbrace>_,/ _\<rbrace>)")
+ "_MTuple" :: "['a, args] \<Rightarrow> 'a * 'b" ("(\<open>indent=2 notation=\<open>mixfix message tuple\<close>\<close>\<lbrace>_,/ _\<rbrace>)")
syntax_consts
- "_MTuple_args" "_MTuple" \<rightleftharpoons> MPair
+ MPair
translations
"\<lbrace>x, y, z\<rbrace>" \<rightleftharpoons> "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
"\<lbrace>x, y\<rbrace>" \<rightleftharpoons> "CONST MPair x y"