diff -r 713424d012fd -r 70076ba563d2 src/Doc/Tutorial/Protocol/Message.thy --- a/src/Doc/Tutorial/Protocol/Message.thy Wed Aug 28 20:46:45 2024 +0200 +++ b/src/Doc/Tutorial/Protocol/Message.thy Wed Aug 28 22:54:45 2024 +0200 @@ -79,13 +79,16 @@ (*<*) text\Concrete syntax: messages appear as \\A,B,NA\\, etc...\ +nonterminal mtuple_args syntax - "_MTuple" :: "['a, args] \ 'a * 'b" ("(2\_,/ _\)") + "" :: "'a \ mtuple_args" ("_") + "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" ("_,/ _") + "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" ("(2\_,/ _\)") syntax_consts - "_MTuple" == MPair + "_MTuple_args" "_MTuple" \ MPair translations - "\x, y, z\" == "\x, \y, z\\" - "\x, y\" == "CONST MPair x y" + "\x, y, z\" \ "\x, \y, z\\" + "\x, y\" \ "CONST MPair x y" definition keysFor :: "msg set \ key set" where