src/HOL/Set.thy
author haftmann
Mon, 09 Mar 2009 14:43:51 +0100
changeset 30387 0affb9975547
parent 30352 047f183c43b0
child 30531 ab3d61baf66a
permissions -rw-r--r--
NameSpace.base_name ~> Long_Name.base_name

(*  Title:      HOL/Set.thy
    Author:     Tobias Nipkow, Lawrence C Paulson and Markus Wenzel
*)

header {* Set theory for higher-order logic *}

theory Set
imports Lattices
begin

subsection {* Basic operations *}

subsubsection {* Comprehension and membership *}

text {* A set in HOL is simply a predicate. *}

global

types 'a set = "'a => bool"

consts
  Collect :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set"
  "op :" :: "'a => 'a set => bool"

local

syntax
  "@Coll"       :: "pttrn => bool => 'a set"              ("(1{_./ _})")

translations
  "{x. P}"      == "Collect (%x. P)"

notation
  "op :"  ("op :") and
  "op :"  ("(_/ : _)" [50, 51] 50)

abbreviation
  "not_mem x A == ~ (x : A)" -- "non-membership"

notation
  not_mem  ("op ~:") and
  not_mem  ("(_/ ~: _)" [50, 51] 50)

notation (xsymbols)
  "op :"  ("op \<in>") and
  "op :"  ("(_/ \<in> _)" [50, 51] 50) and
  not_mem  ("op \<notin>") and
  not_mem  ("(_/ \<notin> _)" [50, 51] 50)

notation (HTML output)
  "op :"  ("op \<in>") and
  "op :"  ("(_/ \<in> _)" [50, 51] 50) and
  not_mem  ("op \<notin>") and
  not_mem  ("(_/ \<notin> _)" [50, 51] 50)

defs
  Collect_def [code]: "Collect P \<equiv> P"
  mem_def [code]: "x \<in> S \<equiv> S x"

text {* Relating predicates and sets *}

lemma mem_Collect_eq [iff]: "(a : {x. P(x)}) = P(a)"
  by (simp add: Collect_def mem_def)

lemma Collect_mem_eq [simp]: "{x. x:A} = A"
  by (simp add: Collect_def mem_def)

lemma CollectI: "P(a) ==> a : {x. P(x)}"
  by simp

lemma CollectD: "a : {x. P(x)} ==> P(a)"
  by simp

lemma Collect_cong: "(!!x. P x = Q x) ==> {x. P(x)} = {x. Q(x)}"
  by simp

lemmas CollectE = CollectD [elim_format]

lemma set_ext: assumes prem: "(!!x. (x:A) = (x:B))" shows "A = B"
  apply (rule prem [THEN ext, THEN arg_cong, THEN box_equals])
   apply (rule Collect_mem_eq)
  apply (rule Collect_mem_eq)
  done

(* Due to Brian Huffman *)
lemma expand_set_eq: "(A = B) = (ALL x. (x:A) = (x:B))"
by(auto intro:set_ext)

lemma equalityCE [elim]:
    "A = B ==> (c \<in> A ==> c \<in> B ==> P) ==> (c \<notin> A ==> c \<notin> B ==> P) ==> P"
  by blast

lemma eqset_imp_iff: "A = B ==> (x : A) = (x : B)"
  by simp

lemma eqelem_imp_iff: "x = y ==> (x : A) = (y : A)"
  by simp


subsubsection {* Subset relation, empty and universal set *}

abbreviation
  subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
  "subset \<equiv> less"

abbreviation
  subset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
  "subset_eq \<equiv> less_eq"

notation (output)
  subset  ("op <") and
  subset  ("(_/ < _)" [50, 51] 50) and
  subset_eq  ("op <=") and
  subset_eq  ("(_/ <= _)" [50, 51] 50)

notation (xsymbols)
  subset  ("op \<subset>") and
  subset  ("(_/ \<subset> _)" [50, 51] 50) and
  subset_eq  ("op \<subseteq>") and
  subset_eq  ("(_/ \<subseteq> _)" [50, 51] 50)

notation (HTML output)
  subset  ("op \<subset>") and
  subset  ("(_/ \<subset> _)" [50, 51] 50) and
  subset_eq  ("op \<subseteq>") and
  subset_eq  ("(_/ \<subseteq> _)" [50, 51] 50)

