haftmann@44103: (* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel; Florian Haftmann, TU Muenchen *) wenzelm@11979: haftmann@44104: header {* Complete lattices *} haftmann@32077: haftmann@44860: theory Complete_Lattices haftmann@32139: imports Set haftmann@32139: begin haftmann@32077: haftmann@32077: notation haftmann@34007: less_eq (infix "\" 50) and haftmann@46691: less (infix "\" 50) haftmann@32077: haftmann@32139: haftmann@32879: subsection {* Syntactic infimum and supremum operations *} haftmann@32879: haftmann@32879: class Inf = haftmann@32879: fixes Inf :: "'a set \ 'a" ("\_" [900] 900) haftmann@32879: haftmann@32879: class Sup = haftmann@32879: fixes Sup :: "'a set \ 'a" ("\_" [900] 900) haftmann@32879: haftmann@46691: haftmann@32139: subsection {* Abstract complete lattices *} haftmann@32139: haftmann@52729: text {* A complete lattice always has a bottom and a top, haftmann@52729: so we include them into the following type class, haftmann@52729: along with assumptions that define bottom and top haftmann@52729: in terms of infimum and supremum. *} haftmann@52729: haftmann@52729: class complete_lattice = lattice + Inf + Sup + bot + top + haftmann@32077: assumes Inf_lower: "x \ A \ \A \ x" haftmann@32077: and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" haftmann@32077: assumes Sup_upper: "x \ A \ x \ \A" haftmann@32077: and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" haftmann@52729: assumes Inf_empty [simp]: "\{} = \" haftmann@52729: assumes Sup_empty [simp]: "\{} = \" haftmann@32077: begin haftmann@32077: haftmann@52729: subclass bounded_lattice haftmann@52729: proof haftmann@52729: fix a haftmann@52729: show "\ \ a" by (auto intro: Sup_least simp only: Sup_empty [symmetric]) haftmann@52729: show "a \ \" by (auto intro: Inf_greatest simp only: Inf_empty [symmetric]) haftmann@52729: qed haftmann@52729: haftmann@32678: lemma dual_complete_lattice: krauss@44845: "class.complete_lattice Sup Inf sup (op \) (op >) inf \ \" haftmann@52729: by (auto intro!: class.complete_lattice.intro dual_lattice) haftmann@52729: (unfold_locales, (fact Inf_empty Sup_empty haftmann@34007: Sup_upper Sup_least Inf_lower Inf_greatest)+) haftmann@32678: haftmann@44040: definition INFI :: "'b set \ ('b \ 'a) \ 'a" where haftmann@44040: INF_def: "INFI A f = \(f ` A)" haftmann@44040: haftmann@44040: definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where haftmann@44040: SUP_def: "SUPR A f = \(f ` A)" haftmann@44040: haftmann@44040: text {* haftmann@44040: Note: must use names @{const INFI} and @{const SUPR} here instead of haftmann@44040: @{text INF} and @{text SUP} to allow the following syntax coexist haftmann@44040: with the plain constant names. haftmann@44040: *} haftmann@44040: haftmann@44040: end haftmann@44040: haftmann@44040: syntax haftmann@44040: "_INF1" :: "pttrns \ 'b \ 'b" ("(3INF _./ _)" [0, 10] 10) haftmann@44040: "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3INF _:_./ _)" [0, 0, 10] 10) haftmann@44040: "_SUP1" :: "pttrns \ 'b \ 'b" ("(3SUP _./ _)" [0, 10] 10) haftmann@44040: "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3SUP _:_./ _)" [0, 0, 10] 10) haftmann@44040: haftmann@44040: syntax (xsymbols) haftmann@44040: "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@44040: "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@44040: "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@44040: "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@44040: haftmann@44040: translations haftmann@44040: "INF x y. B" == "INF x. INF y. B" haftmann@44040: "INF x. B" == "CONST INFI CONST UNIV (%x. B)" haftmann@44040: "INF x. B" == "INF x:CONST UNIV. B" haftmann@44040: "INF x:A. B" == "CONST INFI A (%x. B)" haftmann@44040: "SUP x y. B" == "SUP x. SUP y. B" haftmann@44040: "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)" haftmann@44040: "SUP x. B" == "SUP x:CONST UNIV. B" haftmann@44040: "SUP x:A. B" == "CONST SUPR A (%x. B)" haftmann@44040: haftmann@44040: print_translation {* haftmann@44040: [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"}, haftmann@44040: Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}] haftmann@44040: *} -- {* to avoid eta-contraction of body *} haftmann@44040: haftmann@44040: context complete_lattice haftmann@44040: begin haftmann@32077: blanchet@54147: lemma INF_foundation_dual: haftmann@44040: "complete_lattice.SUPR Inf = INFI" huffman@44921: by (simp add: fun_eq_iff INF_def huffman@44921: complete_lattice.SUP_def [OF dual_complete_lattice]) haftmann@44040: blanchet@54147: lemma SUP_foundation_dual: haftmann@44040: "complete_lattice.INFI Sup = SUPR" huffman@44921: by (simp add: fun_eq_iff SUP_def huffman@44921: complete_lattice.INF_def [OF dual_complete_lattice]) haftmann@44040: hoelzl@51328: lemma Sup_eqI: hoelzl@51328: "(\y. y \ A \ y \ x) \ (\y. (\z. z \ A \ z \ y) \ x \ y) \ \A = x" hoelzl@51328: by (blast intro: antisym Sup_least Sup_upper) hoelzl@51328: hoelzl@51328: lemma Inf_eqI: hoelzl@51328: "(\i. i \ A \ x \ i) \ (\y. (\i. i \ A \ y \ i) \ y \ x) \ \A = x" hoelzl@51328: by (blast intro: antisym Inf_greatest Inf_lower) hoelzl@51328: hoelzl@51328: lemma SUP_eqI: hoelzl@51328: "(\i. i \ A \ f i \ x) \ (\y. (\i. i \ A \ f i \ y) \ x \ y) \ (\i\A. f i) = x" hoelzl@51328: unfolding SUP_def by (rule Sup_eqI) auto hoelzl@51328: hoelzl@51328: lemma INF_eqI: hoelzl@51328: "(\i. i \ A \ x \ f i) \ (\y. (\i. i \ A \ f i \ y) \ x \ y) \ (\i\A. f i) = x" hoelzl@51328: unfolding INF_def by (rule Inf_eqI) auto hoelzl@51328: haftmann@44103: lemma INF_lower: "i \ A \ (\i\A. f i) \ f i" haftmann@44040: by (auto simp add: INF_def intro: Inf_lower) haftmann@44040: haftmann@44103: lemma INF_greatest: "(\i. i \ A \ u \ f i) \ u \ (\i\A. f i)" haftmann@44103: by (auto simp add: INF_def intro: Inf_greatest) haftmann@44103: haftmann@44103: lemma SUP_upper: "i \ A \ f i \ (\i\A. f i)" haftmann@44040: by (auto simp add: SUP_def intro: Sup_upper) haftmann@44040: haftmann@44103: lemma SUP_least: "(\i. i \ A \ f i \ u) \ (\i\A. f i) \ u" haftmann@44040: by (auto simp add: SUP_def intro: Sup_least) haftmann@44040: haftmann@44040: lemma Inf_lower2: "u \ A \ u \ v \ \A \ v" haftmann@44040: using Inf_lower [of u A] by auto haftmann@44040: haftmann@44103: lemma INF_lower2: "i \ A \ f i \ u \ (\i\A. f i) \ u" haftmann@44103: using INF_lower [of i A f] by auto haftmann@44040: haftmann@44040: lemma Sup_upper2: "u \ A \ v \ u \ v \ \A" haftmann@44040: using Sup_upper [of u A] by auto haftmann@44040: haftmann@44103: lemma SUP_upper2: "i \ A \ u \ f i \ u \ (\i\A. f i)" haftmann@44103: using SUP_upper [of i A f] by auto haftmann@44040: noschinl@44918: lemma le_Inf_iff: "b \ \A \ (\a\A. b \ a)" haftmann@44040: by (auto intro: Inf_greatest dest: Inf_lower) haftmann@44040: noschinl@44918: lemma le_INF_iff: "u \ (\i\A. f i) \ (\i\A. u \ f i)" haftmann@44040: by (auto simp add: INF_def le_Inf_iff) haftmann@44040: noschinl@44918: lemma Sup_le_iff: "\A \ b \ (\a\A. a \ b)" haftmann@44040: by (auto intro: Sup_least dest: Sup_upper) haftmann@44040: noschinl@44918: lemma SUP_le_iff: "(\i\A. f i) \ u \ (\i\A. f i \ u)" haftmann@44040: by (auto simp add: SUP_def Sup_le_iff) haftmann@32077: haftmann@52729: lemma Inf_insert [simp]: "\insert a A = a \ \A" haftmann@52729: by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) haftmann@52729: haftmann@52729: lemma INF_insert: "(\x\insert a A. f x) = f a \ INFI A f" haftmann@52729: by (simp add: INF_def) haftmann@52729: haftmann@52729: lemma Sup_insert [simp]: "\insert a A = a \ \A" haftmann@52729: by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) haftmann@52729: haftmann@52729: lemma SUP_insert: "(\x\insert a A. f x) = f a \ SUPR A f" haftmann@52729: by (simp add: SUP_def) haftmann@32077: huffman@44067: lemma INF_empty [simp]: "(\x\{}. f x) = \" haftmann@44040: by (simp add: INF_def) haftmann@44040: huffman@44067: lemma SUP_empty [simp]: "(\x\{}. f x) = \" haftmann@44040: by (simp add: SUP_def) haftmann@44040: haftmann@41080: lemma Inf_UNIV [simp]: haftmann@41080: "\UNIV = \" haftmann@44040: by (auto intro!: antisym Inf_lower) haftmann@41080: haftmann@41080: lemma Sup_UNIV [simp]: haftmann@41080: "\UNIV = \" haftmann@44040: by (auto intro!: antisym Sup_upper) haftmann@41080: noschinl@44918: lemma INF_image [simp]: "(\x\f`A. g x) = (\x\A. g (f x))" huffman@44068: by (simp add: INF_def image_image) huffman@44068: noschinl@44918: lemma SUP_image [simp]: "(\x\f`A. g x) = (\x\A. g (f x))" huffman@44068: by (simp add: SUP_def image_image) huffman@44068: haftmann@44040: lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" haftmann@44040: by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) haftmann@44040: haftmann@44040: lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" haftmann@44040: by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) haftmann@44040: haftmann@43899: lemma Inf_superset_mono: "B \ A \ \A \ \B" haftmann@43899: by (auto intro: Inf_greatest Inf_lower) haftmann@43899: haftmann@43899: lemma Sup_subset_mono: "A \ B \ \A \ \B" haftmann@43899: by (auto intro: Sup_least Sup_upper) haftmann@43899: haftmann@44041: lemma INF_cong: haftmann@44041: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" haftmann@44041: by (simp add: INF_def image_def) haftmann@44041: haftmann@44041: lemma SUP_cong: haftmann@44041: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" haftmann@44041: by (simp add: SUP_def image_def) haftmann@44041: hoelzl@38705: lemma Inf_mono: hoelzl@41971: assumes "\b. b \ B \ \a\A. a \ b" haftmann@43741: shows "\A \ \B" hoelzl@38705: proof (rule Inf_greatest) hoelzl@38705: fix b assume "b \ B" hoelzl@41971: with assms obtain a where "a \ A" and "a \ b" by blast haftmann@43741: from `a \ A` have "\A \ a" by (rule Inf_lower) haftmann@43741: with `a \ b` show "\A \ b" by auto hoelzl@38705: qed hoelzl@38705: haftmann@44041: lemma INF_mono: haftmann@44041: "(\m. m \ B \ \n\A. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" noschinl@44918: unfolding INF_def by (rule Inf_mono) fast haftmann@44041: haftmann@41082: lemma Sup_mono: hoelzl@41971: assumes "\a. a \ A \ \b\B. a \ b" haftmann@43741: shows "\A \ \B" haftmann@41082: proof (rule Sup_least) haftmann@41082: fix a assume "a \ A" hoelzl@41971: with assms obtain b where "b \ B" and "a \ b" by blast haftmann@43741: from `b \ B` have "b \ \B" by (rule Sup_upper) haftmann@43741: with `a \ b` show "a \ \B" by auto haftmann@41082: qed haftmann@32077: haftmann@44041: lemma SUP_mono: haftmann@44041: "(\n. n \ A \ \m\B. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" noschinl@44918: unfolding SUP_def by (rule Sup_mono) fast haftmann@44041: haftmann@44041: lemma INF_superset_mono: haftmann@44041: "B \ A \ (\x. x \ B \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" haftmann@44041: -- {* The last inclusion is POSITIVE! *} haftmann@44041: by (blast intro: INF_mono dest: subsetD) haftmann@44041: haftmann@44041: lemma SUP_subset_mono: haftmann@44041: "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" haftmann@44041: by (blast intro: SUP_mono dest: subsetD) haftmann@44041: haftmann@43868: lemma Inf_less_eq: haftmann@43868: assumes "\v. v \ A \ v \ u" haftmann@43868: and "A \ {}" haftmann@43868: shows "\A \ u" haftmann@43868: proof - haftmann@43868: from `A \ {}` obtain v where "v \ A" by blast wenzelm@53374: moreover from `v \ A` assms(1) have "v \ u" by blast haftmann@43868: ultimately show ?thesis by (rule Inf_lower2) haftmann@43868: qed haftmann@43868: haftmann@43868: lemma less_eq_Sup: haftmann@43868: assumes "\v. v \ A \ u \ v" haftmann@43868: and "A \ {}" haftmann@43868: shows "u \ \A" haftmann@43868: proof - haftmann@43868: from `A \ {}` obtain v where "v \ A" by blast wenzelm@53374: moreover from `v \ A` assms(1) have "u \ v" by blast haftmann@43868: ultimately show ?thesis by (rule Sup_upper2) haftmann@43868: qed haftmann@43868: hoelzl@51328: lemma SUPR_eq: hoelzl@51328: assumes "\i. i \ A \ \j\B. f i \ g j" hoelzl@51328: assumes "\j. j \ B \ \i\A. g j \ f i" hoelzl@51328: shows "(SUP i:A. f i) = (SUP j:B. g j)" hoelzl@51328: by (intro antisym SUP_least) (blast intro: SUP_upper2 dest: assms)+ hoelzl@51328: hoelzl@51328: lemma INFI_eq: hoelzl@51328: assumes "\i. i \ A \ \j\B. f i \ g j" hoelzl@51328: assumes "\j. j \ B \ \i\A. g j \ f i" hoelzl@51328: shows "(INF i:A. f i) = (INF j:B. g j)" hoelzl@51328: by (intro antisym INF_greatest) (blast intro: INF_lower2 dest: assms)+ hoelzl@51328: haftmann@43899: lemma less_eq_Inf_inter: "\A \ \B \ \(A \ B)" haftmann@43868: by (auto intro: Inf_greatest Inf_lower) haftmann@43868: haftmann@43899: lemma Sup_inter_less_eq: "\(A \ B) \ \A \ \B " haftmann@43868: by (auto intro: Sup_least Sup_upper) haftmann@43868: haftmann@43868: lemma Inf_union_distrib: "\(A \ B) = \A \ \B" haftmann@43868: by (rule antisym) (auto intro: Inf_greatest Inf_lower le_infI1 le_infI2) haftmann@43868: haftmann@44041: lemma INF_union: haftmann@44041: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" haftmann@44103: by (auto intro!: antisym INF_mono intro: le_infI1 le_infI2 INF_greatest INF_lower) haftmann@44041: haftmann@43868: lemma Sup_union_distrib: "\(A \ B) = \A \ \B" haftmann@43868: by (rule antisym) (auto intro: Sup_least Sup_upper le_supI1 le_supI2) haftmann@43868: haftmann@44041: lemma SUP_union: haftmann@44041: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" haftmann@44103: by (auto intro!: antisym SUP_mono intro: le_supI1 le_supI2 SUP_least SUP_upper) haftmann@44041: haftmann@44041: lemma INF_inf_distrib: "(\a\A. f a) \ (\a\A. g a) = (\a\A. f a \ g a)" haftmann@44103: by (rule antisym) (rule INF_greatest, auto intro: le_infI1 le_infI2 INF_lower INF_mono) haftmann@44041: noschinl@44918: lemma SUP_sup_distrib: "(\a\A. f a) \ (\a\A. g a) = (\a\A. f a \ g a)" (is "?L = ?R") noschinl@44918: proof (rule antisym) noschinl@44918: show "?L \ ?R" by (auto intro: le_supI1 le_supI2 SUP_upper SUP_mono) noschinl@44918: next noschinl@44918: show "?R \ ?L" by (rule SUP_least) (auto intro: le_supI1 le_supI2 SUP_upper) noschinl@44918: qed haftmann@44041: blanchet@54147: lemma Inf_top_conv [simp]: haftmann@43868: "\A = \ \ (\x\A. x = \)" haftmann@43868: "\ = \A \ (\x\A. x = \)" haftmann@43868: proof - haftmann@43868: show "\A = \ \ (\x\A. x = \)" haftmann@43868: proof haftmann@43868: assume "\x\A. x = \" haftmann@43868: then have "A = {} \ A = {\}" by auto noschinl@44919: then show "\A = \" by auto haftmann@43868: next haftmann@43868: assume "\A = \" haftmann@43868: show "\x\A. x = \" haftmann@43868: proof (rule ccontr) haftmann@43868: assume "\ (\x\A. x = \)" haftmann@43868: then obtain x where "x \ A" and "x \ \" by blast haftmann@43868: then obtain B where "A = insert x B" by blast noschinl@44919: with `\A = \` `x \ \` show False by simp haftmann@43868: qed haftmann@43868: qed haftmann@43868: then show "\ = \A \ (\x\A. x = \)" by auto haftmann@43868: qed haftmann@43868: noschinl@44918: lemma INF_top_conv [simp]: haftmann@44041: "(\x\A. B x) = \ \ (\x\A. B x = \)" haftmann@44041: "\ = (\x\A. B x) \ (\x\A. B x = \)" noschinl@44919: by (auto simp add: INF_def) haftmann@44041: blanchet@54147: lemma Sup_bot_conv [simp]: haftmann@43868: "\A = \ \ (\x\A. x = \)" (is ?P) haftmann@43868: "\ = \A \ (\x\A. x = \)" (is ?Q) huffman@44920: using dual_complete_lattice huffman@44920: by (rule complete_lattice.Inf_top_conv)+ haftmann@43868: noschinl@44918: lemma SUP_bot_conv [simp]: haftmann@44041: "(\x\A. B x) = \ \ (\x\A. B x = \)" haftmann@44041: "\ = (\x\A. B x) \ (\x\A. B x = \)" noschinl@44919: by (auto simp add: SUP_def) haftmann@44041: haftmann@43865: lemma INF_const [simp]: "A \ {} \ (\i\A. f) = f" haftmann@44103: by (auto intro: antisym INF_lower INF_greatest) haftmann@32077: haftmann@43870: lemma SUP_const [simp]: "A \ {} \ (\i\A. f) = f" haftmann@44103: by (auto intro: antisym SUP_upper SUP_least) haftmann@43870: noschinl@44918: lemma INF_top [simp]: "(\x\A. \) = \" huffman@44921: by (cases "A = {}") simp_all haftmann@43900: noschinl@44918: lemma SUP_bot [simp]: "(\x\A. \) = \" huffman@44921: by (cases "A = {}") simp_all haftmann@43900: haftmann@43865: lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" haftmann@44103: by (iprover intro: INF_lower INF_greatest order_trans antisym) haftmann@43865: haftmann@43870: lemma SUP_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" haftmann@44103: by (iprover intro: SUP_upper SUP_least order_trans antisym) haftmann@43870: haftmann@43871: lemma INF_absorb: haftmann@43868: assumes "k \ I" haftmann@43868: shows "A k \ (\i\I. A i) = (\i\I. A i)" haftmann@43868: proof - haftmann@43868: from assms obtain J where "I = insert k J" by blast haftmann@43868: then show ?thesis by (simp add: INF_insert) haftmann@43868: qed haftmann@43868: haftmann@43871: lemma SUP_absorb: haftmann@43871: assumes "k \ I" haftmann@43871: shows "A k \ (\i\I. A i) = (\i\I. A i)" haftmann@43871: proof - haftmann@43871: from assms obtain J where "I = insert k J" by blast haftmann@43871: then show ?thesis by (simp add: SUP_insert) haftmann@43871: qed haftmann@43871: haftmann@43871: lemma INF_constant: haftmann@43868: "(\y\A. c) = (if A = {} then \ else c)" huffman@44921: by simp haftmann@43868: haftmann@43871: lemma SUP_constant: haftmann@43871: "(\y\A. c) = (if A = {} then \ else c)" huffman@44921: by simp haftmann@43871: haftmann@43943: lemma less_INF_D: haftmann@43943: assumes "y < (\i\A. f i)" "i \ A" shows "y < f i" haftmann@43943: proof - haftmann@43943: note `y < (\i\A. f i)` haftmann@43943: also have "(\i\A. f i) \ f i" using `i \ A` haftmann@44103: by (rule INF_lower) haftmann@43943: finally show "y < f i" . haftmann@43943: qed haftmann@43943: haftmann@43943: lemma SUP_lessD: haftmann@43943: assumes "(\i\A. f i) < y" "i \ A" shows "f i < y" haftmann@43943: proof - haftmann@43943: have "f i \ (\i\A. f i)" using `i \ A` haftmann@44103: by (rule SUP_upper) haftmann@43943: also note `(\i\A. f i) < y` haftmann@43943: finally show "f i < y" . haftmann@43943: qed haftmann@43943: haftmann@43873: lemma INF_UNIV_bool_expand: haftmann@43868: "(\b. A b) = A True \ A False" huffman@44921: by (simp add: UNIV_bool INF_insert inf_commute) haftmann@43868: haftmann@43873: lemma SUP_UNIV_bool_expand: haftmann@43871: "(\b. A b) = A True \ A False" huffman@44921: by (simp add: UNIV_bool SUP_insert sup_commute) haftmann@43871: hoelzl@51328: lemma Inf_le_Sup: "A \ {} \ Inf A \ Sup A" hoelzl@51328: by (blast intro: Sup_upper2 Inf_lower ex_in_conv) hoelzl@51328: hoelzl@51328: lemma INF_le_SUP: "A \ {} \ INFI A f \ SUPR A f" hoelzl@51328: unfolding INF_def SUP_def by (rule Inf_le_Sup) auto hoelzl@51328: haftmann@32077: end haftmann@32077: haftmann@44024: class complete_distrib_lattice = complete_lattice + haftmann@44039: assumes sup_Inf: "a \ \B = (\b\B. a \ b)" haftmann@44024: assumes inf_Sup: "a \ \B = (\b\B. a \ b)" haftmann@44024: begin haftmann@44024: haftmann@44039: lemma sup_INF: haftmann@44039: "a \ (\b\B. f b) = (\b\B. a \ f b)" haftmann@44039: by (simp add: INF_def sup_Inf image_image) haftmann@44039: haftmann@44039: lemma inf_SUP: haftmann@44039: "a \ (\b\B. f b) = (\b\B. a \ f b)" haftmann@44039: by (simp add: SUP_def inf_Sup image_image) haftmann@44039: haftmann@44032: lemma dual_complete_distrib_lattice: krauss@44845: "class.complete_distrib_lattice Sup Inf sup (op \) (op >) inf \ \" haftmann@44024: apply (rule class.complete_distrib_lattice.intro) haftmann@44024: apply (fact dual_complete_lattice) haftmann@44024: apply (rule class.complete_distrib_lattice_axioms.intro) haftmann@44032: apply (simp_all only: INF_foundation_dual SUP_foundation_dual inf_Sup sup_Inf) haftmann@44032: done haftmann@44024: haftmann@44322: subclass distrib_lattice proof haftmann@44024: fix a b c haftmann@44024: from sup_Inf have "a \ \{b, c} = (\d\{b, c}. a \ d)" . noschinl@44919: then show "a \ b \ c = (a \ b) \ (a \ c)" by (simp add: INF_def) haftmann@44024: qed haftmann@44024: haftmann@44039: lemma Inf_sup: haftmann@44039: "\B \ a = (\b\B. b \ a)" haftmann@44039: by (simp add: sup_Inf sup_commute) haftmann@44039: haftmann@44039: lemma Sup_inf: haftmann@44039: "\B \ a = (\b\B. b \ a)" haftmann@44039: by (simp add: inf_Sup inf_commute) haftmann@44039: haftmann@44039: lemma INF_sup: haftmann@44039: "(\b\B. f b) \ a = (\b\B. f b \ a)" haftmann@44039: by (simp add: sup_INF sup_commute) haftmann@44039: haftmann@44039: lemma SUP_inf: haftmann@44039: "(\b\B. f b) \ a = (\b\B. f b \ a)" haftmann@44039: by (simp add: inf_SUP inf_commute) haftmann@44039: haftmann@44039: lemma Inf_sup_eq_top_iff: haftmann@44039: "(\B \ a = \) \ (\b\B. b \ a = \)" haftmann@44039: by (simp only: Inf_sup INF_top_conv) haftmann@44039: haftmann@44039: lemma Sup_inf_eq_bot_iff: haftmann@44039: "(\B \ a = \) \ (\b\B. b \ a = \)" haftmann@44039: by (simp only: Sup_inf SUP_bot_conv) haftmann@44039: haftmann@44039: lemma INF_sup_distrib2: haftmann@44039: "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" haftmann@44039: by (subst INF_commute) (simp add: sup_INF INF_sup) haftmann@44039: haftmann@44039: lemma SUP_inf_distrib2: haftmann@44039: "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" haftmann@44039: by (subst SUP_commute) (simp add: inf_SUP SUP_inf) haftmann@44039: haftmann@44024: end haftmann@44024: haftmann@44032: class complete_boolean_algebra = boolean_algebra + complete_distrib_lattice haftmann@43873: begin haftmann@43873: haftmann@43943: lemma dual_complete_boolean_algebra: krauss@44845: "class.complete_boolean_algebra Sup Inf sup (op \) (op >) inf \ \ (\x y. x \ - y) uminus" haftmann@44032: by (rule class.complete_boolean_algebra.intro, rule dual_complete_distrib_lattice, rule dual_boolean_algebra) haftmann@43943: haftmann@43873: lemma uminus_Inf: haftmann@43873: "- (\A) = \(uminus ` A)" haftmann@43873: proof (rule antisym) haftmann@43873: show "- \A \ \(uminus ` A)" haftmann@43873: by (rule compl_le_swap2, rule Inf_greatest, rule compl_le_swap2, rule Sup_upper) simp haftmann@43873: show "\(uminus ` A) \ - \A" haftmann@43873: by (rule Sup_least, rule compl_le_swap1, rule Inf_lower) auto haftmann@43873: qed haftmann@43873: haftmann@44041: lemma uminus_INF: "- (\x\A. B x) = (\x\A. - B x)" haftmann@44041: by (simp add: INF_def SUP_def uminus_Inf image_image) haftmann@44041: haftmann@43873: lemma uminus_Sup: haftmann@43873: "- (\A) = \(uminus ` A)" haftmann@43873: proof - haftmann@43873: have "\A = - \(uminus ` A)" by (simp add: image_image uminus_Inf) haftmann@43873: then show ?thesis by simp haftmann@43873: qed haftmann@43873: haftmann@43873: lemma uminus_SUP: "- (\x\A. B x) = (\x\A. - B x)" haftmann@43873: by (simp add: INF_def SUP_def uminus_Sup image_image) haftmann@43873: haftmann@43873: end haftmann@43873: haftmann@43940: class complete_linorder = linorder + complete_lattice haftmann@43940: begin haftmann@43940: haftmann@43943: lemma dual_complete_linorder: krauss@44845: "class.complete_linorder Sup Inf sup (op \) (op >) inf \ \" haftmann@43943: by (rule class.complete_linorder.intro, rule dual_complete_lattice, rule dual_linorder) haftmann@43943: haftmann@51386: lemma complete_linorder_inf_min: "inf = min" haftmann@51540: by (auto intro: antisym simp add: min_def fun_eq_iff) haftmann@51386: haftmann@51386: lemma complete_linorder_sup_max: "sup = max" haftmann@51540: by (auto intro: antisym simp add: max_def fun_eq_iff) haftmann@51386: noschinl@44918: lemma Inf_less_iff: haftmann@43940: "\S \ a \ (\x\S. x \ a)" haftmann@43940: unfolding not_le [symmetric] le_Inf_iff by auto haftmann@43940: noschinl@44918: lemma INF_less_iff: haftmann@44041: "(\i\A. f i) \ a \ (\x\A. f x \ a)" haftmann@44041: unfolding INF_def Inf_less_iff by auto haftmann@44041: noschinl@44918: lemma less_Sup_iff: haftmann@43940: "a \ \S \ (\x\S. a \ x)" haftmann@43940: unfolding not_le [symmetric] Sup_le_iff by auto haftmann@43940: noschinl@44918: lemma less_SUP_iff: haftmann@43940: "a \ (\i\A. f i) \ (\x\A. a \ f x)" haftmann@43940: unfolding SUP_def less_Sup_iff by auto haftmann@43940: noschinl@44918: lemma Sup_eq_top_iff [simp]: haftmann@43943: "\A = \ \ (\x<\. \i\A. x < i)" haftmann@43943: proof haftmann@43943: assume *: "\A = \" haftmann@43943: show "(\x<\. \i\A. x < i)" unfolding * [symmetric] haftmann@43943: proof (intro allI impI) haftmann@43943: fix x assume "x < \A" then show "\i\A. x < i" haftmann@43943: unfolding less_Sup_iff by auto haftmann@43943: qed haftmann@43943: next haftmann@43943: assume *: "\x<\. \i\A. x < i" haftmann@43943: show "\A = \" haftmann@43943: proof (rule ccontr) haftmann@43943: assume "\A \ \" haftmann@43943: with top_greatest [of "\A"] haftmann@43943: have "\A < \" unfolding le_less by auto haftmann@43943: then have "\A < \A" haftmann@43943: using * unfolding less_Sup_iff by auto haftmann@43943: then show False by auto haftmann@43943: qed haftmann@43943: qed haftmann@43943: noschinl@44918: lemma SUP_eq_top_iff [simp]: haftmann@44041: "(\i\A. f i) = \ \ (\x<\. \i\A. x < f i)" noschinl@44919: unfolding SUP_def by auto haftmann@44041: noschinl@44918: lemma Inf_eq_bot_iff [simp]: haftmann@43943: "\A = \ \ (\x>\. \i\A. i < x)" huffman@44920: using dual_complete_linorder huffman@44920: by (rule complete_linorder.Sup_eq_top_iff) haftmann@43943: noschinl@44918: lemma INF_eq_bot_iff [simp]: haftmann@43967: "(\i\A. f i) = \ \ (\x>\. \i\A. f i < x)" noschinl@44919: unfolding INF_def by auto haftmann@43967: hoelzl@51328: lemma le_Sup_iff: "x \ \A \ (\ya\A. y < a)" hoelzl@51328: proof safe hoelzl@51328: fix y assume "x \ \A" "y < x" hoelzl@51328: then have "y < \A" by auto hoelzl@51328: then show "\a\A. y < a" hoelzl@51328: unfolding less_Sup_iff . hoelzl@51328: qed (auto elim!: allE[of _ "\A"] simp add: not_le[symmetric] Sup_upper) hoelzl@51328: hoelzl@51328: lemma le_SUP_iff: "x \ SUPR A f \ (\yi\A. y < f i)" hoelzl@51328: unfolding le_Sup_iff SUP_def by simp hoelzl@51328: hoelzl@51328: lemma Inf_le_iff: "\A \ x \ (\y>x. \a\A. y > a)" hoelzl@51328: proof safe hoelzl@51328: fix y assume "x \ \A" "y > x" hoelzl@51328: then have "y > \A" by auto hoelzl@51328: then show "\a\A. y > a" hoelzl@51328: unfolding Inf_less_iff . hoelzl@51328: qed (auto elim!: allE[of _ "\A"] simp add: not_le[symmetric] Inf_lower) hoelzl@51328: hoelzl@51328: lemma INF_le_iff: hoelzl@51328: "INFI A f \ x \ (\y>x. \i\A. y > f i)" hoelzl@51328: unfolding Inf_le_iff INF_def by simp hoelzl@51328: haftmann@51386: subclass complete_distrib_lattice haftmann@51386: proof haftmann@51386: fix a and B haftmann@51386: show "a \ \B = (\b\B. a \ b)" and "a \ \B = (\b\B. a \ b)" haftmann@51386: by (safe intro!: INF_eqI [symmetric] sup_mono Inf_lower SUP_eqI [symmetric] inf_mono Sup_upper) haftmann@51386: (auto simp: not_less [symmetric] Inf_less_iff less_Sup_iff haftmann@51386: le_max_iff_disj complete_linorder_sup_max min_le_iff_disj complete_linorder_inf_min) haftmann@51386: qed haftmann@51386: haftmann@43940: end haftmann@43940: hoelzl@51341: haftmann@46631: subsection {* Complete lattice on @{typ bool} *} haftmann@32077: haftmann@44024: instantiation bool :: complete_lattice haftmann@32077: begin haftmann@32077: haftmann@32077: definition haftmann@46154: [simp, code]: "\A \ False \ A" haftmann@32077: haftmann@32077: definition haftmann@46154: [simp, code]: "\A \ True \ A" haftmann@32077: haftmann@32077: instance proof haftmann@44322: qed (auto intro: bool_induct) haftmann@32077: haftmann@32077: end haftmann@32077: haftmann@49905: lemma not_False_in_image_Ball [simp]: haftmann@49905: "False \ P ` A \ Ball A P" haftmann@49905: by auto haftmann@49905: haftmann@49905: lemma True_in_image_Bex [simp]: haftmann@49905: "True \ P ` A \ Bex A P" haftmann@49905: by auto haftmann@49905: haftmann@43873: lemma INF_bool_eq [simp]: haftmann@32120: "INFI = Ball" haftmann@49905: by (simp add: fun_eq_iff INF_def) haftmann@32120: haftmann@43873: lemma SUP_bool_eq [simp]: haftmann@32120: "SUPR = Bex" haftmann@49905: by (simp add: fun_eq_iff SUP_def) haftmann@32120: haftmann@44032: instance bool :: complete_boolean_algebra proof haftmann@44322: qed (auto intro: bool_induct) haftmann@44024: haftmann@46631: haftmann@46631: subsection {* Complete lattice on @{typ "_ \ _"} *} haftmann@46631: haftmann@32077: instantiation "fun" :: (type, complete_lattice) complete_lattice haftmann@32077: begin haftmann@32077: haftmann@32077: definition haftmann@44024: "\A = (\x. \f\A. f x)" haftmann@41080: noschinl@46882: lemma Inf_apply [simp, code]: haftmann@44024: "(\A) x = (\f\A. f x)" haftmann@41080: by (simp add: Inf_fun_def) haftmann@32077: haftmann@32077: definition haftmann@44024: "\A = (\x. \f\A. f x)" haftmann@41080: noschinl@46882: lemma Sup_apply [simp, code]: haftmann@44024: "(\A) x = (\f\A. f x)" haftmann@41080: by (simp add: Sup_fun_def) haftmann@32077: haftmann@32077: instance proof noschinl@46884: qed (auto simp add: le_fun_def intro: INF_lower INF_greatest SUP_upper SUP_least) haftmann@32077: haftmann@32077: end haftmann@32077: noschinl@46882: lemma INF_apply [simp]: haftmann@41080: "(\y\A. f y) x = (\y\A. f y x)" noschinl@46884: by (auto intro: arg_cong [of _ _ Inf] simp add: INF_def) hoelzl@38705: noschinl@46882: lemma SUP_apply [simp]: haftmann@41080: "(\y\A. f y) x = (\y\A. f y x)" noschinl@46884: by (auto intro: arg_cong [of _ _ Sup] simp add: SUP_def) haftmann@32077: haftmann@44024: instance "fun" :: (type, complete_distrib_lattice) complete_distrib_lattice proof noschinl@46884: qed (auto simp add: INF_def SUP_def inf_Sup sup_Inf image_image) haftmann@44024: haftmann@43873: instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra .. haftmann@43873: haftmann@46631: haftmann@46631: subsection {* Complete lattice on unary and binary predicates *} haftmann@46631: haftmann@46631: lemma INF1_iff: "(\x\A. B x) b = (\x\A. B x b)" noschinl@46884: by simp haftmann@46631: haftmann@46631: lemma INF2_iff: "(\x\A. B x) b c = (\x\A. B x b c)" noschinl@46884: by simp haftmann@46631: haftmann@46631: lemma INF1_I: "(\x. x \ A \ B x b) \ (\x\A. B x) b" noschinl@46884: by auto haftmann@46631: haftmann@46631: lemma INF2_I: "(\x. x \ A \ B x b c) \ (\x\A. B x) b c" noschinl@46884: by auto haftmann@46631: haftmann@46631: lemma INF1_D: "(\x\A. B x) b \ a \ A \ B a b" noschinl@46884: by auto haftmann@46631: haftmann@46631: lemma INF2_D: "(\x\A. B x) b c \ a \ A \ B a b c" noschinl@46884: by auto haftmann@46631: haftmann@46631: lemma INF1_E: "(\x\A. B x) b \ (B a b \ R) \ (a \ A \ R) \ R" noschinl@46884: by auto haftmann@46631: haftmann@46631: lemma INF2_E: "(\x\A. B x) b c \ (B a b c \ R) \ (a \ A \ R) \ R" noschinl@46884: by auto haftmann@46631: haftmann@46631: lemma SUP1_iff: "(\x\A. B x) b = (\x\A. B x b)" noschinl@46884: by simp haftmann@46631: haftmann@46631: lemma SUP2_iff: "(\x\A. B x) b c = (\x\A. B x b c)" noschinl@46884: by simp haftmann@46631: haftmann@46631: lemma SUP1_I: "a \ A \ B a b \ (\x\A. B x) b" noschinl@46884: by auto haftmann@46631: haftmann@46631: lemma SUP2_I: "a \ A \ B a b c \ (\x\A. B x) b c" noschinl@46884: by auto haftmann@46631: haftmann@46631: lemma SUP1_E: "(\x\A. B x) b \ (\x. x \ A \ B x b \ R) \ R" noschinl@46884: by auto haftmann@46631: haftmann@46631: lemma SUP2_E: "(\x\A. B x) b c \ (\x. x \ A \ B x b c \ R) \ R" noschinl@46884: by auto haftmann@46631: haftmann@46631: haftmann@46631: subsection {* Complete lattice on @{typ "_ set"} *} haftmann@46631: haftmann@45960: instantiation "set" :: (type) complete_lattice haftmann@45960: begin haftmann@45960: haftmann@45960: definition haftmann@45960: "\A = {x. \((\B. x \ B) ` A)}" haftmann@45960: haftmann@45960: definition haftmann@45960: "\A = {x. \((\B. x \ B) ` A)}" haftmann@45960: haftmann@45960: instance proof haftmann@51386: qed (auto simp add: less_eq_set_def Inf_set_def Sup_set_def le_fun_def) haftmann@45960: haftmann@45960: end haftmann@45960: haftmann@45960: instance "set" :: (type) complete_boolean_algebra haftmann@45960: proof haftmann@45960: qed (auto simp add: INF_def SUP_def Inf_set_def Sup_set_def image_def) haftmann@45960: haftmann@32077: haftmann@46631: subsubsection {* Inter *} haftmann@41082: haftmann@41082: abbreviation Inter :: "'a set set \ 'a set" where haftmann@41082: "Inter S \ \S" haftmann@41082: haftmann@41082: notation (xsymbols) haftmann@52141: Inter ("\_" [900] 900) haftmann@41082: haftmann@41082: lemma Inter_eq: haftmann@41082: "\A = {x. \B \ A. x \ B}" haftmann@41082: proof (rule set_eqI) haftmann@41082: fix x haftmann@41082: have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" haftmann@41082: by auto haftmann@41082: then show "x \ \A \ x \ {x. \B \ A. x \ B}" haftmann@45960: by (simp add: Inf_set_def image_def) haftmann@41082: qed haftmann@41082: blanchet@54147: lemma Inter_iff [simp]: "A \ \C \ (\X\C. A \ X)" haftmann@41082: by (unfold Inter_eq) blast haftmann@41082: haftmann@43741: lemma InterI [intro!]: "(\X. X \ C \ A \ X) \ A \ \C" haftmann@41082: by (simp add: Inter_eq) haftmann@41082: haftmann@41082: text {* haftmann@41082: \medskip A ``destruct'' rule -- every @{term X} in @{term C} haftmann@43741: contains @{term A} as an element, but @{prop "A \ X"} can hold when haftmann@43741: @{prop "X \ C"} does not! This rule is analogous to @{text spec}. haftmann@41082: *} haftmann@41082: haftmann@43741: lemma InterD [elim, Pure.elim]: "A \ \C \ X \ C \ A \ X" haftmann@41082: by auto haftmann@41082: haftmann@43741: lemma InterE [elim]: "A \ \C \ (X \ C \ R) \ (A \ X \ R) \ R" haftmann@41082: -- {* ``Classical'' elimination rule -- does not require proving haftmann@43741: @{prop "X \ C"}. *} haftmann@41082: by (unfold Inter_eq) blast haftmann@41082: haftmann@43741: lemma Inter_lower: "B \ A \ \A \ B" haftmann@43740: by (fact Inf_lower) haftmann@43740: haftmann@41082: lemma Inter_subset: haftmann@43755: "(\X. X \ A \ X \ B) \ A \ {} \ \A \ B" haftmann@43740: by (fact Inf_less_eq) haftmann@41082: haftmann@43755: lemma Inter_greatest: "(\X. X \ A \ C \ X) \ C \ Inter A" haftmann@43740: by (fact Inf_greatest) haftmann@41082: huffman@44067: lemma Inter_empty: "\{} = UNIV" huffman@44067: by (fact Inf_empty) (* already simp *) haftmann@41082: huffman@44067: lemma Inter_UNIV: "\UNIV = {}" huffman@44067: by (fact Inf_UNIV) (* already simp *) haftmann@41082: huffman@44920: lemma Inter_insert: "\(insert a B) = a \ \B" huffman@44920: by (fact Inf_insert) (* already simp *) haftmann@41082: haftmann@41082: lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" haftmann@43899: by (fact less_eq_Inf_inter) haftmann@41082: haftmann@41082: lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" haftmann@43756: by (fact Inf_union_distrib) haftmann@43756: blanchet@54147: lemma Inter_UNIV_conv [simp]: haftmann@43741: "\A = UNIV \ (\x\A. x = UNIV)" haftmann@43741: "UNIV = \A \ (\x\A. x = UNIV)" haftmann@43801: by (fact Inf_top_conv)+ haftmann@41082: haftmann@43741: lemma Inter_anti_mono: "B \ A \ \A \ \B" haftmann@43899: by (fact Inf_superset_mono) haftmann@41082: haftmann@41082: haftmann@46631: subsubsection {* Intersections of families *} haftmann@41082: haftmann@41082: abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" where haftmann@41082: "INTER \ INFI" haftmann@41082: haftmann@43872: text {* haftmann@43872: Note: must use name @{const INTER} here instead of @{text INT} haftmann@43872: to allow the following syntax coexist with the plain constant name. haftmann@43872: *} haftmann@43872: haftmann@41082: syntax haftmann@41082: "_INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) haftmann@41082: "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 0, 10] 10) haftmann@41082: haftmann@41082: syntax (xsymbols) haftmann@41082: "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) haftmann@41082: "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@41082: haftmann@41082: syntax (latex output) haftmann@41082: "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) haftmann@41082: "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) haftmann@41082: haftmann@41082: translations haftmann@41082: "INT x y. B" == "INT x. INT y. B" haftmann@41082: "INT x. B" == "CONST INTER CONST UNIV (%x. B)" haftmann@41082: "INT x. B" == "INT x:CONST UNIV. B" haftmann@41082: "INT x:A. B" == "CONST INTER A (%x. B)" haftmann@41082: haftmann@41082: print_translation {* wenzelm@42284: [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}] haftmann@41082: *} -- {* to avoid eta-contraction of body *} haftmann@41082: haftmann@44085: lemma INTER_eq: haftmann@41082: "(\x\A. B x) = {y. \x\A. y \ B x}" haftmann@44085: by (auto simp add: INF_def) haftmann@41082: haftmann@41082: lemma Inter_image_eq [simp]: haftmann@41082: "\(B`A) = (\x\A. B x)" haftmann@43872: by (rule sym) (fact INF_def) haftmann@41082: haftmann@43817: lemma INT_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" haftmann@44085: by (auto simp add: INF_def image_def) haftmann@41082: haftmann@43817: lemma INT_I [intro!]: "(\x. x \ A \ b \ B x) \ b \ (\x\A. B x)" haftmann@44085: by (auto simp add: INF_def image_def) haftmann@41082: haftmann@43852: lemma INT_D [elim, Pure.elim]: "b \ (\x\A. B x) \ a \ A \ b \ B a" haftmann@41082: by auto haftmann@41082: haftmann@43852: lemma INT_E [elim]: "b \ (\x\A. B x) \ (b \ B a \ R) \ (a \ A \ R) \ R" haftmann@43852: -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a\A"}. *} haftmann@44085: by (auto simp add: INF_def image_def) haftmann@41082: haftmann@41082: lemma INT_cong [cong]: haftmann@43854: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" haftmann@43865: by (fact INF_cong) haftmann@41082: haftmann@41082: lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" haftmann@41082: by blast haftmann@41082: haftmann@41082: lemma Collect_all_eq: "{x. \y. P x y} = (\y. {x. P x y})" haftmann@41082: by blast haftmann@41082: haftmann@43817: lemma INT_lower: "a \ A \ (\x\A. B x) \ B a" haftmann@44103: by (fact INF_lower) haftmann@41082: haftmann@43817: lemma INT_greatest: "(\x. x \ A \ C \ B x) \ C \ (\x\A. B x)" haftmann@44103: by (fact INF_greatest) haftmann@41082: huffman@44067: lemma INT_empty: "(\x\{}. B x) = UNIV" haftmann@44085: by (fact INF_empty) haftmann@43854: haftmann@43817: lemma INT_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" haftmann@43872: by (fact INF_absorb) haftmann@41082: haftmann@43854: lemma INT_subset_iff: "B \ (\i\I. A i) \ (\i\I. B \ A i)" haftmann@41082: by (fact le_INF_iff) haftmann@41082: haftmann@41082: lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" haftmann@43865: by (fact INF_insert) haftmann@43865: haftmann@43865: lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" haftmann@43865: by (fact INF_union) haftmann@43865: haftmann@43865: lemma INT_insert_distrib: haftmann@43865: "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" haftmann@43865: by blast haftmann@43854: haftmann@41082: lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" haftmann@43865: by (fact INF_constant) haftmann@43865: huffman@44920: lemma INTER_UNIV_conv: haftmann@43817: "(UNIV = (\x\A. B x)) = (\x\A. B x = UNIV)" haftmann@43817: "((\x\A. B x) = UNIV) = (\x\A. B x = UNIV)" huffman@44920: by (fact INF_top_conv)+ (* already simp *) haftmann@43865: haftmann@43865: lemma INT_bool_eq: "(\b. A b) = A True \ A False" haftmann@43873: by (fact INF_UNIV_bool_expand) haftmann@43865: haftmann@43865: lemma INT_anti_mono: haftmann@43900: "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\A. g x)" haftmann@43865: -- {* The last inclusion is POSITIVE! *} haftmann@43940: by (fact INF_superset_mono) haftmann@41082: haftmann@41082: lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" haftmann@41082: by blast haftmann@41082: haftmann@43817: lemma vimage_INT: "f -` (\x\A. B x) = (\x\A. f -` B x)" haftmann@41082: by blast haftmann@41082: haftmann@41082: haftmann@46631: subsubsection {* Union *} haftmann@32115: haftmann@32587: abbreviation Union :: "'a set set \ 'a set" where haftmann@32587: "Union S \ \S" haftmann@32115: haftmann@32115: notation (xsymbols) haftmann@52141: Union ("\_" [900] 900) haftmann@32115: haftmann@32135: lemma Union_eq: haftmann@32135: "\A = {x. \B \ A. x \ B}" nipkow@39302: proof (rule set_eqI) haftmann@32115: fix x haftmann@32135: have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" haftmann@32115: by auto haftmann@32135: then show "x \ \A \ x \ {x. \B\A. x \ B}" haftmann@45960: by (simp add: Sup_set_def image_def) haftmann@32115: qed haftmann@32115: blanchet@54147: lemma Union_iff [simp]: haftmann@32115: "A \ \C \ (\X\C. A\X)" haftmann@32115: by (unfold Union_eq) blast haftmann@32115: haftmann@32115: lemma UnionI [intro]: haftmann@32115: "X \ C \ A \ X \ A \ \C" haftmann@32115: -- {* The order of the premises presupposes that @{term C} is rigid; haftmann@32115: @{term A} may be flexible. *} haftmann@32115: by auto haftmann@32115: haftmann@32115: lemma UnionE [elim!]: haftmann@43817: "A \ \C \ (\X. A \ X \ X \ C \ R) \ R" haftmann@32115: by auto haftmann@32115: haftmann@43817: lemma Union_upper: "B \ A \ B \ \A" haftmann@43901: by (fact Sup_upper) haftmann@32135: haftmann@43817: lemma Union_least: "(\X. X \ A \ X \ C) \ \A \ C" haftmann@43901: by (fact Sup_least) haftmann@32135: huffman@44920: lemma Union_empty: "\{} = {}" huffman@44920: by (fact Sup_empty) (* already simp *) haftmann@32135: huffman@44920: lemma Union_UNIV: "\UNIV = UNIV" huffman@44920: by (fact Sup_UNIV) (* already simp *) haftmann@32135: huffman@44920: lemma Union_insert: "\insert a B = a \ \B" huffman@44920: by (fact Sup_insert) (* already simp *) haftmann@32135: haftmann@43817: lemma Union_Un_distrib [simp]: "\(A \ B) = \A \ \B" haftmann@43901: by (fact Sup_union_distrib) haftmann@32135: haftmann@32135: lemma Union_Int_subset: "\(A \ B) \ \A \ \B" haftmann@43901: by (fact Sup_inter_less_eq) haftmann@32135: blanchet@54147: lemma Union_empty_conv: "(\A = {}) \ (\x\A. x = {})" huffman@44920: by (fact Sup_bot_conv) (* already simp *) haftmann@32135: blanchet@54147: lemma empty_Union_conv: "({} = \A) \ (\x\A. x = {})" huffman@44920: by (fact Sup_bot_conv) (* already simp *) haftmann@32135: haftmann@32135: lemma subset_Pow_Union: "A \ Pow (\A)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Union_Pow_eq [simp]: "\(Pow A) = A" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma Union_mono: "A \ B \ \A \ \B" haftmann@43901: by (fact Sup_subset_mono) haftmann@32135: haftmann@32115: haftmann@46631: subsubsection {* Unions of families *} haftmann@32077: haftmann@32606: abbreviation UNION :: "'a set \ ('a \ 'b set) \ 'b set" where haftmann@32606: "UNION \ SUPR" haftmann@32077: haftmann@43872: text {* haftmann@43872: Note: must use name @{const UNION} here instead of @{text UN} haftmann@43872: to allow the following syntax coexist with the plain constant name. haftmann@43872: *} haftmann@43872: haftmann@32077: syntax wenzelm@35115: "_UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) huffman@36364: "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 0, 10] 10) haftmann@32077: haftmann@32077: syntax (xsymbols) wenzelm@35115: "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) huffman@36364: "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@32077: haftmann@32077: syntax (latex output) wenzelm@35115: "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) huffman@36364: "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) haftmann@32077: haftmann@32077: translations haftmann@32077: "UN x y. B" == "UN x. UN y. B" haftmann@32077: "UN x. B" == "CONST UNION CONST UNIV (%x. B)" haftmann@32077: "UN x. B" == "UN x:CONST UNIV. B" haftmann@32077: "UN x:A. B" == "CONST UNION A (%x. B)" haftmann@32077: haftmann@32077: text {* haftmann@32077: Note the difference between ordinary xsymbol syntax of indexed wenzelm@53015: unions and intersections (e.g.\ @{text"\a\<^sub>1\A\<^sub>1. B"}) wenzelm@53015: and their \LaTeX\ rendition: @{term"\a\<^sub>1\A\<^sub>1. B"}. The haftmann@32077: former does not make the index expression a subscript of the haftmann@32077: union/intersection symbol because this leads to problems with nested haftmann@32077: subscripts in Proof General. haftmann@32077: *} haftmann@32077: wenzelm@35115: print_translation {* wenzelm@42284: [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax UNION} @{syntax_const "_UNION"}] wenzelm@35115: *} -- {* to avoid eta-contraction of body *} haftmann@32077: blanchet@54147: lemma UNION_eq: haftmann@32135: "(\x\A. B x) = {y. \x\A. y \ B x}" haftmann@44085: by (auto simp add: SUP_def) huffman@44920: haftmann@45960: lemma bind_UNION [code]: haftmann@45960: "Set.bind A f = UNION A f" haftmann@45960: by (simp add: bind_def UNION_eq) haftmann@45960: haftmann@46036: lemma member_bind [simp]: haftmann@46036: "x \ Set.bind P f \ x \ UNION P f " haftmann@46036: by (simp add: bind_UNION) haftmann@46036: haftmann@32115: lemma Union_image_eq [simp]: haftmann@43817: "\(B ` A) = (\x\A. B x)" huffman@44920: by (rule sym) (fact SUP_def) huffman@44920: haftmann@46036: lemma UN_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" haftmann@44085: by (auto simp add: SUP_def image_def) wenzelm@11979: haftmann@43852: lemma UN_I [intro]: "a \ A \ b \ B a \ b \ (\x\A. B x)" wenzelm@11979: -- {* The order of the premises presupposes that @{term A} is rigid; wenzelm@11979: @{term b} may be flexible. *} wenzelm@11979: by auto wenzelm@11979: haftmann@43852: lemma UN_E [elim!]: "b \ (\x\A. B x) \ (\x. x\A \ b \ B x \ R) \ R" haftmann@44085: by (auto simp add: SUP_def image_def) clasohm@923: wenzelm@11979: lemma UN_cong [cong]: haftmann@43900: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" haftmann@43900: by (fact SUP_cong) wenzelm@11979: berghofe@29691: lemma strong_UN_cong: haftmann@43900: "A = B \ (\x. x \ B =simp=> C x = D x) \ (\x\A. C x) = (\x\B. D x)" haftmann@43900: by (unfold simp_implies_def) (fact UN_cong) berghofe@29691: haftmann@43817: lemma image_eq_UN: "f ` A = (\x\A. {f x})" haftmann@32077: by blast haftmann@32077: haftmann@43817: lemma UN_upper: "a \ A \ B a \ (\x\A. B x)" haftmann@44103: by (fact SUP_upper) haftmann@32135: haftmann@43817: lemma UN_least: "(\x. x \ A \ B x \ C) \ (\x\A. B x) \ C" haftmann@44103: by (fact SUP_least) haftmann@32135: blanchet@54147: lemma Collect_bex_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma UN_insert_distrib: "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" haftmann@32135: by blast haftmann@32135: blanchet@54147: lemma UN_empty: "(\x\{}. B x) = {}" haftmann@44085: by (fact SUP_empty) haftmann@32135: huffman@44920: lemma UN_empty2: "(\x\A. {}) = {}" huffman@44920: by (fact SUP_bot) (* already simp *) haftmann@32135: haftmann@43817: lemma UN_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" haftmann@43900: by (fact SUP_absorb) haftmann@32135: haftmann@32135: lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" haftmann@43900: by (fact SUP_insert) haftmann@32135: haftmann@44085: lemma UN_Un [simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" haftmann@43900: by (fact SUP_union) haftmann@32135: haftmann@43967: lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" huffman@35629: by (fact SUP_le_iff) haftmann@32135: haftmann@32135: lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" haftmann@43900: by (fact SUP_constant) haftmann@32135: haftmann@43944: lemma image_Union: "f ` \S = (\x\S. f ` x)" haftmann@32135: by blast haftmann@32135: huffman@44920: lemma UNION_empty_conv: haftmann@43817: "{} = (\x\A. B x) \ (\x\A. B x = {})" haftmann@43817: "(\x\A. B x) = {} \ (\x\A. B x = {})" huffman@44920: by (fact SUP_bot_conv)+ (* already simp *) haftmann@32135: blanchet@54147: lemma Collect_ex_eq: "{x. \y. P x y} = (\y. {x. P x y})" haftmann@32135: by blast haftmann@32135: haftmann@43900: lemma ball_UN: "(\z \ UNION A B. P z) \ (\x\A. \z \ B x. P z)" haftmann@32135: by blast haftmann@32135: haftmann@43900: lemma bex_UN: "(\z \ UNION A B. P z) \ (\x\A. \z\B x. P z)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" haftmann@32135: by (auto simp add: split_if_mem2) haftmann@32135: haftmann@43817: lemma UN_bool_eq: "(\b. A b) = (A True \ A False)" haftmann@43900: by (fact SUP_UNIV_bool_expand) haftmann@32135: haftmann@32135: lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma UN_mono: haftmann@43817: "A \ B \ (\x. x \ A \ f x \ g x) \ haftmann@32135: (\x\A. f x) \ (\x\B. g x)" haftmann@43940: by (fact SUP_subset_mono) haftmann@32135: haftmann@43817: lemma vimage_Union: "f -` (\A) = (\X\A. f -` X)" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma vimage_UN: "f -` (\x\A. B x) = (\x\A. f -` B x)" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma vimage_eq_UN: "f -` B = (\y\B. f -` {y})" haftmann@32135: -- {* NOT suitable for rewriting *} haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma image_UN: "f ` UNION A B = (\x\A. f ` B x)" haftmann@43817: by blast haftmann@32135: haftmann@45013: lemma UN_singleton [simp]: "(\x\A. {x}) = A" haftmann@45013: by blast haftmann@45013: wenzelm@11979: haftmann@46631: subsubsection {* Distributive laws *} wenzelm@12897: wenzelm@12897: lemma Int_Union: "A \ \B = (\C\B. A \ C)" haftmann@44032: by (fact inf_Sup) wenzelm@12897: haftmann@44039: lemma Un_Inter: "A \ \B = (\C\B. A \ C)" haftmann@44039: by (fact sup_Inf) haftmann@44039: wenzelm@12897: lemma Int_Union2: "\B \ A = (\C\B. C \ A)" haftmann@44039: by (fact Sup_inf) haftmann@44039: haftmann@44039: lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" haftmann@44039: by (rule sym) (rule INF_inf_distrib) haftmann@44039: haftmann@44039: lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" haftmann@44039: by (rule sym) (rule SUP_sup_distrib) haftmann@44039: haftmann@44039: lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" haftmann@44039: by (simp only: INT_Int_distrib INF_def) wenzelm@12897: haftmann@43817: lemma Un_Union_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" wenzelm@12897: -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *} wenzelm@12897: -- {* Union of a family of unions *} haftmann@44039: by (simp only: UN_Un_distrib SUP_def) wenzelm@12897: haftmann@44039: lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" haftmann@44039: by (fact sup_INF) wenzelm@12897: wenzelm@12897: lemma Int_UN_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" wenzelm@12897: -- {* Halmos, Naive Set Theory, page 35. *} haftmann@44039: by (fact inf_SUP) wenzelm@12897: wenzelm@12897: lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" haftmann@44039: by (fact SUP_inf_distrib2) wenzelm@12897: wenzelm@12897: lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" haftmann@44039: by (fact INF_sup_distrib2) haftmann@44039: haftmann@44039: lemma Union_disjoint: "(\C \ A = {}) \ (\B\C. B \ A = {})" haftmann@44039: by (fact Sup_inf_eq_bot_iff) wenzelm@12897: wenzelm@12897: haftmann@46631: subsubsection {* Complement *} haftmann@32135: haftmann@43873: lemma Compl_INT [simp]: "- (\x\A. B x) = (\x\A. -B x)" haftmann@43873: by (fact uminus_INF) wenzelm@12897: haftmann@43873: lemma Compl_UN [simp]: "- (\x\A. B x) = (\x\A. -B x)" haftmann@43873: by (fact uminus_SUP) wenzelm@12897: wenzelm@12897: haftmann@46631: subsubsection {* Miniscoping and maxiscoping *} wenzelm@12897: paulson@13860: text {* \medskip Miniscoping: pushing in quantifiers and big Unions paulson@13860: and Intersections. *} wenzelm@12897: wenzelm@12897: lemma UN_simps [simp]: haftmann@43817: "\a B C. (\x\C. insert a (B x)) = (if C={} then {} else insert a (\x\C. B x))" haftmann@44032: "\A B C. (\x\C. A x \ B) = ((if C={} then {} else (\x\C. A x) \ B))" haftmann@43852: "\A B C. (\x\C. A \ B x) = ((if C={} then {} else A \ (\x\C. B x)))" haftmann@44032: "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \ B)" haftmann@43852: "\A B C. (\x\C. A \ B x) = (A \(\x\C. B x))" haftmann@43852: "\A B C. (\x\C. A x - B) = ((\x\C. A x) - B)" haftmann@43852: "\A B C. (\x\C. A - B x) = (A - (\x\C. B x))" haftmann@43852: "\A B. (\x\\A. B x) = (\y\A. \x\y. B x)" haftmann@43852: "\A B C. (\z\UNION A B. C z) = (\x\A. \z\B x. C z)" haftmann@43831: "\A B f. (\x\f`A. B x) = (\a\A. B (f a))" wenzelm@12897: by auto wenzelm@12897: wenzelm@12897: lemma INT_simps [simp]: haftmann@44032: "\A B C. (\x\C. A x \ B) = (if C={} then UNIV else (\x\C. A x) \ B)" haftmann@43831: "\A B C. (\x\C. A \ B x) = (if C={} then UNIV else A \(\x\C. B x))" haftmann@43852: "\A B C. (\x\C. A x - B) = (if C={} then UNIV else (\x\C. A x) - B)" haftmann@43852: "\A B C. (\x\C. A - B x) = (if C={} then UNIV else A - (\x\C. B x))" haftmann@43817: "\a B C. (\x\C. insert a (B x)) = insert a (\x\C. B x)" haftmann@43852: "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \ B)" haftmann@43852: "\A B C. (\x\C. A \ B x) = (A \ (\x\C. B x))" haftmann@43852: "\A B. (\x\\A. B x) = (\y\A. \x\y. B x)" haftmann@43852: "\A B C. (\z\UNION A B. C z) = (\x\A. \z\B x. C z)" haftmann@43852: "\A B f. (\x\f`A. B x) = (\a\A. B (f a))" wenzelm@12897: by auto wenzelm@12897: blanchet@54147: lemma UN_ball_bex_simps [simp]: haftmann@43852: "\A P. (\x\\A. P x) \ (\y\A. \x\y. P x)" haftmann@43967: "\A B P. (\x\UNION A B. P x) = (\a\A. \x\ B a. P x)" haftmann@43852: "\A P. (\x\\A. P x) \ (\y\A. \x\y. P x)" haftmann@43852: "\A B P. (\x\UNION A B. P x) \ (\a\A. \x\B a. P x)" wenzelm@12897: by auto wenzelm@12897: haftmann@43943: paulson@13860: text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *} paulson@13860: paulson@13860: lemma UN_extend_simps: haftmann@43817: "\a B C. insert a (\x\C. B x) = (if C={} then {a} else (\x\C. insert a (B x)))" haftmann@44032: "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" haftmann@43852: "\A B C. A \ (\x\C. B x) = (if C={} then A else (\x\C. A \ B x))" haftmann@43852: "\A B C. ((\x\C. A x) \ B) = (\x\C. A x \ B)" haftmann@43852: "\A B C. (A \ (\x\C. B x)) = (\x\C. A \ B x)" haftmann@43817: "\A B C. ((\x\C. A x) - B) = (\x\C. A x - B)" haftmann@43817: "\A B C. (A - (\x\C. B x)) = (\x\C. A - B x)" haftmann@43852: "\A B. (\y\A. \x\y. B x) = (\x\\A. B x)" haftmann@43852: "\A B C. (\x\A. \z\B x. C z) = (\z\UNION A B. C z)" haftmann@43831: "\A B f. (\a\A. B (f a)) = (\x\f`A. B x)" paulson@13860: by auto paulson@13860: paulson@13860: lemma INT_extend_simps: haftmann@43852: "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" haftmann@43852: "\A B C. A \ (\x\C. B x) = (if C={} then A else (\x\C. A \ B x))" haftmann@43852: "\A B C. (\x\C. A x) - B = (if C={} then UNIV - B else (\x\C. A x - B))" haftmann@43852: "\A B C. A - (\x\C. B x) = (if C={} then A else (\x\C. A - B x))" haftmann@43817: "\a B C. insert a (\x\C. B x) = (\x\C. insert a (B x))" haftmann@43852: "\A B C. ((\x\C. A x) \ B) = (\x\C. A x \ B)" haftmann@43852: "\A B C. A \ (\x\C. B x) = (\x\C. A \ B x)" haftmann@43852: "\A B. (\y\A. \x\y. B x) = (\x\\A. B x)" haftmann@43852: "\A B C. (\x\A. \z\B x. C z) = (\z\UNION A B. C z)" haftmann@43852: "\A B f. (\a\A. B (f a)) = (\x\f`A. B x)" paulson@13860: by auto paulson@13860: haftmann@43872: text {* Finally *} haftmann@43872: haftmann@32135: no_notation haftmann@46691: less_eq (infix "\" 50) and haftmann@46691: less (infix "\" 50) haftmann@32135: haftmann@30596: lemmas mem_simps = haftmann@30596: insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff haftmann@30596: mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff haftmann@30596: -- {* Each of these has ALREADY been added @{text "[simp]"} above. *} wenzelm@21669: wenzelm@11979: end haftmann@49905: