diff -r 46f59511b7bb -r d97fdabd9e2b src/HOL/Metis_Examples/Message.thy --- a/src/HOL/Metis_Examples/Message.thy Fri Sep 20 19:07:10 2024 +0200 +++ b/src/HOL/Metis_Examples/Message.thy Fri Sep 20 19:51:08 2024 +0200 @@ -50,9 +50,9 @@ text\Concrete syntax: messages appear as \\A,B,NA\\, etc...\ nonterminal mtuple_args syntax - "" :: "'a \ mtuple_args" ("_") - "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" ("_,/ _") - "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" ("(2\_,/ _\)") + "" :: "'a \ mtuple_args" (\_\) + "_MTuple_args" :: "'a \ mtuple_args \ mtuple_args" (\_,/ _\) + "_MTuple" :: "['a, mtuple_args] \ 'a * 'b" (\(2\_,/ _\)\) syntax_consts "_MTuple_args" "_MTuple" \ MPair translations @@ -60,7 +60,7 @@ "\x, y\" \ "CONST MPair x y" -definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where +definition HPair :: "[msg,msg] => msg" (\(4Hash[_] /_)\ [0, 1000]) where \ \Message Y paired with a MAC computed with the help of X\ "Hash[X] Y == \ Hash\X,Y\, Y\"