abbreviation (input)
  supset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
  "supset \<equiv> greater"

abbreviation (input)
  supset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
  "supset_eq \<equiv> greater_eq"

notation (xsymbols)
  supset  ("op \<supset>") and
  supset  ("(_/ \<supset> _)" [50, 51] 50) and
  supset_eq  ("op \<supseteq>") and
  supset_eq  ("(_/ \<supseteq> _)" [50, 51] 50)

definition empty :: "'a set" ("{}") where
  "empty \<equiv> {x. False}"

definition UNIV :: "'a set" where
  "UNIV \<equiv> {x. True}"
 
lemma subsetI [atp,intro!]: "(!!x. x:A ==> x:B) ==> A \<subseteq> B"
  by (auto simp add: mem_def intro: predicate1I)

text {*
  \medskip Map the type @{text "'a set => anything"} to just @{typ
  'a}; for overloading constants whose first argument has type @{typ
  "'a set"}.
*}

lemma subsetD [elim]: "A \<subseteq> B ==> c \<in> A ==> c \<in> B"
  -- {* Rule in Modus Ponens style. *}
  by (unfold mem_def) blast

declare subsetD [intro?] -- FIXME

lemma rev_subsetD: "c \<in> A ==> A \<subseteq> B ==> c \<in> B"
  -- {* The same, with reversed premises for use with @{text erule} --
      cf @{text rev_mp}. *}
  by (rule subsetD)

declare rev_subsetD [intro?] -- FIXME

text {*
  \medskip Converts @{prop "A \<subseteq> B"} to @{prop "x \<in> A ==> x \<in> B"}.
*}

ML {*
  fun impOfSubs th = th RSN (2, @{thm rev_subsetD})
*}

lemma subsetCE [elim]: "A \<subseteq>  B ==> (c \<notin> A ==> P) ==> (c \<in> B ==> P) ==> P"
  -- {* Classical elimination rule. *}
  by (unfold mem_def) blast

text {*
  \medskip Takes assumptions @{prop "A \<subseteq> B"}; @{prop "c \<in> A"} and
  creates the assumption @{prop "c \<in> B"}.
*}

ML {*
  fun set_mp_tac i = etac @{thm subsetCE} i THEN mp_tac i
*}

lemma contra_subsetD: "A \<subseteq> B ==> c \<notin> B ==> c \<notin> A"
  by blast

lemma subset_refl [simp,atp]: "A \<subseteq> A"
  by fast

lemma subset_trans: "A \<subseteq> B ==> B \<subseteq> C ==> A \<subseteq> C"
  by blast

lemma subset_antisym [intro!]: "A \<subseteq> B ==> B \<subseteq> A ==> A = B"
  -- {* Anti-symmetry of the subset relation. *}
  by (iprover intro: set_ext subsetD)

text {*
  \medskip Equality rules from ZF set theory -- are they appropriate
  here?
*}

lemma equalityD1: "A = B ==> A \<subseteq> B"
  by (simp add: subset_refl)

lemma equalityD2: "A = B ==> B \<subseteq> A"
  by (simp add: subset_refl)

