diff -r ff2a637a449e -r fc5066122e68 src/ZF/ZF_Base.thy --- a/src/ZF/ZF_Base.thy Fri Oct 18 11:44:05 2024 +0200 +++ b/src/ZF/ZF_Base.thy Fri Oct 18 14:20:09 2024 +0200 @@ -56,6 +56,8 @@ syntax "_Replace" :: "[pttrn, pttrn, i, o] \ i" (\(\indent=1 notation=\mixfix relational replacement\\{_ ./ _ \ _, _})\) +syntax_consts + "_Replace" \ Replace translations "{y. x\A, Q}" \ "CONST Replace(A, \x y. Q)" @@ -67,6 +69,8 @@ syntax "_RepFun" :: "[i, pttrn, i] \ i" (\(\indent=1 notation=\mixfix functional replacement\\{_ ./ _ \ _})\ [51,0,51]) +syntax_consts + "_RepFun" \ RepFun translations "{b. x\A}" \ "CONST RepFun(A, \x. b)" @@ -77,6 +81,8 @@ syntax "_Collect" :: "[pttrn, i, o] \ i" (\(\indent=1 notation=\mixfix set comprehension\\{_ \ _ ./ _})\) +syntax_consts + "_Collect" \ Collect translations "{x\A. P}" \ "CONST Collect(A, \x. P)"