Theory HOL.Set
section ‹Set theory for higher-order logic›
theory Set
imports Lattices Boolean_Algebras
begin
subsection ‹Sets as predicates›
typedecl 'a set
axiomatization Collect :: "('a ⇒ bool) ⇒ 'a set"
and member :: "'a ⇒ 'a set ⇒ bool"
where mem_Collect_eq [iff, code_unfold]: "member a (Collect P) = P a"
and Collect_mem_eq [simp]: "Collect (λx. member x A) = A"
notation
member ("'(∈')") and
member ("(_/ ∈ _)" [51, 51] 50)
abbreviation not_member
where "not_member x A ≡ ¬ (x ∈ A)"
notation
not_member ("'(∉')") and
not_member ("(_/ ∉ _)" [51, 51] 50)
notation (ASCII)
member ("'(:')") and
member ("(_/ : _)" [51, 51] 50) and
not_member ("'(~:')") and
not_member ("(_/ ~: _)" [51, 51] 50)
text ‹Set comprehensions›
syntax
"_Coll" :: "pttrn ⇒ bool ⇒ 'a set" ("(1{_./ _})")
translations
"{x. P}" ⇌ "CONST Collect (λx. P)"
syntax (ASCII)
"_Collect" :: "pttrn ⇒ 'a set ⇒ bool ⇒ 'a set" ("(1{(_/: _)./ _})")
syntax
"_Collect" :: "pttrn ⇒ 'a set ⇒ bool ⇒ 'a set" ("(1{(_/ ∈ _)./ _})")
translations
"{p:A. P}" ⇀ "CONST Collect (λp. p ∈ A ∧ P)"
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
text ‹
Simproc for pulling ‹x = t› in ‹{x. … ∧ x = t ∧ …}›
to the front (and similarly for ‹t = x›):
›
simproc_setup defined_Collect ("{x. P x ∧ Q x}") = ‹
K (Quantifier1.rearrange_Collect
(fn ctxt =>
resolve_tac ctxt @{thms Collect_cong} 1 THEN
resolve_tac ctxt @{thms iffI} 1 THEN
ALLGOALS
(EVERY' [REPEAT_DETERM o eresolve_tac ctxt @{thms conjE},
DEPTH_SOLVE_1 o (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms conjI})])))
›
lemmas CollectE = CollectD [elim_format]
lemma set_eqI:
assumes "⋀x. x ∈ A ⟷ x ∈ B"
shows "A = B"
proof -
from assms have "{x. x ∈ A} = {x. x ∈ B}"
by simp
then show ?thesis by simp
qed
lemma set_eq_iff: "A = B ⟷ (∀x. x ∈ A ⟷ x ∈ B)"
by (auto intro:set_eqI)
lemma Collect_eqI:
assumes "⋀x. P x = Q x"
shows "Collect P = Collect Q"
using assms by (auto intro: set_eqI)
text ‹Lifting of predicate class instances›
instantiation set :: (type) boolean_algebra
begin
definition less_eq_set
where "A ≤ B ⟷ (λx. member x A) ≤ (λx. member x B)"
definition less_set
where "A < B ⟷ (λx. member x A) < (λx. member x B)"
definition inf_set
where "A ⊓ B = Collect ((λx. member x A) ⊓ (λx. member x B))"
definition sup_set
where "A ⊔ B = Collect ((λx. member x A) ⊔ (λx. member x B))"
definition bot_set
where "⊥ = Collect ⊥"
definition top_set
where "⊤ = Collect ⊤"
definition uminus_set
where "- A = Collect (- (λx. member x A))"
definition minus_set
where "A - B = Collect ((λx. member x A) - (λx. member x B))"
instance
by standard
(simp_all add: less_eq_set_def less_set_def inf_set_def sup_set_def
bot_set_def top_set_def uminus_set_def minus_set_def
less_le_not_le sup_inf_distrib1 diff_eq set_eqI fun_eq_iff
del: inf_apply sup_apply bot_apply top_apply minus_apply uminus_apply)
end
text ‹Set enumerations›
abbreviation empty :: "'a set" ("{}")
where "{} ≡ bot"
definition insert :: "'a ⇒ 'a set ⇒ 'a set"
where insert_compr: "insert a B = {x. x = a ∨ x ∈ B}"
syntax
"_Finset" :: "args ⇒ 'a set" ("{(_)}")
translations
"{x, xs}" ⇌ "CONST insert x {xs}"
"{x}" ⇌ "CONST insert x {}"
subsection ‹Subsets and bounded quantifiers›
abbreviation subset :: "'a set ⇒ 'a set ⇒ bool"
where "subset ≡ less"
abbreviation subset_eq :: "'a set ⇒ 'a set ⇒ bool"
where "subset_eq ≡ less_eq"
notation
subset ("'(⊂')") and
subset ("(_/ ⊂ _)" [51, 51] 50) and
subset_eq ("'(⊆')") and
subset_eq ("(_/ ⊆ _)" [51, 51] 50)
abbreviation (input)
supset :: "'a set ⇒ 'a set ⇒ bool" where
"supset ≡ greater"
abbreviation (input)
supset_eq :: "'a set ⇒ 'a set ⇒ bool" where
"supset_eq ≡ greater_eq"
notation
supset ("'(⊃')") and
supset ("(_/ ⊃ _)" [51, 51] 50) and
supset_eq ("'(⊇')") and
supset_eq ("(_/ ⊇ _)" [51, 51] 50)
notation (ASCII output)
subset ("'(<')") and
subset ("(_/ < _)" [51, 51] 50) and
subset_eq ("'(<=')") and
subset_eq ("(_/ <= _)" [51, 51] 50)
definition Ball :: "'a set ⇒ ('a ⇒ bool) ⇒ bool"
where "Ball A P ⟷ (∀x. x ∈ A ⟶ P x)"
definition Bex :: "'a set ⇒ ('a ⇒ bool) ⇒ bool"
where "Bex A P ⟷ (∃x. x ∈ A ∧ P x)"
syntax (ASCII)
"_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 (input)
"_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
"_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)
"_Bleast" :: "id ⇒ 'a set ⇒ bool ⇒ 'a" ("(3LEAST(_/∈_)./ _)" [0, 0, 10] 10)
translations
"∀x∈A. P" ⇌ "CONST Ball A (λx. P)"
"∃x∈A. P" ⇌ "CONST Bex A (λx. P)"
"∃!x∈A. P" ⇀ "∃!x. x ∈ A ∧ P"
"LEAST x:A. P" ⇀ "LEAST x. x ∈ A ∧ P"
syntax (ASCII 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
"_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)
translations
"∀A⊂B. P" ⇀ "∀A. A ⊂ B ⟶ P"
"∃A⊂B. P" ⇀ "∃A. A ⊂ B ∧ P"
"∀A⊆B. P" ⇀ "∀A. A ⊆ B ⟶ P"
"∃A⊆B. P" ⇀ "∃A. A ⊆ B ∧ P"
"∃!A⊆B. P" ⇀ "∃!A. A ⊆ B ∧ P"
print_translation ‹
let
val All_binder = Mixfix.binder_name \<^const_syntax>‹All›;
val Ex_binder = Mixfix.binder_name \<^const_syntax>‹Ex›;
val impl = \<^const_syntax>‹HOL.implies›;
val conj = \<^const_syntax>‹HOL.conj›;
val sbset = \<^const_syntax>‹subset›;
val sbset_eq = \<^const_syntax>‹subset_eq›;
val trans =
[((All_binder, impl, sbset), \<^syntax_const>‹_setlessAll›),
((All_binder, impl, sbset_eq), \<^syntax_const>‹_setleAll›),
((Ex_binder, conj, sbset), \<^syntax_const>‹_setlessEx›),
((Ex_binder, conj, sbset_eq), \<^syntax_const>‹_setleEx›)];
fun mk v (v', T) c n P =
if v = v' andalso not (Term.exists_subterm (fn Free (x, _) => x = v | _ => false) n)
then Syntax.const c $ Syntax_Trans.mark_bound_body (v', T) $ n $ P
else raise Match;
fun tr' q = (q, fn _ =>
(fn [Const (\<^syntax_const>‹_bound›, _) $ Free (v, Type (\<^type_name>‹set›, _)),
Const (c, _) $
(Const (d, _) $ (Const (\<^syntax_const>‹_bound›, _) $ Free (v', T)) $ n) $ P] =>
(case AList.lookup (=) trans (q, c, d) of
NONE => raise Match
| SOME l => mk v (v', T) l n P)
| _ => raise Match));
in
[tr' All_binder, tr' Ex_binder]
end
›
text ‹
┉
Translate between ‹{e | x1…xn. P}› and ‹{u. ∃x1…xn. u = e ∧ P}›;
‹{y. ∃x1…xn. y = e ∧ P}› is only translated if ‹[0..n] ⊆ bvs e›.
›
syntax
"_Setcompr" :: "'a ⇒ idts ⇒ bool ⇒ 'a set" ("(1{_ |/_./ _})")
parse_translation ‹
let
val ex_tr = snd (Syntax_Trans.mk_binder_tr ("EX ", \<^const_syntax>‹Ex›));
fun nvars (Const (\<^syntax_const>‹_idts›, _) $ _ $ idts) = nvars idts + 1
| nvars _ = 1;
fun setcompr_tr ctxt [e, idts, b] =
let
val eq = Syntax.const \<^const_syntax>‹HOL.eq› $ Bound (nvars idts) $ e;
val P = Syntax.const \<^const_syntax>‹HOL.conj› $ eq $ b;
val exP = ex_tr ctxt [idts, P];
in Syntax.const \<^const_syntax>‹Collect› $ absdummy dummyT exP end;
in [(\<^syntax_const>‹_Setcompr›, setcompr_tr)] end
›
print_translation ‹
[Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>‹Ball› \<^syntax_const>‹_Ball›,
Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>‹Bex› \<^syntax_const>‹_Bex›]
›
print_translation ‹
let
val ex_tr' = snd (Syntax_Trans.mk_binder_tr' (\<^const_syntax>‹Ex›, "DUMMY"));
fun setcompr_tr' ctxt [Abs (abs as (_, _, P))] =
let
fun check (Const (\<^const_syntax>‹Ex›, _) $ Abs (_, _, P), n) = check (P, n + 1)
| check (Const (\<^const_syntax>‹HOL.conj›, _) $
(Const (\<^const_syntax>‹HOL.eq›, _) $ Bound m $ e) $ P, n) =
n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso
subset (=) (0 upto (n - 1), add_loose_bnos (e, 0, []))
| check _ = false;
fun tr' (_ $ abs) =
let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr' ctxt [abs]
in Syntax.const \<^syntax_const>‹_Setcompr› $ e $ idts $ Q end;
in
if check (P, 0) then tr' P
else
let
val (x as _ $ Free(xN, _), t) = Syntax_Trans.atomic_abs_tr' abs;
val M = Syntax.const \<^syntax_const>‹_Coll› $ x $ t;
in
case t of
Const (\<^const_syntax>‹HOL.conj›, _) $
(Const (\<^const_syntax>‹Set.member›, _) $
(Const (\<^syntax_const>‹_bound›, _) $ Free (yN, _)) $ A) $ P =>
if xN = yN then Syntax.const \<^syntax_const>‹_Collect› $ x $ A $ P else M
| _ => M
end
end;
in [(\<^const_syntax>‹Collect›, setcompr_tr')] end
›
simproc_setup defined_Bex ("∃x∈A. P x ∧ Q x") = ‹
K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def}))
›
simproc_setup defined_All ("∀x∈A. P x ⟶ Q x") = ‹
K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms Ball_def}))
›
lemma ballI [intro!]: "(⋀x. x ∈ A ⟹ P x) ⟹ ∀x∈A. P x"
by (simp add: Ball_def)
lemmas strip = impI allI ballI
lemma bspec [dest?]: "∀x∈A. P x ⟹ x ∈ A ⟹ P x"
by (simp add: Ball_def)
text ‹Gives better instantiation for bound:›
setup ‹
map_theory_claset (fn ctxt =>
ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt'))
›
ML ‹
structure Simpdata =
struct
open Simpdata;
val mksimps_pairs = [(\<^const_name>‹Ball›, @{thms bspec})] @ mksimps_pairs;
end;
open Simpdata;
›
declaration ‹fn _ => Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))›
lemma ballE [elim]: "∀x∈A. P x ⟹ (P x ⟹ Q) ⟹ (x ∉ A ⟹ Q) ⟹ Q"
unfolding Ball_def by blast
lemma bexI [intro]: "P x ⟹ x ∈ A ⟹ ∃x∈A. P x"
unfolding Bex_def by blast
lemma rev_bexI [intro?]: "x ∈ A ⟹ P x ⟹ ∃x∈A. P x"
unfolding Bex_def by blast
lemma bexCI: "(∀x∈A. ¬ P x ⟹ P a) ⟹ a ∈ A ⟹ ∃x∈A. P x"
unfolding Bex_def by blast
lemma bexE [elim!]: "∃x∈A. P x ⟹ (⋀x. x ∈ A ⟹ P x ⟹ Q) ⟹ Q"
unfolding Bex_def by blast
lemma ball_triv [simp]: "(∀x∈A. P) ⟷ ((∃x. x ∈ A) ⟶ P)"
by (simp add: Ball_def)
lemma bex_triv [simp]: "(∃x∈A. P) ⟷ ((∃x. x ∈ A) ∧ P)"
by (simp add: Bex_def)
lemma bex_triv_one_point1 [simp]: "(∃x∈A. x = a) ⟷ a ∈ A"
by blast
lemma bex_triv_one_point2 [simp]: "(∃x∈A. a = x) ⟷ a ∈ A"
by blast
lemma bex_one_point1 [simp]: "(∃x∈A. x = a ∧ P x) ⟷ a ∈ A ∧ P a"
by blast
lemma bex_one_point2 [simp]: "(∃x∈A. a = x ∧ P x) ⟷ a ∈ A ∧ P a"
by blast
lemma ball_one_point1 [simp]: "(∀x∈A. x = a ⟶ P x) ⟷ (a ∈ A ⟶ P a)"
by blast
lemma ball_one_point2 [simp]: "(∀x∈A. a = x ⟶ P x) ⟷ (a ∈ A ⟶ P a)"
by blast
lemma ball_conj_distrib: "(∀x∈A. P x ∧ Q x) ⟷ (∀x∈A. P x) ∧ (∀x∈A. Q x)"
by blast
lemma bex_disj_distrib: "(∃x∈A. P x ∨ Q x) ⟷ (∃x∈A. P x) ∨ (∃x∈A. Q x)"
by blast
text ‹Congruence rules›
lemma ball_cong:
"⟦ A = B; ⋀x. x ∈ B ⟹ P x ⟷ Q x ⟧ ⟹
(∀x∈A. P x) ⟷ (∀x∈B. Q x)"
by (simp add: Ball_def)
lemma ball_cong_simp [cong]:
"⟦ A = B; ⋀x. x ∈ B =simp=> P x ⟷ Q x ⟧ ⟹
(∀x∈A. P x) ⟷ (∀x∈B. Q x)"
by (simp add: simp_implies_def Ball_def)
lemma bex_cong:
"⟦ A = B; ⋀x. x ∈ B ⟹ P x ⟷ Q x ⟧ ⟹
(∃x∈A. P x) ⟷ (∃x∈B. Q x)"
by (simp add: Bex_def cong: conj_cong)
lemma bex_cong_simp [cong]:
"⟦ A = B; ⋀x. x ∈ B =simp=> P x ⟷ Q x ⟧ ⟹
(∃x∈A. P x) ⟷ (∃x∈B. Q x)"
by (simp add: simp_implies_def Bex_def cong: conj_cong)
lemma bex1_def: "(∃!x∈X. P x) ⟷ (∃x∈X. P x) ∧ (∀x∈X. ∀y∈X. P x ⟶ P y ⟶ x = y)"
by auto
subsection ‹Basic operations›
subsubsection ‹Subsets›
lemma subsetI [intro!]: "(⋀x. x ∈ A ⟹ x ∈ B) ⟹ A ⊆ B"
by (simp add: less_eq_set_def le_fun_def)
text ‹
┉
Map the type ‹'a set ⇒ anything› to just ‹'a›; for overloading constants
whose first argument has type ‹'a set›.
›
lemma subsetD [elim, intro?]: "A ⊆ B ⟹ c ∈ A ⟹ c ∈ B"
by (simp add: less_eq_set_def le_fun_def)
lemma rev_subsetD [intro?,no_atp]: "c ∈ A ⟹ A ⊆ B ⟹ c ∈ B"
by (rule subsetD)
lemma subsetCE [elim,no_atp]: "A ⊆ B ⟹ (c ∉ A ⟹ P) ⟹ (c ∈ B ⟹ P) ⟹ P"
by (auto simp add: less_eq_set_def le_fun_def)
lemma subset_eq: "A ⊆ B ⟷ (∀x∈A. x ∈ B)"
by blast
lemma contra_subsetD [no_atp]: "A ⊆ B ⟹ c ∉ B ⟹ c ∉ A"
by blast
lemma subset_refl: "A ⊆ A"
by (fact order_refl)
lemma subset_trans: "A ⊆ B ⟹ B ⊆ C ⟹ A ⊆ C"
by (fact order_trans)
lemma subset_not_subset_eq [code]: "A ⊂ B ⟷ A ⊆ B ∧ ¬ B ⊆ A"
by (fact less_le_not_le)
lemma eq_mem_trans: "a = b ⟹ b ∈ A ⟹ a ∈ A"
by simp
lemmas basic_trans_rules [trans] =
order_trans_rules rev_subsetD subsetD eq_mem_trans
subsubsection ‹Equality›
lemma subset_antisym [intro!]: "A ⊆ B ⟹ B ⊆ A ⟹ A = B"
by (iprover intro: set_eqI subsetD)
text ‹┉ Equality rules from ZF set theory -- are they appropriate here?›
lemma equalityD1: "A = B ⟹ A ⊆ B"
by simp
lemma equalityD2: "A = B ⟹ B ⊆ A"
by simp
text ‹
┉
Be careful when adding this to the claset as ‹subset_empty› is in the
simpset: \<^prop>‹A = {}› goes to \<^prop>‹{} ⊆ A› and \<^prop>‹A ⊆ {}›
and then back to \<^prop>‹A = {}›!
›
lemma equalityE: "A = B ⟹ (A ⊆ B ⟹ B ⊆ A ⟹ P) ⟹ P"
by simp
lemma equalityCE [elim]: "A = B ⟹ (c ∈ A ⟹ c ∈ B ⟹ P) ⟹ (c ∉ A ⟹ c ∉ 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 ‹The empty set›
lemma empty_def: "{} = {x. False}"
by (simp add: bot_set_def bot_fun_def)
lemma empty_iff [simp]: "c ∈ {} ⟷ False"
by (simp add: empty_def)
lemma emptyE [elim!]: "a ∈ {} ⟹ P"
by simp
lemma empty_subsetI [iff]: "{} ⊆ A"
by blast
lemma equals0I: "(⋀y. y ∈ A ⟹ False) ⟹ A = {}"
by blast
lemma equals0D: "A = {} ⟹ a ∉ A"
by blast
lemma ball_empty [simp]: "Ball {} P ⟷ True"
by (simp add: Ball_def)
lemma bex_empty [simp]: "Bex {} P ⟷ False"
by (simp add: Bex_def)
subsubsection ‹The universal set -- UNIV›
abbreviation UNIV :: "'a set"
where "UNIV ≡ top"
lemma UNIV_def: "UNIV = {x. True}"
by (simp add: top_set_def top_fun_def)
lemma UNIV_I [simp]: "x ∈ UNIV"
by (simp add: UNIV_def)
declare UNIV_I [intro]
lemma UNIV_witness [intro?]: "∃x. x ∈ UNIV"
by simp
lemma subset_UNIV: "A ⊆ UNIV"
by (fact top_greatest)
text ‹
┉
Eta-contracting these two rules (to remove ‹P›) causes them
to be ignored because of their interaction with congruence rules.
›
lemma ball_UNIV [simp]: "Ball UNIV P ⟷ All P"
by (simp add: Ball_def)
lemma bex_UNIV [simp]: "Bex UNIV P ⟷ Ex P"
by (simp add: Bex_def)
lemma UNIV_eq_I: "(⋀x. x ∈ A) ⟹ UNIV = A"
by auto
lemma UNIV_not_empty [iff]: "UNIV ≠ {}"
by (blast elim: equalityE)
lemma empty_not_UNIV[simp]: "{} ≠ UNIV"
by blast
subsubsection ‹The Powerset operator -- Pow›
definition Pow :: "'a set ⇒ 'a set set"
where Pow_def: "Pow A = {B. B ⊆ A}"
lemma Pow_iff [iff]: "A ∈ Pow B ⟷ A ⊆ B"
by (simp add: Pow_def)
lemma PowI: "A ⊆ B ⟹ A ∈ Pow B"
by (simp add: Pow_def)
lemma PowD: "A ∈ Pow B ⟹ A ⊆ B"
by (simp add: Pow_def)
lemma Pow_bottom: "{} ∈ Pow B"
by simp
lemma Pow_top: "A ∈ Pow A"
by simp
lemma Pow_not_empty: "Pow A ≠ {}"
using Pow_top by blast
subsubsection ‹Set complement›
lemma Compl_iff [simp]: "c ∈ - A ⟷ c ∉ A"
by (simp add: fun_Compl_def uminus_set_def)
lemma ComplI [intro!]: "(c ∈ A ⟹ False) ⟹ c ∈ - A"
by (simp add: fun_Compl_def uminus_set_def) blast
text ‹
┉
This form, with negated conclusion, works well with the Classical prover.
Negated assumptions behave like formulae on the right side of the
notional turnstile \dots
›
lemma ComplD [dest!]: "c ∈ - A ⟹ c ∉ A"
by simp
lemmas ComplE = ComplD [elim_format]
lemma Compl_eq: "- A = {x. ¬ x ∈ A}"
by blast
subsubsection ‹Binary intersection›
abbreviation inter :: "'a set ⇒ 'a set ⇒ 'a set" (infixl "∩" 70)
where "(∩) ≡ inf"
notation (ASCII)
inter (infixl "Int" 70)
lemma Int_def: "A ∩ B = {x. x ∈ A ∧ x ∈ B}"
by (simp add: inf_set_def inf_fun_def)
lemma Int_iff [simp]: "c ∈ A ∩ B ⟷ c ∈ A ∧ c ∈ B"
unfolding Int_def by blast
lemma IntI [intro!]: "c ∈ A ⟹ c ∈ B ⟹ c ∈ A ∩ B"
by simp
lemma IntD1: "c ∈ A ∩ B ⟹ c ∈ A"
by simp
lemma IntD2: "c ∈ A ∩ B ⟹ c ∈ B"
by simp
lemma IntE [elim!]: "c ∈ A ∩ B ⟹ (c ∈ A ⟹ c ∈ B ⟹ P) ⟹ P"
by simp
subsubsection ‹Binary union›
abbreviation union :: "'a set ⇒ 'a set ⇒ 'a set" (infixl "∪" 65)
where "union ≡ sup"
notation (ASCII)
union (infixl "Un" 65)
lemma Un_def: "A ∪ B = {x. x ∈ A ∨ x ∈ B}"
by (simp add: sup_set_def sup_fun_def)
lemma Un_iff [simp]: "c ∈ A ∪ B ⟷ c ∈ A ∨ c ∈ B"
unfolding Un_def by blast
lemma UnI1 [elim?]: "c ∈ A ⟹ c ∈ A ∪ B"
by simp
lemma UnI2 [elim?]: "c ∈ B ⟹ c ∈ A ∪ B"
by simp
text ‹┉ Classical introduction rule: no commitment to ‹A› vs. ‹B›.›
lemma UnCI [intro!]: "(c ∉ B ⟹ c ∈ A) ⟹ c ∈ A ∪ B"
by auto
lemma UnE [elim!]: "c ∈ A ∪ B ⟹ (c ∈ A ⟹ P) ⟹ (c ∈ B ⟹ P) ⟹ P"
unfolding Un_def by blast
lemma insert_def: "insert a B = {x. x = a} ∪ B"
by (simp add: insert_compr Un_def)
subsubsection ‹Set difference›
lemma Diff_iff [simp]: "c ∈ A - B ⟷ c ∈ A ∧ c ∉ B"
by (simp add: minus_set_def fun_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
abbreviation sym_diff :: "'a set ⇒ 'a set ⇒ 'a set" where
"sym_diff A B ≡ ((A - B) ∪ (B-A))"
subsubsection ‹Augmenting a set -- \<^const>‹insert››
lemma insert_iff [simp]: "a ∈ insert b A ⟷ a = b ∨ a ∈ A"
unfolding insert_def by 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"
unfolding insert_def by blast
lemma insertCI [intro!]: "(a ∉ B ⟹ a = b) ⟹ a ∈ insert b B"
by auto
lemma subset_insert_iff: "A ⊆ insert x B ⟷ (if x ∈ A then A - {x} ⊆ B else A ⊆ B)"
by auto
lemma set_insert:
assumes "x ∈ A"
obtains B where "A = insert x B" and "x ∉ B"
proof
show "A = insert x (A - {x})" using assms by blast
show "x ∉ A - {x}" by blast
qed
lemma insert_ident: "x ∉ A ⟹ x ∉ B ⟹ insert x A = insert x B ⟷ A = B"
by auto
lemma insert_eq_iff:
assumes "a ∉ A" "b ∉ B"
shows "insert a A = insert b B ⟷
(if a = b then A = B else ∃C. A = insert b C ∧ b ∉ C ∧ B = insert a C ∧ a ∉ C)"
(is "?L ⟷ ?R")
proof
show ?R if ?L
proof (cases "a = b")
case True
with assms ‹?L› show ?R
by (simp add: insert_ident)
next
case False
let ?C = "A - {b}"
have "A = insert b ?C ∧ b ∉ ?C ∧ B = insert a ?C ∧ a ∉ ?C"
using assms ‹?L› ‹a ≠ b› by auto
then show ?R using ‹a ≠ b› by auto
qed
show ?L if ?R
using that by (auto split: if_splits)
qed
lemma insert_UNIV[simp]: "insert x UNIV = UNIV"
by auto
subsubsection ‹Singletons, using insert›
lemma singletonI [intro!]: "a ∈ {a}"
by (rule insertI1)
lemma singletonD [dest!]: "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]: "{b} = insert a A ⟷ a = b ∧ A ⊆ {b}"
by blast
lemma singleton_insert_inj_eq' [iff]: "insert a A = {b} ⟷ a = b ∧ A ⊆ {b}"
by blast
lemma subset_singletonD: "A ⊆ {x} ⟹ A = {} ∨ A = {x}"
by fast
lemma subset_singleton_iff: "X ⊆ {a} ⟷ X = {} ∨ X = {a}"
by blast
lemma subset_singleton_iff_Uniq: "(∃a. A ⊆ {a}) ⟷ (∃⇩≤⇩1x. x ∈ A)"
unfolding Uniq_def by blast
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} ⊆ B ⟹ A ⊆ insert x B"
by blast
lemma subset_Diff_insert: "A ⊆ B - insert x C ⟷ A ⊆ B - C ∧ x ∉ A"
by blast
lemma doubleton_eq_iff: "{a, b} = {c, d} ⟷ a = c ∧ b = d ∨ a = d ∧ b = c"
by (blast elim: equalityE)
lemma Un_singleton_iff: "A ∪ B = {x} ⟷ A = {} ∧ B = {x} ∨ A = {x} ∧ B = {} ∨ A = {x} ∧ B = {x}"
by auto
lemma singleton_Un_iff: "{x} = A ∪ B ⟷ A = {} ∧ B = {x} ∨ A = {x} ∧ B = {} ∨ A = {x} ∧ B = {x}"
by auto
subsubsection ‹Image of a set under a function›
text ‹Frequently ‹b› does not have the syntactic form of ‹f x›.›
definition image :: "('a ⇒ 'b) ⇒ 'a set ⇒ 'b set" (infixr "`" 90)
where "f ` A = {y. ∃x∈A. y = f x}"
lemma image_eqI [simp, intro]: "b = f x ⟹ x ∈ A ⟹ b ∈ f ` A"
unfolding image_def by blast
lemma imageI: "x ∈ A ⟹ f x ∈ f ` A"
by (rule image_eqI) (rule refl)
lemma rev_image_eqI: "x ∈ A ⟹ b = f x ⟹ b ∈ f ` A"
by (rule image_eqI)
lemma imageE [elim!]:
assumes "b ∈ (λx. f x) ` A"
obtains x where "b = f x" and "x ∈ A"
using assms unfolding image_def by blast
lemma Compr_image_eq: "{x ∈ f ` A. P x} = f ` {x ∈ A. P (f x)}"
by auto
lemma image_Un: "f ` (A ∪ B) = f ` A ∪ f ` B"
by blast
lemma image_iff: "z ∈ f ` A ⟷ (∃x∈A. z = f x)"
by blast
lemma image_subsetI: "(⋀x. x ∈ A ⟹ f x ∈ B) ⟹ f ` A ⊆ B"
by blast
lemma image_subset_iff: "f ` A ⊆ B ⟷ (∀x∈A. f x ∈ B)"
by blast
lemma subset_imageE:
assumes "B ⊆ f ` A"
obtains C where "C ⊆ A" and "B = f ` C"
proof -
from assms have "B = f ` {a ∈ A. f a ∈ B}" by fast
moreover have "{a ∈ A. f a ∈ B} ⊆ A" by blast
ultimately show thesis by (blast intro: that)
qed
lemma subset_image_iff: "B ⊆ f ` A ⟷ (∃AA⊆A. B = f ` AA)"
by (blast elim: subset_imageE)
lemma image_ident [simp]: "(λx. x) ` Y = Y"
by blast
lemma image_empty [simp]: "f ` {} = {}"
by blast
lemma image_insert [simp]: "f ` insert a B = insert (f a) (f ` B)"
by blast
lemma image_constant: "x ∈ A ⟹ (λx. c) ` A = {c}"
by auto
lemma image_constant_conv: "(λx. c) ` A = (if A = {} then {} else {c})"
by auto
lemma image_image: "f ` (g ` A) = (λx. f (g x)) ` A"
by blast
lemma insert_image [simp]: "x ∈ A ⟹ insert (f x) (f ` A) = f ` A"
by blast
lemma image_is_empty [iff]: "f ` A = {} ⟷ A = {}"
by blast
lemma empty_is_image [iff]: "{} = f ` A ⟷ A = {}"
by blast
lemma image_Collect: "f ` {x. P x} = {f x | x. P x}"
by blast
lemma if_image_distrib [simp]:
"(λx. if P x then f x else g x) ` S = f ` (S ∩ {x. P x}) ∪ g ` (S ∩ {x. ¬ P x})"
by auto
lemma image_cong:
"f ` M = g ` N" if "M = N" "⋀x. x ∈ N ⟹ f x = g x"
using that by (simp add: image_def)
lemma image_cong_simp [cong]:
"f ` M = g ` N" if "M = N" "⋀x. x ∈ N =simp=> f x = g x"
using that image_cong [of M N f g] by (simp add: simp_implies_def)
lemma image_Int_subset: "f ` (A ∩ B) ⊆ f ` A ∩ f ` B"
by blast
lemma image_diff_subset: "f ` A - f ` B ⊆ f ` (A - B)"
by blast
lemma Setcompr_eq_image: "{f x |x. x ∈ A} = f ` A"
by blast
lemma setcompr_eq_image: "{f x |x. P x} = f ` {x. P x}"
by auto
lemma ball_imageD: "∀x∈f ` A. P x ⟹ ∀x∈A. P (f x)"
by simp
lemma bex_imageD: "∃x∈f ` A. P x ⟹ ∃x∈A. P (f x)"
by auto
lemma image_add_0 [simp]: "(+) (0::'a::comm_monoid_add) ` S = S"
by auto
theorem Cantors_theorem: "∄f. f ` A = Pow A"
proof
assume "∃f. f ` A = Pow A"
then obtain f where f: "f ` A = Pow A" ..
let ?X = "{a ∈ A. a ∉ f a}"
have "?X ∈ Pow A" by blast
then have "?X ∈ f ` A" by (simp only: f)
then obtain x where "x ∈ A" and "f x = ?X" by blast
then show False by blast
qed
text ‹┉ Range of a function -- just an abbreviation for image!›
abbreviation range :: "('a ⇒ 'b) ⇒ 'b set"
where "range f ≡ f ` UNIV"
lemma range_eqI: "b = f x ⟹ b ∈ range f"
by simp
lemma rangeI: "f x ∈ range f"
by simp
lemma rangeE [elim?]: "b ∈ range (λx. f x) ⟹ (⋀x. b = f x ⟹ P) ⟹ P"
by (rule imageE)
lemma range_subsetD: "range f ⊆ B ⟹ f i ∈ B"
by blast
lemma full_SetCompr_eq: "{u. ∃x. u = f x} = range f"
by auto
lemma range_composition: "range (λx. f (g x)) = f ` range g"
by auto
lemma range_constant [simp]: "range (λ_. x) = {x}"
by (simp add: image_constant)
lemma range_eq_singletonD: "range f = {a} ⟹ f x = a"
by auto
subsubsection ‹Some rules with ‹if››
text ‹Elimination of ‹{x. … ∧ x = t ∧ …}›.›
lemma Collect_conv_if: "{x. x = a ∧ P x} = (if P a then {a} else {})"
by auto
lemma Collect_conv_if2: "{x. a = x ∧ P x} = (if P a then {a} else {})"
by auto
text ‹
Rewrite rules for boolean case-splitting: faster than ‹if_split [split]›.
›
lemma if_split_eq1: "(if Q then x else y) = b ⟷ (Q ⟶ x = b) ∧ (¬ Q ⟶ y = b)"
by (rule if_split)
lemma if_split_eq2: "a = (if Q then x else y) ⟷ (Q ⟶ a = x) ∧ (¬ Q ⟶ a = y)"
by (rule if_split)
text ‹
Split ifs on either side of the membership relation.
Not for ‹[simp]› -- can cause goals to blow up!
›
lemma if_split_mem1: "(if Q then x else y) ∈ b ⟷ (Q ⟶ x ∈ b) ∧ (¬ Q ⟶ y ∈ b)"
by (rule if_split)
lemma if_split_mem2: "(a ∈ (if Q then x else y)) ⟷ (Q ⟶ a ∈ x) ∧ (¬ Q ⟶ a ∈ y)"
by (rule if_split [where P = "λS. a ∈ S"])
lemmas split_ifs = if_bool_eq_conj if_split_eq1 if_split_eq2 if_split_mem1 if_split_mem2
subsection ‹Further operations and lemmas›
subsubsection ‹The ``proper subset'' relation›
lemma psubsetI [intro!]: "A ⊆ B ⟹ A ≠ B ⟹ A ⊂ B"
unfolding less_le by blast
lemma psubsetE [elim!]: "A ⊂ B ⟹ (A ⊆ B ⟹ ¬ B ⊆ A ⟹ R) ⟹ R"
unfolding less_le by blast
lemma psubset_insert_iff:
"A ⊂ insert x B ⟷ (if x ∈ B then A ⊂ B else if x ∈ A then A - {x} ⊂ B else A ⊆ B)"
by (auto simp add: less_le subset_insert_iff)
lemma psubset_eq: "A ⊂ B ⟷ A ⊆ B ∧ A ≠ B"
by (simp only: less_le)
lemma psubset_imp_subset: "A ⊂ B ⟹ A ⊆ B"
by (simp add: psubset_eq)
lemma psubset_trans: "A ⊂ B ⟹ B ⊂ C ⟹ A ⊂ C"
unfolding less_le by (auto dest: subset_antisym)
lemma psubsetD: "A ⊂ B ⟹ c ∈ A ⟹ c ∈ B"
unfolding less_le by (auto dest: subsetD)
lemma psubset_subset_trans: "A ⊂ B ⟹ B ⊆ C ⟹ A ⊂ C"
by (auto simp add: psubset_eq)
lemma subset_psubset_trans: "A ⊆ B ⟹ B ⊂ C ⟹ A ⊂ C"
by (auto simp add: psubset_eq)
lemma psubset_imp_ex_mem: "A ⊂ B ⟹ ∃b. b ∈ B - A"
unfolding less_le by blast
lemma atomize_ball: "(⋀x. x ∈ A ⟹ P x) ≡ Trueprop (∀x∈A. P x)"
by (simp only: Ball_def atomize_all atomize_imp)
lemmas [symmetric, rulify] = atomize_ball
and [symmetric, defn] = atomize_ball
lemma image_Pow_mono: "f ` A ⊆ B ⟹ image f ` Pow A ⊆ Pow B"
by blast
lemma image_Pow_surj: "f ` A = B ⟹ image f ` Pow A = Pow B"
by (blast elim: subset_imageE)
subsubsection ‹Derived rules involving subsets.›
text ‹‹insert›.›
lemma subset_insertI: "B ⊆ insert a B"
by (rule subsetI) (erule insertI2)
lemma subset_insertI2: "A ⊆ B ⟹ A ⊆ insert b B"
by blast
lemma subset_insert: "x ∉ A ⟹ A ⊆ insert x B ⟷ A ⊆ B"
by blast
text ‹┉ Finite Union -- the least upper bound of two sets.›
lemma Un_upper1: "A ⊆ A ∪ B"
by (fact sup_ge1)
lemma Un_upper2: "B ⊆ A ∪ B"
by (fact sup_ge2)
lemma Un_least: "A ⊆ C ⟹ B ⊆ C ⟹ A ∪ B ⊆ C"
by (fact sup_least)
text ‹┉ Finite Intersection -- the greatest lower bound of two sets.›
lemma Int_lower1: "A ∩ B ⊆ A"
by (fact inf_le1)
lemma Int_lower2: "A ∩ B ⊆ B"
by (fact inf_le2)
lemma Int_greatest: "C ⊆ A ⟹ C ⊆ B ⟹ C ⊆ A ∩ B"
by (fact inf_greatest)
text ‹┉ Set difference.›
lemma Diff_subset[simp]: "A - B ⊆ A"
by blast
lemma Diff_subset_conv: "A - B ⊆ C ⟷ A ⊆ B ∪ C"
by blast
subsubsection ‹Equalities involving union, intersection, inclusion, etc.›
text ‹‹{}›.›
lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})"
by auto
lemma subset_empty [simp]: "A ⊆ {} ⟷ A = {}"
by (fact bot_unique)
lemma not_psubset_empty [iff]: "¬ (A < {})"
by (fact not_less_bot)
lemma Collect_subset [simp]: "{x∈A. P x} ⊆ A" by auto
lemma Collect_empty_eq [simp]: "Collect P = {} ⟷ (∀x. ¬ P x)"
by blast
lemma empty_Collect_eq [simp]: "{} = Collect P ⟷ (∀x. ¬ P x)"
by blast
lemma Collect_neg_eq: "{x. ¬ P x} = - {x. P x}"
by blast
lemma Collect_disj_eq: "{x. P x ∨ Q x} = {x. P x} ∪ {x. Q x}"
by blast
lemma Collect_imp_eq: "{x. P x ⟶ Q x} = - {x. P x} ∪ {x. Q x}"
by blast
lemma Collect_conj_eq: "{x. P x ∧ Q x} = {x. P x} ∩ {x. Q x}"
by blast
lemma Collect_mono_iff: "Collect P ⊆ Collect Q ⟷ (∀x. P x ⟶ Q x)"
by blast
text ‹┉ ‹insert›.›
lemma insert_is_Un: "insert a A = {a} ∪ A"
by blast
lemma insert_not_empty [simp]: "insert a A ≠ {}"
and empty_not_insert [simp]: "{} ≠ insert a A"
by blast+
lemma insert_absorb: "a ∈ A ⟹ insert a A = A"
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 ⊆ B ⟷ x ∈ B ∧ A ⊆ B"
by blast
lemma mk_disjoint_insert: "a ∈ A ⟹ ∃B. A = insert a B ∧ a ∉ B"
by (rule exI [where x = "A - {a}"]) blast
lemma insert_Collect: "insert a (Collect P) = {u. u ≠ a ⟶ P u}"
by auto
lemma insert_inter_insert [simp]: "insert a A ∩ insert a B = insert a (A ∩ B)"
by blast
lemma insert_disjoint [simp]:
"insert a A ∩ B = {} ⟷ a ∉ B ∧ A ∩ B = {}"
"{} = insert a A ∩ B ⟷ a ∉ B ∧ {} = A ∩ B"
by auto
lemma disjoint_insert [simp]:
"B ∩ insert a A = {} ⟷ a ∉ B ∧ B ∩ A = {}"
"{} = A ∩ insert b B ⟷ b ∉ A ∧ {} = A ∩ B"
by auto
text ‹┉ ‹Int››
lemma Int_absorb: "A ∩ A = A"
by (fact inf_idem)
lemma Int_left_absorb: "A ∩ (A ∩ B) = A ∩ B"
by (fact inf_left_idem)
lemma Int_commute: "A ∩ B = B ∩ A"
by (fact inf_commute)
lemma Int_left_commute: "A ∩ (B ∩ C) = B ∩ (A ∩ C)"
by (fact inf_left_commute)
lemma Int_assoc: "(A ∩ B) ∩ C = A ∩ (B ∩ C)"
by (fact inf_assoc)
lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute
lemma Int_absorb1: "B ⊆ A ⟹ A ∩ B = B"
by (fact inf_absorb2)
lemma Int_absorb2: "A ⊆ B ⟹ A ∩ B = A"
by (fact inf_absorb1)
lemma Int_empty_left: "{} ∩ B = {}"
by (fact inf_bot_left)
lemma Int_empty_right: "A ∩ {} = {}"
by (fact inf_bot_right)
lemma disjoint_eq_subset_Compl: "A ∩ B = {} ⟷ A ⊆ - B"
by blast
lemma disjoint_iff: "A ∩ B = {} ⟷ (∀x. x∈A ⟶ x ∉ B)"
by blast
lemma disjoint_iff_not_equal: "A ∩ B = {} ⟷ (∀x∈A. ∀y∈B. x ≠ y)"
by blast
lemma Int_UNIV_left: "UNIV ∩ B = B"
by (fact inf_top_left)
lemma Int_UNIV_right: "A ∩ UNIV = A"
by (fact inf_top_right)
lemma Int_Un_distrib: "A ∩ (B ∪ C) = (A ∩ B) ∪ (A ∩ C)"
by (fact inf_sup_distrib1)
lemma Int_Un_distrib2: "(B ∪ C) ∩ A = (B ∩ A) ∪ (C ∩ A)"
by (fact inf_sup_distrib2)
lemma Int_UNIV: "A ∩ B = UNIV ⟷ A = UNIV ∧ B = UNIV"
by (fact inf_eq_top_iff)
lemma Int_subset_iff: "C ⊆ A ∩ B ⟷ C ⊆ A ∧ C ⊆ B"
by (fact le_inf_iff)
lemma Int_Collect: "x ∈ A ∩ {x. P x} ⟷ x ∈ A ∧ P x"
by blast
text ‹┉ ‹Un›.›
lemma Un_absorb: "A ∪ A = A"
by (fact sup_idem)
lemma Un_left_absorb: "A ∪ (A ∪ B) = A ∪ B"
by (fact sup_left_idem)
lemma Un_commute: "A ∪ B = B ∪ A"
by (fact sup_commute)
lemma Un_left_commute: "A ∪ (B ∪ C) = B ∪ (A ∪ C)"
by (fact sup_left_commute)
lemma Un_assoc: "(A ∪ B) ∪ C = A ∪ (B ∪ C)"
by (fact sup_assoc)
lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute
lemma Un_absorb1: "A ⊆ B ⟹ A ∪ B = B"
by (fact sup_absorb2)
lemma Un_absorb2: "B ⊆ A ⟹ A ∪ B = A"
by (fact sup_absorb1)
lemma Un_empty_left: "{} ∪ B = B"
by (fact sup_bot_left)
lemma Un_empty_right: "A ∪ {} = A"
by (fact sup_bot_right)
lemma Un_UNIV_left: "UNIV ∪ B = UNIV"
by (fact sup_top_left)
lemma Un_UNIV_right: "A ∪ UNIV = UNIV"
by (fact sup_top_right)
lemma Un_insert_left [simp]: "(insert a B) ∪ C = insert a (B ∪ C)"
by blast
lemma Un_insert_right [simp]: "A ∪ (insert a B) = insert a (A ∪ B)"
by blast
lemma Int_insert_left: "(insert a B) ∩ C = (if a ∈ C then insert a (B ∩ C) else B ∩ C)"
by auto
lemma Int_insert_left_if0 [simp]: "a ∉ C ⟹ (insert a B) ∩ C = B ∩ C"
by auto
lemma Int_insert_left_if1 [simp]: "a ∈ C ⟹ (insert a B) ∩ C = insert a (B ∩ C)"
by auto
lemma Int_insert_right: "A ∩ (insert a B) = (if a ∈ A then insert a (A ∩ B) else A ∩ B)"
by auto
lemma Int_insert_right_if0 [simp]: "a ∉ A ⟹ A ∩ (insert a B) = A ∩ B"
by auto
lemma Int_insert_right_if1 [simp]: "a ∈ A ⟹ A ∩ (insert a B) = insert a (A ∩ B)"
by auto
lemma Un_Int_distrib: "A ∪ (B ∩ C) = (A ∪ B) ∩ (A ∪ C)"
by (fact sup_inf_distrib1)
lemma Un_Int_distrib2: "(B ∩ C) ∪ A = (B ∪ A) ∩ (C ∪ A)"
by (fact sup_inf_distrib2)
lemma Un_Int_crazy: "(A ∩ B) ∪ (B ∩ C) ∪ (C ∩ A) = (A ∪ B) ∩ (B ∪ C) ∩ (C ∪ A)"
by blast
lemma subset_Un_eq: "A ⊆ B ⟷ A ∪ B = B"
by (fact le_iff_sup)
lemma Un_empty [iff]: "A ∪ B = {} ⟷ A = {} ∧ B = {}"
by (fact sup_eq_bot_iff)
lemma Un_subset_iff: "A ∪ B ⊆ C ⟷ A ⊆ C ∧ B ⊆ C"
by (fact le_sup_iff)
lemma Un_Diff_Int: "(A - B) ∪ (A ∩ B) = A"
by blast
lemma Diff_Int2: "A ∩ C - B ∩ C = A ∩ C - B"
by blast
lemma subset_UnE:
assumes "C ⊆ A ∪ B"
obtains A' B' where "A' ⊆ A" "B' ⊆ B" "C = A' ∪ B'"
proof
show "C ∩ A ⊆ A" "C ∩ B ⊆ B" "C = (C ∩ A) ∪ (C ∩ B)"
using assms by blast+
qed
lemma Un_Int_eq [simp]: "(S ∪ T) ∩ S = S" "(S ∪ T) ∩ T = T" "S ∩ (S ∪ T) = S" "T ∩ (S ∪ T) = T"
by auto
lemma Int_Un_eq [simp]: "(S ∩ T) ∪ S = S" "(S ∩ T) ∪ T = T" "S ∪ (S ∩ T) = S" "T ∪ (S ∩ T) = T"
by auto
text ‹┉ Set complement›
lemma Compl_disjoint [simp]: "A ∩ - A = {}"
by (fact inf_compl_bot)
lemma Compl_disjoint2 [simp]: "- A ∩ A = {}"
by (fact compl_inf_bot)
lemma Compl_partition: "A ∪ - A = UNIV"
by (fact sup_compl_top)
lemma Compl_partition2: "- A ∪ A = UNIV"
by (fact compl_sup_top)
lemma double_complement: "- (-A) = A" for A :: "'a set"
by (fact double_compl)
lemma Compl_Un: "- (A ∪ B) = (- A) ∩ (- B)"
by (fact compl_sup)
lemma Compl_Int: "- (A ∩ B) = (- A) ∪ (- B)"
by (fact compl_inf)
lemma subset_Compl_self_eq: "A ⊆ - A ⟷ A = {}"
by blast
lemma Un_Int_assoc_eq: "(A ∩ B) ∪ C = A ∩ (B ∪ C) ⟷ C ⊆ A"
by blast
lemma Compl_UNIV_eq: "- UNIV = {}"
by (fact compl_top_eq)
lemma Compl_empty_eq: "- {} = UNIV"
by (fact compl_bot_eq)
lemma Compl_subset_Compl_iff [iff]: "- A ⊆ - B ⟷ B ⊆ A"
by (fact compl_le_compl_iff)
lemma Compl_eq_Compl_iff [iff]: "- A = - B ⟷ A = B"
for A B :: "'a set"
by (fact compl_eq_compl_iff)
lemma Compl_insert: "- insert x A = (- A) - {x}"
by blast
text ‹┉ Bounded quantifiers.
The following are not added to the default simpset because
(a) they duplicate the body and (b) there are no similar rules for ‹Int›.
›
lemma ball_Un: "(∀x ∈ A ∪ B. P x) ⟷ (∀x∈A. P x) ∧ (∀x∈B. P x)"
by blast
lemma bex_Un: "(∃x ∈ A ∪ B. P x) ⟷ (∃x∈A. P x) ∨ (∃x∈B. P x)"
by blast
text ‹┉ Set difference.›
lemma Diff_eq: "A - B = A ∩ (- B)"
by(rule boolean_algebra_class.diff_eq)
lemma Diff_eq_empty_iff: "A - B = {} ⟷ A ⊆ B"
by(rule boolean_algebra_class.diff_shunt_var)
lemma Diff_cancel [simp]: "A - A = {}"
by blast
lemma Diff_idemp [simp]: "(A - B) - B = A - B"
for A B :: "'a set"
by blast
lemma Diff_triv: "A ∩ B = {} ⟹ A - B = A"
by (blast elim: equalityE)
lemma empty_Diff [simp]: "{} - A = {}"
by blast
lemma Diff_empty [simp]: "A - {} = A"
by blast
lemma Diff_UNIV [simp]: "A - UNIV = {}"
by blast
lemma Diff_insert0 [simp]: "x ∉ A ⟹ A - insert x B = A - B"
by blast
lemma Diff_insert: "A - insert a B = A - B - {a}"
by blast
lemma Diff_insert2: "A - insert a B = A - {a} - B"
by blast
lemma insert_Diff_if: "insert x A - B = (if x ∈ B then A - B else insert x (A - B))"
by auto
lemma insert_Diff1 [simp]: "x ∈ B ⟹ insert x A - B = A - B"
by blast
lemma insert_Diff_single[simp]: "insert a (A - {a}) = insert a A"
by blast
lemma insert_Diff: "a ∈ A ⟹ insert a (A - {a}) = A"
by blast
lemma Diff_insert_absorb: "x ∉ A ⟹ (insert x A) - {x} = A"
by auto
lemma Diff_disjoint [simp]: "A ∩ (B - A) = {}"
by blast
lemma Diff_partition: "A ⊆ B ⟹ A ∪ (B - A) = B"
by blast
lemma double_diff: "A ⊆ B ⟹ B ⊆ C ⟹ B - (C - A) = A"
by blast
lemma Un_Diff_cancel [simp]: "A ∪ (B - A) = A ∪ B"
by blast
lemma Un_Diff_cancel2 [simp]: "(B - A) ∪ A = B ∪ A"
by blast
lemma Diff_Un: "A - (B ∪ C) = (A - B) ∩ (A - C)"
by blast
lemma Diff_Int: "A - (B ∩ C) = (A - B) ∪ (A - C)"
by blast
lemma Diff_Diff_Int: "A - (A - B) = A ∩ B"
by blast
lemma Un_Diff: "(A ∪ B) - C = (A - C) ∪ (B - C)"
by blast
lemma Int_Diff: "(A ∩ B) - C = A ∩ (B - C)"
by blast
lemma Diff_Int_distrib: "C ∩ (A - B) = (C ∩ A) - (C ∩ B)"
by blast
lemma Diff_Int_distrib2: "(A - B) ∩ C = (A ∩ C) - (B ∩ C)"
by blast
lemma Diff_Compl [simp]: "A - (- B) = A ∩ B"
by auto
lemma Compl_Diff_eq [simp]: "- (A - B) = - A ∪ B"
by blast
lemma subset_Compl_singleton [simp]: "A ⊆ - {b} ⟷ b ∉ A"
by blast
text ‹┉ Quantification over type \<^typ>‹bool›.›
lemma bool_induct: "P True ⟹ P False ⟹ P x"
by (cases x) auto
lemma all_bool_eq: "(∀b. P b) ⟷ P True ∧ P False"
by (auto intro: bool_induct)
lemma bool_contrapos: "P x ⟹ ¬ P False ⟹ P True"
by (cases x) auto
lemma ex_bool_eq: "(∃b. P b) ⟷ P True ∨ P False"
by (auto intro: bool_contrapos)
lemma UNIV_bool: "UNIV = {False, True}"
by (auto intro: bool_induct)
text ‹┉ ‹Pow››
lemma Pow_empty [simp]: "Pow {} = {{}}"
by (auto simp add: Pow_def)
lemma Pow_singleton_iff [simp]: "Pow X = {Y} ⟷ X = {} ∧ Y = {}"
by blast
lemma Pow_insert: "Pow (insert a A) = Pow A ∪ (insert a ` Pow A)"
by (blast intro: image_eqI [where ?x = "u - {a}" for u])
lemma Pow_Compl: "Pow (- A) = {- B | B. A ∈ Pow B}"
by (blast intro: exI [where ?x = "- u" for u])
lemma Pow_UNIV [simp]: "Pow UNIV = UNIV"
by blast
lemma Un_Pow_subset: "Pow A ∪ Pow B ⊆ Pow (A ∪ B)"
by blast
lemma Pow_Int_eq [simp]: "Pow (A ∩ B) = Pow A ∩ Pow B"
by blast
text ‹┉ Miscellany.›
lemma Int_Diff_disjoint: "A ∩ B ∩ (A - B) = {}"
by blast
lemma Int_Diff_Un: "A ∩ B ∪ (A - B) = A"
by blast
lemma set_eq_subset: "A = B ⟷ A ⊆ B ∧ B ⊆ A"
by blast
lemma subset_iff: "A ⊆ B ⟷ (∀t. t ∈ A ⟶ t ∈ B)"
by blast
lemma subset_iff_psubset_eq: "A ⊆ B ⟷ A ⊂ B ∨ A = B"
unfolding less_le by blast
lemma all_not_in_conv [simp]: "(∀x. x ∉ A) ⟷ A = {}"
by blast
lemma ex_in_conv: "(∃x. x ∈ A) ⟷ A ≠ {}"
by blast
lemma ball_simps [simp, no_atp]:
"⋀A P Q. (∀x∈A. P x ∨ Q) ⟷ ((∀x∈A. P x) ∨ Q)"
"⋀A P Q. (∀x∈A. P ∨ Q x) ⟷ (P ∨ (∀x∈A. Q x))"
"⋀A P Q. (∀x∈A. P ⟶ Q x) ⟷ (P ⟶ (∀x∈A. Q x))"
"⋀A P Q. (∀x∈A. P x ⟶ Q) ⟷ ((∃x∈A. P x) ⟶ Q)"
"⋀P. (∀x∈{}. P x) ⟷ True"
"⋀P. (∀x∈UNIV. P x) ⟷ (∀x. P x)"
"⋀a B P. (∀x∈insert a B. P x) ⟷ (P a ∧ (∀x∈B. P x))"
"⋀P Q. (∀x∈Collect Q. P x) ⟷ (∀x. Q x ⟶ P x)"
"⋀A P f. (∀x∈f`A. P x) ⟷ (∀x∈A. P (f x))"
"⋀A P. (¬ (∀x∈A. P x)) ⟷ (∃x∈A. ¬ P x)"
by auto
lemma bex_simps [simp, no_atp]:
"⋀A P Q. (∃x∈A. P x ∧ Q) ⟷ ((∃x∈A. P x) ∧ Q)"
"⋀A P Q. (∃x∈A. P ∧ Q x) ⟷ (P ∧ (∃x∈A. Q x))"
"⋀P. (∃x∈{}. P x) ⟷ False"
"⋀P. (∃x∈UNIV. P x) ⟷ (∃x. P x)"
"⋀a B P. (∃x∈insert a B. P x) ⟷ (P a ∨ (∃x∈B. P x))"
"⋀P Q. (∃x∈Collect Q. P x) ⟷ (∃x. Q x ∧ P x)"
"⋀A P f. (∃x∈f`A. P x) ⟷ (∃x∈A. P (f x))"
"⋀A P. (¬(∃x∈A. P x)) ⟷ (∀x∈A. ¬ P x)"
by auto
lemma ex_image_cong_iff [simp, no_atp]:
"(∃x. x∈f`A) ⟷ A ≠ {}" "(∃x. x∈f`A ∧ P x) ⟷ (∃x∈A. P (f x))"
by auto
subsubsection ‹Monotonicity of various operations›
lemma image_mono: "A ⊆ B ⟹ f ` A ⊆ f ` B"
by blast
lemma Pow_mono: "A ⊆ B ⟹ Pow A ⊆ Pow B"
by blast
lemma insert_mono: "C ⊆ D ⟹ insert a C ⊆ insert a D"
by blast
lemma Un_mono: "A ⊆ C ⟹ B ⊆ D ⟹ A ∪ B ⊆ C ∪ D"
by (fact sup_mono)
lemma Int_mono: "A ⊆ C ⟹ B ⊆ D ⟹ A ∩ B ⊆ C ∩ D"
by (fact inf_mono)
lemma Diff_mono: "A ⊆ C ⟹ D ⊆ B ⟹ A - B ⊆ C - D"
by blast
lemma Compl_anti_mono: "A ⊆ B ⟹ - B ⊆ - A"
by (fact compl_mono)
text ‹┉ Monotonicity of implications.›
lemma in_mono: "A ⊆ B ⟹ x ∈ A ⟶ x ∈ B"
by (rule impI) (erule subsetD)
lemma conj_mono: "P1 ⟶ Q1 ⟹ P2 ⟶ Q2 ⟹ (P1 ∧ P2) ⟶ (Q1 ∧ Q2)"
by iprover
lemma disj_mono: "P1 ⟶ Q1 ⟹ P2 ⟶ Q2 ⟹ (P1 ∨ P2) ⟶ (Q1 ∨ Q2)"
by iprover
lemma imp_mono: "Q1 ⟶ P1 ⟹ P2 ⟶ Q2 ⟹ (P1 ⟶ P2) ⟶ (Q1 ⟶ Q2)"
by iprover
lemma imp_refl: "P ⟶ P" ..
lemma not_mono: "Q ⟶ P ⟹ ¬ P ⟶ ¬ Q"
by iprover
lemma ex_mono: "(⋀x. P x ⟶ Q x) ⟹ (∃x. P x) ⟶ (∃x. Q x)"
by iprover
lemma all_mono: "(⋀x. P x ⟶ Q x) ⟹ (∀x. P x) ⟶ (∀x. Q x)"
by iprover
lemma Collect_mono: "(⋀x. P x ⟶ Q x) ⟹ Collect P ⊆ Collect Q"
by blast
lemma Int_Collect_mono: "A ⊆ B ⟹ (⋀x. x ∈ A ⟹ P x ⟶ Q x) ⟹ A ∩ Collect P ⊆ B ∩ Collect Q"
by blast
lemmas basic_monos =
subset_refl imp_refl disj_mono conj_mono ex_mono Collect_mono in_mono
lemma eq_to_mono: "a = b ⟹ c = d ⟹ b ⟶ d ⟹ a ⟶ c"
by iprover
subsubsection ‹Inverse image of a function›
definition vimage :: "('a ⇒ 'b) ⇒ 'b set ⇒ 'a set" (infixr "-`" 90)
where "f -` B ≡ {x. f x ∈ B}"
lemma vimage_eq [simp]: "a ∈ f -` B ⟷ f a ∈ B"
unfolding vimage_def by blast
lemma vimage_singleton_eq: "a ∈ f -` {b} ⟷ f a = b"
by simp
lemma vimageI [intro]: "f a = b ⟹ b ∈ B ⟹ a ∈ f -` B"
unfolding vimage_def by blast
lemma vimageI2: "f a ∈ A ⟹ a ∈ f -` A"
unfolding vimage_def by fast
lemma vimageE [elim!]: "a ∈ f -` B ⟹ (⋀x. f a = x ⟹ x ∈ B ⟹ P) ⟹ P"
unfolding vimage_def by blast
lemma vimageD: "a ∈ f -` A ⟹ f a ∈ A"
unfolding vimage_def by fast
lemma vimage_empty [simp]: "f -` {} = {}"
by blast
lemma vimage_Compl: "f -` (- A) = - (f -` A)"
by blast
lemma vimage_Un [simp]: "f -` (A ∪ B) = (f -` A) ∪ (f -` B)"
by blast
lemma vimage_Int [simp]: "f -` (A ∩ B) = (f -` A) ∩ (f -` B)"
by fast
lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}"
by blast
lemma vimage_Collect: "(⋀x. P (f x) = Q x) ⟹ f -` (Collect P) = Collect Q"
by blast
lemma vimage_insert: "f -` (insert a B) = (f -` {a}) ∪ (f -` B)"
by blast
lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)"
by blast
lemma vimage_UNIV [simp]: "f -` UNIV = UNIV"
by blast
lemma vimage_mono: "A ⊆ B ⟹ f -` A ⊆ f -` B"
by blast
lemma vimage_image_eq: "f -` (f ` A) = {y. ∃x∈A. f x = f y}"
by (blast intro: sym)
lemma image_vimage_subset: "f ` (f -` A) ⊆ A"
by blast
lemma image_vimage_eq [simp]: "f ` (f -` A) = A ∩ range f"
by blast
lemma image_subset_iff_subset_vimage: "f ` A ⊆ B ⟷ A ⊆ f -` B"
by blast
lemma subset_vimage_iff: "A ⊆ f -` B ⟷ (∀x∈A. f x ∈ B)"
by auto
lemma vimage_const [simp]: "((λx. c) -` A) = (if c ∈ A then UNIV else {})"
by auto
lemma vimage_if [simp]: "((λx. if x ∈ B then c else d) -` A) =
(if c ∈ A then (if d ∈ A then UNIV else B)
else if d ∈ A then - B else {})"
by (auto simp add: vimage_def)
lemma vimage_inter_cong: "(⋀ w. w ∈ S ⟹ f w = g w) ⟹ f -` y ∩ S = g -` y ∩ S"
by auto
lemma vimage_ident [simp]: "(λx. x) -` Y = Y"
by blast
subsubsection ‹Singleton sets›
definition is_singleton :: "'a set ⇒ bool"
where "is_singleton A ⟷ (∃x. A = {x})"
lemma is_singletonI [simp, intro!]: "is_singleton {x}"
unfolding is_singleton_def by simp
lemma is_singletonI': "A ≠ {} ⟹ (⋀x y. x ∈ A ⟹ y ∈ A ⟹ x = y) ⟹ is_singleton A"
unfolding is_singleton_def by blast
lemma is_singletonE: "is_singleton A ⟹ (⋀x. A = {x} ⟹ P) ⟹ P"
unfolding is_singleton_def by blast
subsubsection ‹Getting the contents of a singleton set›
definition the_elem :: "'a set ⇒ 'a"
where "the_elem X = (THE x. X = {x})"
lemma the_elem_eq [simp]: "the_elem {x} = x"
by (simp add: the_elem_def)
lemma is_singleton_the_elem: "is_singleton A ⟷ A = {the_elem A}"
by (auto simp: is_singleton_def)
lemma the_elem_image_unique:
assumes "A ≠ {}"
and *: "⋀y. y ∈ A ⟹ f y = f x"
shows "the_elem (f ` A) = f x"
unfolding the_elem_def
proof (rule the1_equality)
from ‹A ≠ {}› obtain y where "y ∈ A" by auto
with * have "f x = f y" by simp
with ‹y ∈ A› have "f x ∈ f ` A" by blast
with * show "f ` A = {f x}" by auto
then show "∃!x. f ` A = {x}" by auto
qed
subsubsection ‹Monad operation›
definition bind :: "'a set ⇒ ('a ⇒ 'b set) ⇒ 'b set"
where "bind A f = {x. ∃B ∈ f`A. x ∈ B}"
hide_const (open) bind
lemma bind_bind: "Set.bind (Set.bind A B) C = Set.bind A (λx. Set.bind (B x) C)"
for A :: "'a set"
by (auto simp: bind_def)
lemma empty_bind [simp]: "Set.bind {} f = {}"
by (simp add: bind_def)
lemma nonempty_bind_const: "A ≠ {} ⟹ Set.bind A (λ_. B) = B"
by (auto simp: bind_def)
lemma bind_const: "Set.bind A (λ_. B) = (if A = {} then {} else B)"
by (auto simp: bind_def)
lemma bind_singleton_conv_image: "Set.bind A (λx. {f x}) = f ` A"
by (auto simp: bind_def)
subsubsection ‹Operations for execution›
definition is_empty :: "'a set ⇒ bool"
where [code_abbrev]: "is_empty A ⟷ A = {}"
hide_const (open) is_empty
definition remove :: "'a ⇒ 'a set ⇒ 'a set"
where [code_abbrev]: "remove x A = A - {x}"
hide_const (open) remove
lemma member_remove [simp]: "x ∈ Set.remove y A ⟷ x ∈ A ∧ x ≠ y"
by (simp add: remove_def)
definition filter :: "('a ⇒ bool) ⇒ 'a set ⇒ 'a set"
where [code_abbrev]: "filter P A = {a ∈ A. P a}"
hide_const (open) filter
lemma member_filter [simp]: "x ∈ Set.filter P A ⟷ x ∈ A ∧ P x"
by (simp add: filter_def)
instantiation set :: (equal) equal
begin
definition "HOL.equal A B ⟷ A ⊆ B ∧ B ⊆ A"
instance by standard (auto simp add: equal_set_def)
end
text ‹Misc›
definition pairwise :: "('a ⇒ 'a ⇒ bool) ⇒ 'a set ⇒ bool"
where "pairwise R S ⟷ (∀x ∈ S. ∀y ∈ S. x ≠ y ⟶ R x y)"
lemma pairwise_alt: "pairwise R S ⟷ (∀x∈S. ∀y∈S-{x}. R x y)"
by (auto simp add: pairwise_def)
lemma pairwise_trivial [simp]: "pairwise (λi j. j ≠ i) I"
by (auto simp: pairwise_def)
lemma pairwiseI [intro?]:
"pairwise R S" if "⋀x y. x ∈ S ⟹ y ∈ S ⟹ x ≠ y ⟹ R x y"
using that by (simp add: pairwise_def)
lemma pairwiseD:
"R x y" and "R y x"
if "pairwise R S" "x ∈ S" and "y ∈ S" and "x ≠ y"
using that by (simp_all add: pairwise_def)
lemma pairwise_empty [simp]: "pairwise P {}"
by (simp add: pairwise_def)
lemma pairwise_singleton [simp]: "pairwise P {A}"
by (simp add: pairwise_def)
lemma pairwise_insert:
"pairwise r (insert x s) ⟷ (∀y. y ∈ s ∧ y ≠ x ⟶ r x y ∧ r y x) ∧ pairwise r s"
by (force simp: pairwise_def)
lemma pairwise_subset: "pairwise P S ⟹ T ⊆ S ⟹ pairwise P T"
by (force simp: pairwise_def)
lemma pairwise_mono: "⟦pairwise P A; ⋀x y. P x y ⟹ Q x y; B ⊆ A⟧ ⟹ pairwise Q B"
by (fastforce simp: pairwise_def)
lemma pairwise_imageI:
"pairwise P (f ` A)"
if "⋀x y. x ∈ A ⟹ y ∈ A ⟹ x ≠ y ⟹ f x ≠ f y ⟹ P (f x) (f y)"
using that by (auto intro: pairwiseI)
lemma pairwise_image: "pairwise r (f ` s) ⟷ pairwise (λx y. (f x ≠ f y) ⟶ r (f x) (f y)) s"
by (force simp: pairwise_def)
definition disjnt :: "'a set ⇒ 'a set ⇒ bool"
where "disjnt A B ⟷ A ∩ B = {}"
lemma disjnt_self_iff_empty [simp]: "disjnt S S ⟷ S = {}"
by (auto simp: disjnt_def)
lemma disjnt_commute: "disjnt A B = disjnt B A"
by (auto simp: disjnt_def)
lemma disjnt_iff: "disjnt A B ⟷ (∀x. ¬ (x ∈ A ∧ x ∈ B))"
by (force simp: disjnt_def)
lemma disjnt_sym: "disjnt A B ⟹ disjnt B A"
using disjnt_iff by blast
lemma disjnt_empty1 [simp]: "disjnt {} A" and disjnt_empty2 [simp]: "disjnt A {}"
by (auto simp: disjnt_def)
lemma disjnt_insert1 [simp]: "disjnt (insert a X) Y ⟷ a ∉ Y ∧ disjnt X Y"
by (simp add: disjnt_def)
lemma disjnt_insert2 [simp]: "disjnt Y (insert a X) ⟷ a ∉ Y ∧ disjnt Y X"
by (simp add: disjnt_def)
lemma disjnt_subset1 : "⟦disjnt X Y; Z ⊆ X⟧ ⟹ disjnt Z Y"
by (auto simp: disjnt_def)
lemma disjnt_subset2 : "⟦disjnt X Y; Z ⊆ Y⟧ ⟹ disjnt X Z"
by (auto simp: disjnt_def)
lemma disjnt_Un1 [simp]: "disjnt (A ∪ B) C ⟷ disjnt A C ∧ disjnt B C"
by (auto simp: disjnt_def)
lemma disjnt_Un2 [simp]: "disjnt C (A ∪ B) ⟷ disjnt C A ∧ disjnt C B"
by (auto simp: disjnt_def)
lemma disjnt_Diff1: "disjnt (X-Y) (U-V)" and disjnt_Diff2: "disjnt (U-V) (X-Y)" if "X ⊆ V"
using that by (auto simp: disjnt_def)
lemma disjoint_image_subset: "⟦pairwise disjnt 𝒜; ⋀X. X ∈ 𝒜 ⟹ f X ⊆ X⟧ ⟹ pairwise disjnt (f `𝒜)"
unfolding disjnt_def pairwise_def by fast
lemma pairwise_disjnt_iff: "pairwise disjnt 𝒜 ⟷ (∀x. ∃⇩≤⇩1 X. X ∈ 𝒜 ∧ x ∈ X)"
by (auto simp: Uniq_def disjnt_iff pairwise_def)
lemma disjnt_insert:
‹disjnt (insert x M) N› if ‹x ∉ N› ‹disjnt M N›
using that by (simp add: disjnt_def)
lemma Int_emptyI: "(⋀x. x ∈ A ⟹ x ∈ B ⟹ False) ⟹ A ∩ B = {}"
by blast
lemma in_image_insert_iff:
assumes "⋀C. C ∈ B ⟹ x ∉ C"
shows "A ∈ insert x ` B ⟷ x ∈ A ∧ A - {x} ∈ B" (is "?P ⟷ ?Q")
proof
assume ?P then show ?Q
using assms by auto
next
assume ?Q
then have "x ∈ A" and "A - {x} ∈ B"
by simp_all
from ‹A - {x} ∈ B› have "insert x (A - {x}) ∈ insert x ` B"
by (rule imageI)
also from ‹x ∈ A›
have "insert x (A - {x}) = A"
by auto
finally show ?P .
qed
hide_const (open) member not_member
lemmas equalityI = subset_antisym
lemmas set_mp = subsetD
lemmas set_rev_mp = rev_subsetD
ML ‹
val Ball_def = @{thm Ball_def}
val Bex_def = @{thm Bex_def}
val CollectD = @{thm CollectD}
val CollectE = @{thm CollectE}
val CollectI = @{thm CollectI}
val Collect_conj_eq = @{thm Collect_conj_eq}
val Collect_mem_eq = @{thm Collect_mem_eq}
val IntD1 = @{thm IntD1}
val IntD2 = @{thm IntD2}
val IntE = @{thm IntE}
val IntI = @{thm IntI}
val Int_Collect = @{thm Int_Collect}
val UNIV_I = @{thm UNIV_I}
val UNIV_witness = @{thm UNIV_witness}
val UnE = @{thm UnE}
val UnI1 = @{thm UnI1}
val UnI2 = @{thm UnI2}
val ballE = @{thm ballE}
val ballI = @{thm ballI}
val bexCI = @{thm bexCI}
val bexE = @{thm bexE}
val bexI = @{thm bexI}
val bex_triv = @{thm bex_triv}
val bspec = @{thm bspec}
val contra_subsetD = @{thm contra_subsetD}
val equalityCE = @{thm equalityCE}
val equalityD1 = @{thm equalityD1}
val equalityD2 = @{thm equalityD2}
val equalityE = @{thm equalityE}
val equalityI = @{thm equalityI}
val imageE = @{thm imageE}
val imageI = @{thm imageI}
val image_Un = @{thm image_Un}
val image_insert = @{thm image_insert}
val insert_commute = @{thm insert_commute}
val insert_iff = @{thm insert_iff}
val mem_Collect_eq = @{thm mem_Collect_eq}
val rangeE = @{thm rangeE}
val rangeI = @{thm rangeI}
val range_eqI = @{thm range_eqI}
val subsetCE = @{thm subsetCE}
val subsetD = @{thm subsetD}
val subsetI = @{thm subsetI}
val subset_refl = @{thm subset_refl}
val subset_trans = @{thm subset_trans}
val vimageD = @{thm vimageD}
val vimageE = @{thm vimageE}
val vimageI = @{thm vimageI}
val vimageI2 = @{thm vimageI2}
val vimage_Collect = @{thm vimage_Collect}
val vimage_Int = @{thm vimage_Int}
val vimage_Un = @{thm vimage_Un}
›
end