text {*
  \medskip Be careful when adding this to the claset as @{text
  subset_empty} is in the simpset: @{prop "A = {}"} goes to @{prop "{}
  \<subseteq> A"} and @{prop "A \<subseteq> {}"} and then back to @{prop "A = {}"}!
*}

lemma equalityE: "A = B ==> (A \<subseteq> B ==> B \<subseteq> A ==> P) ==> P"
  by (simp add: subset_refl)

lemma empty_iff [simp]: "(c : {}) = False"
  by (simp add: empty_def)

lemma emptyE [elim!]: "a : {} ==> P"
  by simp

lemma empty_subsetI [iff]: "{} \<subseteq> A"
    -- {* One effect is to delete the ASSUMPTION @{prop "{} \<subseteq> A"} *}
  by blast

lemma bot_set_eq: "bot = {}"
  by (iprover intro!: subset_antisym empty_subsetI bot_least)

lemma equals0I: "(!!y. y \<in> A ==> False) ==> A = {}"
  by blast

lemma equals0D: "A = {} ==> a \<notin> A"
    -- {* Use for reasoning about disjointness *}
  by blast

lemma UNIV_I [simp]: "x : UNIV"
  by (simp add: UNIV_def)

declare UNIV_I [intro]  -- {* unsafe makes it less likely to cause problems *}

lemma UNIV_witness [intro?]: "EX x. x : UNIV"
  by simp

lemma subset_UNIV [simp]: "A \<subseteq> UNIV"
  by (rule subsetI) (rule UNIV_I)

lemma top_set_eq: "top = UNIV"
  by (iprover intro!: subset_antisym subset_UNIV top_greatest)

lemma UNIV_eq_I: "(\<And>x. x \<in> A) \<Longrightarrow> UNIV = A"
  by auto

lemma UNIV_not_empty [iff]: "UNIV ~= {}"
  by blast

lemma psubsetI [intro!,noatp]: "A \<subseteq> B ==> A \<noteq> B ==> A \<subset> B"
  by (unfold less_le) blast

lemma psubsetE [elim!,noatp]: 
    "[|A \<subset> B;  [|A \<subseteq> B; ~ (B\<subseteq>A)|] ==> R|] ==> R"
  by (unfold less_le) blast

lemma psubset_eq: "(A \<subset> B) = (A \<subseteq> B & A \<noteq> B)"
  by (simp only: less_le)

lemma psubset_imp_subset: "A \<subset> B ==> A \<subseteq> B"
  by (simp add: psubset_eq)

lemma psubset_trans: "[| A \<subset> B; B \<subset> C |] ==> A \<subset> C"
apply (unfold less_le)
apply (auto dest: subset_antisym)
done

lemma psubsetD: "[| A \<subset> B; c \<in> A |] ==> c \<in> B"
apply (unfold less_le)
apply (auto dest: subsetD)
done

lemma psubset_subset_trans: "A \<subset> B ==> B \<subseteq> C ==> A \<subset> C"
  by (auto simp add: psubset_eq)

lemma subset_psubset_trans: "A \<subseteq> B ==> B \<subset> C ==> A \<subset> C"
  by (auto simp add: psubset_eq)

lemma Collect_empty_eq [simp]: "(Collect P = {}) = (\<forall>x. \<not> P x)"
by blast

lemma empty_Collect_eq [simp]: "({} = Collect P) = (\<forall>x. \<not> P x)"
by blast

subsubsection {* Intersection and union *}

definition Int :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Int" 70) where
  "A Int B \<equiv> {x. x \<in> A \<and> x \<in> B}"

definition Un :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" (infixl "Un" 65) where
  "A Un B \<equiv> {x. x \<in> A \<or> x \<in> B}"

notation (xsymbols)
  "Int"  (infixl "\<inter>" 70) and
  "Un"  (infixl "\<union>" 65)

notation (HTML output)
  "Int"  (infixl "\<inter>" 70) and
  "Un"  (infixl "\<union>" 65)

lemma Int_iff [simp]: "(c : A Int B) = (c:A & c:B)"
  by (unfold Int_def) blast

lemma IntI [intro!]: "c:A ==> c:B ==> c : A Int B"
  by simp

lemma IntD1: "c : A Int B ==> c:A"
  by simp

lemma IntD2: "c : A Int B ==> c:B"
  by simp

lemma IntE [elim!]: "c : A Int B ==> (c:A ==> c:B ==> P) ==> P"
  by simp

lemma Un_iff [simp]: "(c : A Un B) = (c:A | c:B)"
  by (unfold Un_def) blast

lemma UnI1 [elim?]: "c:A ==> c : A Un B"
  by simp

lemma UnI2 [elim?]: "c:B ==> c : A Un B"
  by simp

text {*
  \medskip Classical introduction rule: no commitment to @{prop A} vs
  @{prop B}.
*}

lemma UnCI [intro!]: "(c~:B ==> c:A) ==> c : A Un B"
  by auto

lemma UnE [elim!]: "c : A Un B ==> (c:A ==> P) ==> (c:B ==> P) ==> P"
  by (unfold Un_def) blast

lemma Int_lower1: "A \<inter> B \<subseteq> A"
  by blast

lemma Int_lower2: "A \<inter> B \<subseteq> B"
  by blast

lemma Int_greatest: "C \<subseteq> A ==> C \<subseteq> B ==> C \<subseteq> A \<inter> B"
  by blast

lemma inf_set_eq: "inf A B = A \<inter> B"
  apply (rule subset_antisym)
  apply (rule Int_greatest)
  apply (rule inf_le1)
  apply (rule inf_le2)
  apply (rule inf_greatest)
  apply (rule Int_lower1)
  apply (rule Int_lower2)
  done

lemma Un_upper1: "A \<subseteq> A \<union> B"
  by blast

lemma Un_upper2: "B \<subseteq> A \<union> B"
  by blast

lemma Un_least: "A \<subseteq> C ==> B \<subseteq> C ==> A \<union> B \<subseteq> C"
  by blast

lemma sup_set_eq: "sup A B = A \<union> B"
  apply (rule subset_antisym)
  apply (rule sup_least)
  apply (rule Un_upper1)
  apply (rule Un_upper2)
  apply (rule Un_least)
  apply (rule sup_ge1)
  apply (rule sup_ge2)
  done

lemma Collect_conj_eq: "{x. P x & Q x} = {x. P x} \<inter> {x. Q x}"
  by blast

lemma Collect_disj_eq: "{x. P x | Q x} = {x. P x} \<union> {x. Q x}"
  by blast

lemma subset_empty [simp]: "(A \<subseteq> {}) = (A = {})"
  by blast

lemma not_psubset_empty [iff]: "\<not> (A \<subset> {})"
  by (unfold less_le) blast

lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})"
  -- {* supersedes @{text "Collect_False_empty"} *}
  by auto


