--- a/src/HOL/Metis_Examples/Message.thy Wed Aug 28 20:46:45 2024 +0200
+++ b/src/HOL/Metis_Examples/Message.thy Wed Aug 28 22:54:45 2024 +0200
@@ -48,13 +48,16 @@
text\<open>Concrete syntax: messages appear as \<open>\<lbrace>A,B,NA\<rbrace>\<close>, etc...\<close>
+nonterminal mtuple_args
syntax
- "_MTuple" :: "['a, args] => 'a * 'b" ("(2\<lbrace>_,/ _\<rbrace>)")
+ "" :: "'a \<Rightarrow> mtuple_args" ("_")
+ "_MTuple_args" :: "'a \<Rightarrow> mtuple_args \<Rightarrow> mtuple_args" ("_,/ _")
+ "_MTuple" :: "['a, mtuple_args] \<Rightarrow> 'a * 'b" ("(2\<lbrace>_,/ _\<rbrace>)")
syntax_consts
- "_MTuple" == MPair
+ "_MTuple_args" "_MTuple" \<rightleftharpoons> MPair
translations
- "\<lbrace>x, y, z\<rbrace>" == "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
- "\<lbrace>x, y\<rbrace>" == "CONST MPair x y"
+ "\<lbrace>x, y, z\<rbrace>" \<rightleftharpoons> "\<lbrace>x, \<lbrace>y, z\<rbrace>\<rbrace>"
+ "\<lbrace>x, y\<rbrace>" \<rightleftharpoons> "CONST MPair x y"
definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where