diff -r 4641f0fdaa59 -r be8c0e039a5e src/HOL/Set.thy --- a/src/HOL/Set.thy Sun Aug 25 12:43:43 2024 +0200 +++ b/src/HOL/Set.thy Sun Aug 25 15:02:19 2024 +0200 @@ -40,6 +40,8 @@ syntax "_Coll" :: "pttrn \ bool \ 'a set" ("(1{_./ _})") +syntax_consts + "_Coll" \ Collect translations "{x. P}" \ "CONST Collect (\x. P)" @@ -47,6 +49,8 @@ "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{(_/: _)./ _})") syntax "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{(_/ \ _)./ _})") +syntax_consts + "_Collect" \ Collect translations "{p:A. P}" \ "CONST Collect (\p. p \ A \ P)" @@ -141,6 +145,8 @@ syntax "_Finset" :: "args \ 'a set" ("{(_)}") +syntax_consts + "_Finset" \ insert translations "{x, xs}" \ "CONST insert x {xs}" "{x}" \ "CONST insert x {}" @@ -203,6 +209,12 @@ "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3\!(_/\_)./ _)" [0, 0, 10] 10) "_Bleast" :: "id \ 'a set \ bool \ 'a" ("(3LEAST(_/\_)./ _)" [0, 0, 10] 10) +syntax_consts + "_Ball" \ Ball and + "_Bex" \ Bex and + "_Bex1" \ Ex1 and + "_Bleast" \ Least + translations "\x\A. P" \ "CONST Ball A (\x. P)" "\x\A. P" \ "CONST Bex A (\x. P)" @@ -223,6 +235,11 @@ "_setleEx" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleEx1" :: "[idt, 'a, bool] \ bool" ("(3\!_\_./ _)" [0, 0, 10] 10) +syntax_consts + "_setlessAll" "_setleAll" \ All and + "_setlessEx" "_setleEx" \ Ex and + "_setleEx1" \ Ex1 + translations "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" @@ -272,6 +289,8 @@ syntax "_Setcompr" :: "'a \ idts \ bool \ 'a set" ("(1{_ |/_./ _})") +syntax_consts + "_Setcompr" \ Collect parse_translation \ let