subsubsection {* Complement and set difference *}

instantiation bool :: minus
begin

definition
  bool_diff_def: "A - B \<longleftrightarrow> A \<and> \<not> B"

instance ..

end

instantiation "fun" :: (type, minus) minus
begin

definition
  fun_diff_def: "A - B = (\<lambda>x. A x - B x)"

instance ..

end

instantiation bool :: uminus
begin

definition
  bool_Compl_def: "- A \<longleftrightarrow> \<not> A"

instance ..

end

instantiation "fun" :: (type, uminus) uminus
begin

definition
  fun_Compl_def: "- A = (\<lambda>x. - A x)"

instance ..

end

lemma Compl_iff [simp]: "(c \<in> -A) = (c \<notin> A)"
  by (simp add: mem_def fun_Compl_def bool_Compl_def)

lemma ComplI [intro!]: "(c \<in> A ==> False) ==> c \<in> -A"
  by (unfold mem_def fun_Compl_def bool_Compl_def) blast

text {*
  \medskip This form, with negated conclusion, works well with the
  Classical prover.  Negated assumptions behave like formulae on the
  right side of the notional turnstile ... *}

lemma ComplD [dest!]: "c : -A ==> c~:A"
  by (simp add: mem_def fun_Compl_def bool_Compl_def)

lemmas ComplE = ComplD [elim_format]

lemma Compl_eq: "- A = {x. ~ x : A}" by blast

lemma Diff_iff [simp]: "(c : A - B) = (c:A & c~:B)"
  by (simp add: mem_def fun_diff_def bool_diff_def)

lemma DiffI [intro!]: "c : A ==> c ~: B ==> c : A - B"
  by simp

lemma DiffD1: "c : A - B ==> c : A"
  by simp

lemma DiffD2: "c : A - B ==> c : B ==> P"
  by simp

lemma DiffE [elim!]: "c : A - B ==> (c:A ==> c~:B ==> P) ==> P"
  by simp

lemma set_diff_eq: "A - B = {x. x : A & ~ x : B}" by blast

lemma Compl_eq_Diff_UNIV: "-A = (UNIV - A)"
by blast

lemma psubset_imp_ex_mem: "A \<subset> B ==> \<exists>b. b \<in> (B - A)"
  by (unfold less_le) blast

lemma Diff_subset: "A - B \<subseteq> A"
  by blast

lemma Diff_subset_conv: "(A - B \<subseteq> C) = (A \<subseteq> B \<union> C)"
by blast

lemma Collect_imp_eq: "{x. P x --> Q x} = -{x. P x} \<union> {x. Q x}"
  by blast

lemma Collect_neg_eq: "{x. \<not> P x} = - {x. P x}"
  by blast


