--- a/src/HOL/Library/FSet.thy Wed Oct 02 13:34:03 2024 +0200
+++ b/src/HOL/Library/FSet.thy Wed Oct 02 14:23:28 2024 +0200
@@ -194,24 +194,47 @@
syntax (input)
"_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3! (_/|:|_)./ _)\<close> [0, 0, 10] 10)
"_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3? (_/|:|_)./ _)\<close> [0, 0, 10] 10)
+ "_fBex1" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3?! (_/:_)./ _)\<close> [0, 0, 10] 10)
syntax
"_fBall" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<forall>(_/|\<in>|_)./ _)\<close> [0, 0, 10] 10)
"_fBex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<exists>(_/|\<in>|_)./ _)\<close> [0, 0, 10] 10)
+ "_fBnex" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<nexists>(_/|\<in>|_)./ _)\<close> [0, 0, 10] 10)
+ "_fBex1" :: "pttrn \<Rightarrow> 'a fset \<Rightarrow> bool \<Rightarrow> bool" (\<open>(3\<exists>!(_/|\<in>|_)./ _)\<close> [0, 0, 10] 10)
syntax_consts
- "_fBall" \<rightleftharpoons> FSet.Ball and
- "_fBex" \<rightleftharpoons> FSet.Bex
+ "_fBall" "_fBnex" \<rightleftharpoons> fBall and
+ "_fBex" \<rightleftharpoons> fBex and
+ "_fBex1" \<rightleftharpoons> Ex1
translations
"\<forall>x|\<in>|A. P" \<rightleftharpoons> "CONST FSet.Ball A (\<lambda>x. P)"
"\<exists>x|\<in>|A. P" \<rightleftharpoons> "CONST FSet.Bex A (\<lambda>x. P)"
+ "\<nexists>x|\<in>|A. P" \<rightleftharpoons> "CONST fBall A (\<lambda>x. \<not> P)"
+ "\<exists>!x|\<in>|A. P" \<rightharpoonup> "\<exists>!x. x |\<in>| A \<and> P"
print_translation \<open>
[Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\<open>fBall\<close> \<^syntax_const>\<open>_fBall\<close>,
Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\<open>fBex\<close> \<^syntax_const>\<open>_fBex\<close>]
\<close> \<comment> \<open>to avoid eta-contraction of body\<close>
+syntax
+ "_setlessfAll" :: "[idt, 'a, bool] \<Rightarrow> bool" (\<open>(3\<forall>_|\<subset>|_./ _)\<close> [0, 0, 10] 10)
+ "_setlessfEx" :: "[idt, 'a, bool] \<Rightarrow> bool" (\<open>(3\<exists>_|\<subset>|_./ _)\<close> [0, 0, 10] 10)
+ "_setlefAll" :: "[idt, 'a, bool] \<Rightarrow> bool" (\<open>(3\<forall>_|\<subseteq>|_./ _)\<close> [0, 0, 10] 10)
+ "_setlefEx" :: "[idt, 'a, bool] \<Rightarrow> bool" (\<open>(3\<exists>_|\<subseteq>|_./ _)\<close> [0, 0, 10] 10)
+
+syntax_consts
+ "_setlessfAll" "_setlefAll" \<rightleftharpoons> All and
+ "_setlessfEx" "_setlefEx" \<rightleftharpoons> Ex
+
+translations
+ "\<forall>A|\<subset>|B. P" \<rightharpoonup> "\<forall>A. A |\<subset>| B \<longrightarrow> P"
+ "\<exists>A|\<subset>|B. P" \<rightharpoonup> "\<exists>A. A |\<subset>| B \<and> P"
+ "\<forall>A|\<subseteq>|B. P" \<rightharpoonup> "\<forall>A. A |\<subseteq>| B \<longrightarrow> P"
+ "\<exists>A|\<subseteq>|B. P" \<rightharpoonup> "\<exists>A. A |\<subseteq>| B \<and> P"
+
+
context includes lifting_syntax
begin