subsubsection {* Set enumerations *}

global

consts
  insert        :: "'a => 'a set => 'a set"

local

defs
  insert_def:   "insert a B == {x. x=a} Un B"

syntax
  "@Finset"     :: "args => 'a set"                       ("{(_)}")

translations
  "{x, xs}"     == "insert x {xs}"
  "{x}"         == "insert x {}"

lemma insert_iff [simp]: "(a : insert b A) = (a = b | a:A)"
  by (unfold insert_def) blast

lemma insertI1: "a : insert a B"
  by simp

lemma insertI2: "a : B ==> a : insert b B"
  by simp

lemma insertE [elim!]: "a : insert b A ==> (a = b ==> P) ==> (a:A ==> P) ==> P"
  by (unfold insert_def) blast

lemma insertCI [intro!]: "(a~:B ==> a = b) ==> a: insert b B"
  -- {* Classical introduction rule. *}
  by auto

lemma subset_insert_iff: "(A \<subseteq> insert x B) = (if x:A then A - {x} \<subseteq> B else A \<subseteq> B)"
  by auto

lemma set_insert:
  assumes "x \<in> A"
  obtains B where "A = insert x B" and "x \<notin> B"
proof
  from assms show "A = insert x (A - {x})" by blast
next
  show "x \<notin> A - {x}" by blast
qed

lemma insert_ident: "x ~: A ==> x ~: B ==> (insert x A = insert x B) = (A = B)"
by auto

lemma insert_is_Un: "insert a A = {a} Un A"
  -- {* NOT SUITABLE FOR REWRITING since @{text "{a} == insert a {}"} *}
  by blast

lemma insert_not_empty [simp]: "insert a A \<noteq> {}"
  by blast

lemmas empty_not_insert = insert_not_empty [symmetric, standard]
declare empty_not_insert [simp]

lemma insert_absorb: "a \<in> A ==> insert a A = A"
  -- {* @{text "[simp]"} causes recursive calls when there are nested inserts *}
  -- {* with \emph{quadratic} running time *}
  by blast

lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A"
  by blast

lemma insert_commute: "insert x (insert y A) = insert y (insert x A)"
  by blast

lemma insert_subset [simp]: "(insert x A \<subseteq> B) = (x \<in> B & A \<subseteq> B)"
  by blast

lemma mk_disjoint_insert: "a \<in> A ==> \<exists>B. A = insert a B & a \<notin> B"
  -- {* use new @{text B} rather than @{text "A - {a}"} to avoid infinite unfolding *}
  apply (rule_tac x = "A - {a}" in exI, blast)
  done

lemma insert_Collect: "insert a (Collect P) = {u. u \<noteq> a --> P u}"
  by auto

lemma insert_inter_insert[simp]: "insert a A \<inter> insert a B = insert a (A \<inter> B)"
  by blast

lemma insert_disjoint [simp,noatp]:
 "(insert a A \<inter> B = {}) = (a \<notin> B \<and> A \<inter> B = {})"
 "({} = insert a A \<inter> B) = (a \<notin> B \<and> {} = A \<inter> B)"
  by auto

lemma disjoint_insert [simp,noatp]:
 "(B \<inter> insert a A = {}) = (a \<notin> B \<and> B \<inter> A = {})"
 "({} = A \<inter> insert b B) = (b \<notin> A \<and> {} = A \<inter> B)"
  by auto

text {* Singletons, using insert *}

lemma singletonI [intro!,noatp]: "a : {a}"
    -- {* Redundant? But unlike @{text insertCI}, it proves the subgoal immediately! *}
  by (rule insertI1)

lemma singletonD [dest!,noatp]: "b : {a} ==> b = a"
  by blast

lemmas singletonE = singletonD [elim_format]

lemma singleton_iff: "(b : {a}) = (b = a)"
  by blast

lemma singleton_inject [dest!]: "{a} = {b} ==> a = b"
  by blast

lemma singleton_insert_inj_eq [iff,noatp]:
     "({b} = insert a A) = (a = b & A \<subseteq> {b})"
  by blast

lemma singleton_insert_inj_eq' [iff,noatp]:
     "(insert a A = {b}) = (a = b & A \<subseteq> {b})"
  by blast

lemma subset_singletonD: "A \<subseteq> {x} ==> A = {} | A = {x}"
  by fast

lemma singleton_conv [simp]: "{x. x = a} = {a}"
  by blast

lemma singleton_conv2 [simp]: "{x. a = x} = {a}"
  by blast

lemma diff_single_insert: "A - {x} \<subseteq> B ==> x \<in> A ==> A \<subseteq> insert x B"
  by blast

lemma doubleton_eq_iff: "({a,b} = {c,d}) = (a=c & b=d | a=d & b=c)"
  by (blast elim: equalityE)

lemma psubset_insert_iff:
  "(A \<subset> insert x B) = (if x \<in> B then A \<subset> B else if x \<in> A then A - {x} \<subset> B else A \<subseteq> B)"
  by (auto simp add: less_le subset_insert_iff)

lemma subset_insertI: "B \<subseteq> insert a B"
  by (rule subsetI) (erule insertI2)

lemma subset_insertI2: "A \<subseteq> B \<Longrightarrow> A \<subseteq> insert b B"
  by blast

lemma subset_insert: "x \<notin> A ==> (A \<subseteq> insert x B) = (A \<subseteq> B)"
  by blast


subsubsection {* Bounded quantifiers and operators *}

global

consts
  Ball          :: "'a set => ('a => bool) => bool"      -- "bounded universal quantifiers"
  Bex           :: "'a set => ('a => bool) => bool"      -- "bounded existential quantifiers"
  Bex1          :: "'a set => ('a => bool) => bool"      -- "bounded unique existential quantifiers"

local

defs
  Ball_def:     "Ball A P       == ALL x. x:A --> P(x)"
  Bex_def:      "Bex A P        == EX x. x:A & P(x)"
  Bex1_def:     "Bex1 A P       == EX! x. x:A & P(x)"

syntax
  "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3ALL _:_./ _)" [0, 0, 10] 10)
  "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3EX _:_./ _)" [0, 0, 10] 10)
  "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3EX! _:_./ _)" [0, 0, 10] 10)
  "_Bleast"     :: "id => 'a set => bool => 'a"           ("(3LEAST _:_./ _)" [0, 0, 10] 10)

syntax (HOL)
  "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3! _:_./ _)" [0, 0, 10] 10)
  "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3? _:_./ _)" [0, 0, 10] 10)
  "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3?! _:_./ _)" [0, 0, 10] 10)

syntax (xsymbols)
  "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3\<forall>_\<in>_./ _)" [0, 0, 10] 10)
  "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3\<exists>_\<in>_./ _)" [0, 0, 10] 10)
  "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3\<exists>!_\<in>_./ _)" [0, 0, 10] 10)
  "_Bleast"     :: "id => 'a set => bool => 'a"           ("(3LEAST_\<in>_./ _)" [0, 0, 10] 10)

syntax (HTML output)
  "_Ball"       :: "pttrn => 'a set => bool => bool"      ("(3\<forall>_\<in>_./ _)" [0, 0, 10] 10)
  "_Bex"        :: "pttrn => 'a set => bool => bool"      ("(3\<exists>_\<in>_./ _)" [0, 0, 10] 10)
  "_Bex1"       :: "pttrn => 'a set => bool => bool"      ("(3\<exists>!_\<in>_./ _)" [0, 0, 10] 10)

translations
  "ALL x:A. P"  == "Ball A (%x. P)"
  "EX x:A. P"   == "Bex A (%x. P)"
  "EX! x:A. P"  == "Bex1 A (%x. P)"
  "LEAST x:A. P" => "LEAST x. x:A & P"

syntax (output)
  "_setlessAll" :: "[idt, 'a, bool] => bool"  ("(3ALL _<_./ _)"  [0, 0, 10] 10)
  "_setlessEx"  :: "[idt, 'a, bool] => bool"  ("(3EX _<_./ _)"  [0, 0, 10] 10)
  "_setleAll"   :: "[idt, 'a, bool] => bool"  ("(3ALL _<=_./ _)" [0, 0, 10] 10)
  "_setleEx"    :: "[idt, 'a, bool] => bool"  ("(3EX _<=_./ _)" [0, 0, 10] 10)
  "_setleEx1"   :: "[idt, 'a, bool] => bool"  ("(3EX! _<=_./ _)" [0, 0, 10] 10)

syntax (xsymbols)
  "_setlessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<subset>_./ _)"  [0, 0, 10] 10)
  "_setlessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<subset>_./ _)"  [0, 0, 10] 10)
  "_setleAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<subseteq>_./ _)" [0, 0, 10] 10)
  "_setleEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<subseteq>_./ _)" [0, 0, 10] 10)
  "_setleEx1"   :: "[idt, 'a, bool] => bool"   ("(3\<exists>!_\<subseteq>_./ _)" [0, 0, 10] 10)

syntax (HOL output)
  "_setlessAll" :: "[idt, 'a, bool] => bool"   ("(3! _<_./ _)"  [0, 0, 10] 10)
  "_setlessEx"  :: "[idt, 'a, bool] => bool"   ("(3? _<_./ _)"  [0, 0, 10] 10)
  "_setleAll"   :: "[idt, 'a, bool] => bool"   ("(3! _<=_./ _)" [0, 0, 10] 10)
  "_setleEx"    :: "[idt, 'a, bool] => bool"   ("(3? _<=_./ _)" [0, 0, 10] 10)
  "_setleEx1"   :: "[idt, 'a, bool] => bool"   ("(3?! _<=_./ _)" [0, 0, 10] 10)

syntax (HTML output)
  "_setlessAll" :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<subset>_./ _)"  [0, 0, 10] 10)
  "_setlessEx"  :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<subset>_./ _)"  [0, 0, 10] 10)
  "_setleAll"   :: "[idt, 'a, bool] => bool"   ("(3\<forall>_\<subseteq>_./ _)" [0, 0, 10] 10)
  "_setleEx"    :: "[idt, 'a, bool] => bool"   ("(3\<exists>_\<subseteq>_./ _)" [0, 0, 10] 10)
  "_setleEx1"   :: "[idt, 'a, bool] => bool"   ("(3\<exists>!_\<subseteq>_./ _)" [0, 0, 10] 10)

translations
  "\<forall>A\<subset>B. P"   =>  "ALL A. A \<subset> B --> P"
  "\<exists>A\<subset>B. P"   =>  "EX A. A \<subset> B & P"
  "\<forall>A\<subseteq>B. P"   =>  "ALL A. A \<subseteq> B --> P"
  "\<exists>A\<subseteq>B. P"   =>  "EX A. A \<subseteq> B & P"
  "\<exists>!A\<subseteq>B. P"  =>  "EX! A. A \<subseteq> B & P"

print_translation {*
let
  val Type (set_type, _) = @{typ "'a set"};
  val All_binder = Syntax.binder_name @{const_syntax "All"};
  val Ex_binder = Syntax.binder_name @{const_syntax "Ex"};
  val impl = @{const_syntax "op -->"};
  val conj = @{const_syntax "op &"};
  val sbset = @{const_syntax "subset"};
  val sbset_eq = @{const_syntax "subset_eq"};

  val trans =
   [((All_binder, impl, sbset), "_setlessAll"),
    ((All_binder, impl, sbset_eq), "_setleAll"),
    ((Ex_binder, conj, sbset), "_setlessEx"),
    ((Ex_binder, conj, sbset_eq), "_setleEx")];

  fun mk v v' c n P =
    if v = v' andalso not (Term.exists_subterm (fn Free (x, _) => x = v | _ => false) n)
    then Syntax.const c $ Syntax.mark_bound v' $ n $ P else raise Match;

  fun tr' q = (q,
    fn [Const ("_bound", _) $ Free (v, Type (T, _)), Const (c, _) $ (Const (d, _) $ (Const ("_bound", _) $ Free (v', _)) $ n) $ P] =>
         if T = (set_type) then case AList.lookup (op =) trans (q, c, d)
          of NONE => raise Match
           | SOME l => mk v v' l n P
         else raise Match
     | _ => raise Match);
in
  [tr' All_binder, tr' Ex_binder]
end
*}

text {*
  \medskip Translate between @{text "{e | x1...xn. P}"} and @{text
  "{u. EX x1..xn. u = e & P}"}; @{text "{y. EX x1..xn. y = e & P}"} is
  only translated if @{text "[0..n] subset bvs(e)"}.