haftmann@32139: (* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel; Florian Haftmann, TU Muenchen *) wenzelm@11979: haftmann@32139: header {* Complete lattices, with special focus on sets *} haftmann@32077: haftmann@32139: theory Complete_Lattice haftmann@32139: imports Set haftmann@32139: begin haftmann@32077: haftmann@32077: notation haftmann@34007: less_eq (infix "\" 50) and haftmann@32077: less (infix "\" 50) and haftmann@34007: inf (infixl "\" 70) and haftmann@34007: sup (infixl "\" 65) and haftmann@32678: top ("\") and haftmann@32678: bot ("\") 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@32139: subsection {* Abstract complete lattices *} haftmann@32139: haftmann@34007: class complete_lattice = bounded_lattice + Inf + Sup + 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@32077: begin haftmann@32077: haftmann@32678: lemma dual_complete_lattice: haftmann@43868: "class.complete_lattice Sup Inf (op \) (op >) sup inf \ \" haftmann@36635: by (auto intro!: class.complete_lattice.intro dual_bounded_lattice) haftmann@34007: (unfold_locales, (fact bot_least top_greatest haftmann@34007: Sup_upper Sup_least Inf_lower Inf_greatest)+) haftmann@32678: haftmann@34007: lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" haftmann@32077: by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) haftmann@32077: haftmann@34007: lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" haftmann@32077: by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) haftmann@32077: haftmann@41080: lemma Inf_empty [simp]: haftmann@34007: "\{} = \" haftmann@34007: by (auto intro: antisym Inf_greatest) haftmann@32077: haftmann@41080: lemma Sup_empty [simp]: haftmann@34007: "\{} = \" haftmann@34007: by (auto intro: antisym Sup_least) haftmann@32077: haftmann@41080: lemma Inf_UNIV [simp]: haftmann@41080: "\UNIV = \" haftmann@41080: by (simp add: Sup_Inf Sup_empty [symmetric]) haftmann@41080: haftmann@41080: lemma Sup_UNIV [simp]: haftmann@41080: "\UNIV = \" haftmann@41080: by (simp add: Inf_Sup Inf_empty [symmetric]) haftmann@41080: haftmann@32077: lemma Inf_insert: "\insert a A = a \ \A" haftmann@32077: by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) haftmann@32077: haftmann@32077: lemma Sup_insert: "\insert a A = a \ \A" haftmann@32077: by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) haftmann@32077: haftmann@32077: lemma Inf_singleton [simp]: haftmann@32077: "\{a} = a" haftmann@32077: by (auto intro: antisym Inf_lower Inf_greatest) haftmann@32077: haftmann@32077: lemma Sup_singleton [simp]: haftmann@32077: "\{a} = a" haftmann@32077: by (auto intro: antisym Sup_upper Sup_least) haftmann@32077: haftmann@32077: lemma Inf_binary: haftmann@32077: "\{a, b} = a \ b" haftmann@43831: by (simp add: Inf_insert) haftmann@32077: haftmann@32077: lemma Sup_binary: haftmann@32077: "\{a, b} = a \ b" haftmann@43831: by (simp add: Sup_insert) haftmann@32077: haftmann@43754: lemma le_Inf_iff: "b \ \A \ (\a\A. b \ a)" huffman@35629: by (auto intro: Inf_greatest dest: Inf_lower) huffman@35629: haftmann@43741: lemma Sup_le_iff: "\A \ b \ (\a\A. a \ b)" haftmann@41082: by (auto intro: Sup_least dest: Sup_upper) hoelzl@38705: 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@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@43741: lemma Sup_upper2: "u \ A \ v \ u \ v \ \A" haftmann@43868: using Sup_upper [of u A] by auto hoelzl@41971: haftmann@43741: lemma Inf_lower2: "u \ A \ u \ v \ \A \ v" haftmann@43868: using Inf_lower [of u A] by auto haftmann@43868: 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 haftmann@43868: moreover with assms 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 haftmann@43868: moreover with assms have "u \ v" by blast haftmann@43868: ultimately show ?thesis by (rule Sup_upper2) haftmann@43868: qed haftmann@43868: haftmann@43868: lemma Inf_inter_less_eq: "\A \ \B \ \(A \ B)" haftmann@43868: by (auto intro: Inf_greatest Inf_lower) haftmann@43868: haftmann@43868: lemma Sup_inter_greater_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@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@43868: lemma Inf_top_conv [no_atp]: 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 haftmann@43868: 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 haftmann@43868: with `\A = \` `x \ \` show False by (simp add: Inf_insert) haftmann@43868: qed haftmann@43868: qed haftmann@43868: then show "\ = \A \ (\x\A. x = \)" by auto haftmann@43868: qed haftmann@43868: haftmann@43868: lemma Sup_bot_conv [no_atp]: haftmann@43868: "\A = \ \ (\x\A. x = \)" (is ?P) haftmann@43868: "\ = \A \ (\x\A. x = \)" (is ?Q) haftmann@43868: proof - haftmann@43868: interpret dual: complete_lattice Sup Inf "op \" "op >" sup inf \ \ haftmann@43868: by (fact dual_complete_lattice) haftmann@43868: from dual.Inf_top_conv show ?P and ?Q by simp_all haftmann@43868: qed haftmann@43868: haftmann@43868: lemma Inf_anti_mono: "B \ A \ \A \ \B" haftmann@43868: by (auto intro: Inf_greatest Inf_lower) haftmann@43868: haftmann@43868: lemma Sup_anti_mono: "A \ B \ \A \ \B" haftmann@43868: by (auto intro: Sup_least Sup_upper) hoelzl@41971: haftmann@32077: definition INFI :: "'b set \ ('b \ 'a) \ 'a" where haftmann@32117: "INFI A f = \ (f ` A)" haftmann@32077: haftmann@41082: definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where haftmann@41082: "SUPR A f = \ (f ` A)" haftmann@41082: haftmann@32077: end haftmann@32077: haftmann@32077: syntax haftmann@41082: "_INF1" :: "pttrns \ 'b \ 'b" ("(3INF _./ _)" [0, 10] 10) haftmann@41082: "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3INF _:_./ _)" [0, 0, 10] 10) haftmann@41080: "_SUP1" :: "pttrns \ 'b \ 'b" ("(3SUP _./ _)" [0, 10] 10) haftmann@41080: "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3SUP _:_./ _)" [0, 0, 10] 10) haftmann@41080: haftmann@41080: syntax (xsymbols) haftmann@41082: "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@41082: "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@41080: "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@41080: "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@32077: haftmann@32077: translations haftmann@41082: "INF x y. B" == "INF x. INF y. B" haftmann@41082: "INF x. B" == "CONST INFI CONST UNIV (%x. B)" haftmann@41082: "INF x. B" == "INF x:CONST UNIV. B" haftmann@41082: "INF x:A. B" == "CONST INFI A (%x. B)" haftmann@32077: "SUP x y. B" == "SUP x. SUP y. B" haftmann@32077: "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)" haftmann@32077: "SUP x. B" == "SUP x:CONST UNIV. B" haftmann@32077: "SUP x:A. B" == "CONST SUPR A (%x. B)" haftmann@32077: wenzelm@35115: print_translation {* wenzelm@42284: [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"}, wenzelm@42284: Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}] wenzelm@35115: *} -- {* to avoid eta-contraction of body *} wenzelm@11979: haftmann@32077: context complete_lattice haftmann@32077: begin haftmann@32077: haftmann@43865: lemma INF_empty: "(\x\{}. f x) = \" haftmann@43865: by (simp add: INFI_def) hoelzl@41971: haftmann@43870: lemma SUP_empty: "(\x\{}. f x) = \" haftmann@43870: by (simp add: SUPR_def) haftmann@43870: haftmann@43865: lemma INF_insert: "(\x\insert a A. f x) = f a \ INFI A f" haftmann@43865: by (simp add: INFI_def Inf_insert) haftmann@32077: haftmann@43870: lemma SUP_insert: "(\x\insert a A. f x) = f a \ SUPR A f" haftmann@43870: by (simp add: SUPR_def Sup_insert) haftmann@43870: haftmann@43865: lemma INF_leI: "i \ A \ (\i\A. f i) \ f i" haftmann@32077: by (auto simp add: INFI_def intro: Inf_lower) haftmann@32077: haftmann@43870: lemma le_SUPI: "i \ A \ f i \ (\i\A. f i)" haftmann@43870: by (auto simp add: SUPR_def intro: Sup_upper) haftmann@43870: haftmann@43865: lemma INF_leI2: "i \ A \ f i \ u \ (\i\A. f i) \ u" haftmann@43865: using INF_leI [of i A f] by auto hoelzl@41971: haftmann@43870: lemma le_SUPI2: "i \ A \ u \ f i \ u \ (\i\A. f i)" haftmann@43870: using le_SUPI [of i A f] by auto haftmann@43870: haftmann@43865: lemma le_INFI: "(\i. i \ A \ u \ f i) \ u \ (\i\A. f i)" haftmann@32077: by (auto simp add: INFI_def intro: Inf_greatest) haftmann@32077: haftmann@43870: lemma SUP_leI: "(\i. i \ A \ f i \ u) \ (\i\A. f i) \ u" haftmann@43870: by (auto simp add: SUPR_def intro: Sup_least) haftmann@43870: haftmann@43865: lemma le_INF_iff: "u \ (\i\A. f i) \ (\i \ A. u \ f i)" haftmann@43865: by (auto simp add: INFI_def le_Inf_iff) huffman@35629: haftmann@43870: lemma SUP_le_iff: "(\i\A. f i) \ u \ (\i \ A. f i \ u)" haftmann@43870: by (auto simp add: SUPR_def Sup_le_iff) haftmann@43870: haftmann@43865: lemma INF_const [simp]: "A \ {} \ (\i\A. f) = f" haftmann@32077: by (auto intro: antisym INF_leI le_INFI) haftmann@32077: haftmann@43870: lemma SUP_const [simp]: "A \ {} \ (\i\A. f) = f" haftmann@43870: by (auto intro: antisym SUP_leI le_SUPI) haftmann@43870: haftmann@43865: lemma INF_cong: haftmann@43865: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" haftmann@43865: by (simp add: INFI_def image_def) hoelzl@38705: haftmann@43870: lemma SUP_cong: haftmann@43870: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" haftmann@43870: by (simp add: SUPR_def image_def) haftmann@43870: hoelzl@38705: lemma INF_mono: haftmann@43753: "(\m. m \ B \ \n\A. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" hoelzl@38705: by (force intro!: Inf_mono simp: INFI_def) hoelzl@38705: haftmann@43870: lemma SUP_mono: haftmann@43870: "(\n. n \ A \ \m\B. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" haftmann@43870: by (force intro!: Sup_mono simp: SUPR_def) haftmann@43870: haftmann@43870: lemma INF_subset: haftmann@43870: "A \ B \ INFI B f \ INFI A f" haftmann@43865: by (intro INF_mono) auto haftmann@43865: haftmann@43870: lemma SUP_subset: haftmann@43870: "A \ B \ SUPR A f \ SUPR B f" haftmann@43870: by (intro SUP_mono) auto haftmann@43870: haftmann@43865: lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" haftmann@43865: by (iprover intro: INF_leI le_INFI 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@43870: by (iprover intro: SUP_leI le_SUPI order_trans antisym) haftmann@43870: haftmann@43868: lemma (in complete_lattice) INFI_empty: haftmann@43868: "(\x\{}. B x) = \" haftmann@43868: by (simp add: INFI_def) haftmann@43868: haftmann@43868: lemma (in complete_lattice) INFI_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@43868: lemma (in complete_lattice) INF_union: haftmann@43868: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" haftmann@43868: by (auto intro!: antisym INF_mono intro: le_infI1 le_infI2 le_INFI INF_leI) haftmann@43868: haftmann@43868: lemma (in complete_lattice) INF_constant: haftmann@43868: "(\y\A. c) = (if A = {} then \ else c)" haftmann@43868: by (simp add: INF_empty) haftmann@43868: haftmann@43868: lemma (in complete_lattice) INF_eq: haftmann@43868: "(\x\A. B x) = \({Y. \x\A. Y = B x})" haftmann@43868: by (simp add: INFI_def image_def) haftmann@43868: haftmann@43868: lemma (in complete_lattice) INF_top_conv: haftmann@43868: "\ = (\x\A. B x) \ (\x\A. B x = \)" haftmann@43868: "(\x\A. B x) = \ \ (\x\A. B x = \)" haftmann@43868: by (auto simp add: INFI_def Inf_top_conv) haftmann@43868: haftmann@43868: lemma (in complete_lattice) INFI_UNIV_range: haftmann@43868: "(\x\UNIV. f x) = \range f" haftmann@43868: by (simp add: INFI_def) haftmann@43868: haftmann@43868: lemma (in complete_lattice) INF_bool_eq: haftmann@43868: "(\b. A b) = A True \ A False" haftmann@43868: by (simp add: UNIV_bool INF_empty INF_insert inf_commute) haftmann@43868: haftmann@43868: lemma (in complete_lattice) INF_anti_mono: haftmann@43868: "B \ A \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\B. g x)" haftmann@43868: -- {* The last inclusion is POSITIVE! *} haftmann@43868: by (blast intro: INF_mono dest: subsetD) haftmann@43868: haftmann@32077: end haftmann@32077: haftmann@41082: lemma Inf_less_iff: haftmann@41082: fixes a :: "'a\{complete_lattice,linorder}" haftmann@43753: shows "\S \ a \ (\x\S. x \ a)" haftmann@43754: unfolding not_le [symmetric] le_Inf_iff by auto haftmann@41082: haftmann@43865: lemma INF_less_iff: haftmann@43865: fixes a :: "'a::{complete_lattice,linorder}" haftmann@43865: shows "(\i\A. f i) \ a \ (\x\A. f x \ a)" haftmann@43865: unfolding INFI_def Inf_less_iff by auto haftmann@43865: hoelzl@38705: lemma less_Sup_iff: hoelzl@38705: fixes a :: "'a\{complete_lattice,linorder}" haftmann@43753: shows "a \ \S \ (\x\S. a \ x)" haftmann@43754: unfolding not_le [symmetric] Sup_le_iff by auto hoelzl@38705: hoelzl@40872: lemma less_SUP_iff: hoelzl@40872: fixes a :: "'a::{complete_lattice,linorder}" haftmann@43753: shows "a \ (\i\A. f i) \ (\x\A. a \ f x)" hoelzl@40872: unfolding SUPR_def less_Sup_iff by auto hoelzl@40872: haftmann@32139: subsection {* @{typ bool} and @{typ "_ \ _"} as complete lattice *} haftmann@32077: haftmann@32077: instantiation bool :: complete_lattice haftmann@32077: begin haftmann@32077: haftmann@32077: definition haftmann@41080: "\A \ (\x\A. x)" haftmann@32077: haftmann@32077: definition haftmann@41080: "\A \ (\x\A. x)" haftmann@32077: haftmann@32077: instance proof haftmann@43852: qed (auto simp add: Inf_bool_def Sup_bool_def) haftmann@32077: haftmann@32077: end haftmann@32077: haftmann@41080: lemma INFI_bool_eq [simp]: haftmann@32120: "INFI = Ball" haftmann@32120: proof (rule ext)+ haftmann@32120: fix A :: "'a set" haftmann@32120: fix P :: "'a \ bool" haftmann@43753: show "(\x\A. P x) \ (\x\A. P x)" haftmann@32120: by (auto simp add: Ball_def INFI_def Inf_bool_def) haftmann@32120: qed haftmann@32120: haftmann@41080: lemma SUPR_bool_eq [simp]: haftmann@32120: "SUPR = Bex" haftmann@32120: proof (rule ext)+ haftmann@32120: fix A :: "'a set" haftmann@32120: fix P :: "'a \ bool" haftmann@43753: show "(\x\A. P x) \ (\x\A. P x)" haftmann@32120: by (auto simp add: Bex_def SUPR_def Sup_bool_def) haftmann@32120: qed haftmann@32120: haftmann@32077: instantiation "fun" :: (type, complete_lattice) complete_lattice haftmann@32077: begin haftmann@32077: haftmann@32077: definition haftmann@41080: "\A = (\x. \{y. \f\A. y = f x})" haftmann@41080: haftmann@41080: lemma Inf_apply: haftmann@41080: "(\A) x = \{y. \f\A. y = f x}" haftmann@41080: by (simp add: Inf_fun_def) haftmann@32077: haftmann@32077: definition haftmann@41080: "\A = (\x. \{y. \f\A. y = f x})" haftmann@41080: haftmann@41080: lemma Sup_apply: haftmann@41080: "(\A) x = \{y. \f\A. y = f x}" haftmann@41080: by (simp add: Sup_fun_def) haftmann@32077: haftmann@32077: instance proof haftmann@41080: qed (auto simp add: le_fun_def Inf_apply Sup_apply haftmann@32077: intro: Inf_lower Sup_upper Inf_greatest Sup_least) haftmann@32077: haftmann@32077: end haftmann@32077: haftmann@41080: lemma INFI_apply: haftmann@41080: "(\y\A. f y) x = (\y\A. f y x)" haftmann@41080: by (auto intro: arg_cong [of _ _ Inf] simp add: INFI_def Inf_apply) hoelzl@38705: haftmann@41080: lemma SUPR_apply: haftmann@41080: "(\y\A. f y) x = (\y\A. f y x)" haftmann@41080: by (auto intro: arg_cong [of _ _ Sup] simp add: SUPR_def Sup_apply) haftmann@32077: haftmann@32077: haftmann@41082: subsection {* 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@41082: Inter ("\_" [90] 90) 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@41082: by (simp add: Inf_fun_def Inf_bool_def) (simp add: mem_def) haftmann@41082: qed haftmann@41082: haftmann@43741: lemma Inter_iff [simp,no_atp]: "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: haftmann@41082: lemma Int_eq_Inter: "A \ B = \{A, B}" haftmann@43739: by (fact Inf_binary [symmetric]) haftmann@41082: haftmann@41082: lemma Inter_empty [simp]: "\{} = UNIV" haftmann@41082: by (fact Inf_empty) haftmann@41082: haftmann@41082: lemma Inter_UNIV [simp]: "\UNIV = {}" haftmann@43739: by (fact Inf_UNIV) haftmann@41082: haftmann@41082: lemma Inter_insert [simp]: "\(insert a B) = a \ \B" haftmann@43739: by (fact Inf_insert) haftmann@41082: haftmann@41082: lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" haftmann@43868: by (fact Inf_inter_less_eq) haftmann@41082: haftmann@41082: lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" haftmann@43756: by (fact Inf_union_distrib) haftmann@43756: haftmann@43868: lemma Inter_UNIV_conv [simp, no_atp]: 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@43756: by (fact Inf_anti_mono) haftmann@41082: haftmann@41082: haftmann@41082: subsection {* Intersections of families *} haftmann@41082: haftmann@41082: abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" where haftmann@41082: "INTER \ INFI" haftmann@41082: 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@41082: lemma INTER_eq_Inter_image: haftmann@41082: "(\x\A. B x) = \(B`A)" haftmann@41082: by (fact INFI_def) haftmann@41082: haftmann@41082: lemma Inter_def: haftmann@41082: "\S = (\x\S. x)" haftmann@41082: by (simp add: INTER_eq_Inter_image image_def) haftmann@41082: haftmann@41082: lemma INTER_def: haftmann@41082: "(\x\A. B x) = {y. \x\A. y \ B x}" haftmann@41082: by (auto simp add: INTER_eq_Inter_image Inter_eq) haftmann@41082: haftmann@41082: lemma Inter_image_eq [simp]: haftmann@41082: "\(B`A) = (\x\A. B x)" haftmann@43801: by (rule sym) (fact INFI_def) haftmann@41082: haftmann@43817: lemma INT_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" haftmann@41082: by (unfold INTER_def) blast haftmann@41082: haftmann@43817: lemma INT_I [intro!]: "(\x. x \ A \ b \ B x) \ b \ (\x\A. B x)" haftmann@41082: by (unfold INTER_def) blast 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@41082: by (unfold INTER_def) blast 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@41082: by (fact INF_leI) haftmann@41082: haftmann@43817: lemma INT_greatest: "(\x. x \ A \ C \ B x) \ C \ (\x\A. B x)" haftmann@41082: by (fact le_INFI) haftmann@41082: haftmann@41082: lemma INT_empty [simp]: "(\x\{}. B x) = UNIV" haftmann@43854: by (fact INFI_empty) haftmann@43854: haftmann@43817: lemma INT_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" haftmann@43854: by (fact INFI_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: haftmann@41082: lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" haftmann@41082: -- {* Look: it has an \emph{existential} quantifier *} haftmann@43865: by (fact INF_eq) haftmann@43865: haftmann@43854: lemma INTER_UNIV_conv [simp]: haftmann@43817: "(UNIV = (\x\A. B x)) = (\x\A. B x = UNIV)" haftmann@43817: "((\x\A. B x) = UNIV) = (\x\A. B x = UNIV)" haftmann@43865: by (fact INF_top_conv)+ haftmann@43865: haftmann@43865: lemma INT_bool_eq: "(\b. A b) = A True \ A False" haftmann@43865: by (fact INF_bool_eq) haftmann@43865: haftmann@43865: lemma INT_anti_mono: haftmann@43867: "B \ A \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\B. g x)" haftmann@43865: -- {* The last inclusion is POSITIVE! *} haftmann@43867: by (fact INF_anti_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@32139: subsection {* 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@32115: Union ("\_" [90] 90) 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@32587: by (simp add: Sup_fun_def Sup_bool_def) (simp add: mem_def) haftmann@32115: qed haftmann@32115: blanchet@35828: lemma Union_iff [simp, no_atp]: 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@32135: by (iprover intro: subsetI UnionI) haftmann@32135: haftmann@43817: lemma Union_least: "(\X. X \ A \ X \ C) \ \A \ C" haftmann@32135: by (iprover intro: subsetI elim: UnionE dest: subsetD) haftmann@32135: haftmann@32135: lemma Un_eq_Union: "A \ B = \{A, B}" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma Union_empty [simp]: "\{} = {}" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma Union_UNIV [simp]: "\UNIV = UNIV" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma Union_insert [simp]: "\insert a B = a \ \B" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma Union_Un_distrib [simp]: "\(A \ B) = \A \ \B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma Union_Int_subset: "\(A \ B) \ \A \ \B" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma Union_empty_conv [simp,no_atp]: "(\A = {}) \ (\x\A. x = {})" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma empty_Union_conv [simp,no_atp]: "({} = \A) \ (\x\A. x = {})" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma Union_disjoint: "(\C \ A = {}) \ (\B\C. B \ A = {})" haftmann@32135: by blast 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@32135: by blast haftmann@32135: haftmann@32115: haftmann@32139: subsection {* Unions of families *} haftmann@32077: haftmann@32606: abbreviation UNION :: "'a set \ ('a \ 'b set) \ 'b set" where haftmann@32606: "UNION \ SUPR" haftmann@32077: 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 haftmann@32077: unions and intersections (e.g.\ @{text"\a\<^isub>1\A\<^isub>1. B"}) haftmann@32077: and their \LaTeX\ rendition: @{term"\a\<^isub>1\A\<^isub>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: haftmann@32135: lemma UNION_eq_Union_image: haftmann@43817: "(\x\A. B x) = \(B ` A)" haftmann@32606: by (fact SUPR_def) haftmann@32115: haftmann@32115: lemma Union_def: haftmann@32117: "\S = (\x\S. x)" haftmann@32115: by (simp add: UNION_eq_Union_image image_def) haftmann@32115: blanchet@35828: lemma UNION_def [no_atp]: haftmann@32135: "(\x\A. B x) = {y. \x\A. y \ B x}" haftmann@32117: by (auto simp add: UNION_eq_Union_image Union_eq) haftmann@32115: haftmann@32115: lemma Union_image_eq [simp]: haftmann@43817: "\(B ` A) = (\x\A. B x)" haftmann@32115: by (rule sym) (fact UNION_eq_Union_image) haftmann@32115: haftmann@43852: lemma UN_iff [simp]: "(b \ (\x\A. B x)) = (\x\A. b \ B x)" wenzelm@11979: by (unfold UNION_def) blast 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" wenzelm@11979: by (unfold UNION_def) blast clasohm@923: wenzelm@11979: lemma UN_cong [cong]: haftmann@43852: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" wenzelm@11979: by (simp add: UNION_def) wenzelm@11979: berghofe@29691: lemma strong_UN_cong: haftmann@43852: "A = B \ (\x. x \ B =simp=> C x = D x) \ (\x\A. C x) = (\x\B. D x)" berghofe@29691: by (simp add: UNION_def simp_implies_def) 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@32606: by (fact le_SUPI) haftmann@32135: haftmann@43817: lemma UN_least: "(\x. x \ A \ B x \ C) \ (\x\A. B x) \ C" haftmann@32135: by (iprover intro: subsetI elim: UN_E dest: subsetD) haftmann@32135: blanchet@35828: lemma Collect_bex_eq [no_atp]: "{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@35828: lemma UN_empty [simp,no_atp]: "(\x\{}. B x) = {}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma UN_empty2 [simp]: "(\x\A. {}) = {}" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma UN_singleton [simp]: "(\x\A. {x}) = A" haftmann@32135: by blast haftmann@32135: haftmann@43817: lemma UN_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma UN_Un[simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" haftmann@32135: by blast haftmann@32135: haftmann@32135: 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 image_Union: "f ` \S = (\x\S. f ` x)" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" haftmann@32135: by auto haftmann@32135: haftmann@32135: lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma UNION_empty_conv[simp]: haftmann@43817: "{} = (\x\A. B x) \ (\x\A. B x = {})" haftmann@43817: "(\x\A. B x) = {} \ (\x\A. B x = {})" haftmann@32135: by blast+ haftmann@32135: blanchet@35828: lemma Collect_ex_eq [no_atp]: "{x. \y. P x y} = (\y. {x. P x y})" haftmann@32135: by blast haftmann@32135: haftmann@32135: lemma ball_UN: "(\z \ UNION A B. P z) = (\x\A. \z \ B x. P z)" haftmann@32135: by blast haftmann@32135: haftmann@32135: 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@32135: by (auto intro: bool_contrapos) 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@32135: by (blast dest: subsetD) 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: wenzelm@11979: haftmann@32139: subsection {* Distributive laws *} wenzelm@12897: wenzelm@12897: lemma Int_Union: "A \ \B = (\C\B. A \ C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_Union2: "\B \ A = (\C\B. C \ A)" wenzelm@12897: by blast 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 *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" wenzelm@12897: -- {* Equivalent version *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_Inter: "A \ \B = (\C\B. A \ C)" wenzelm@12897: by blast wenzelm@12897: haftmann@43817: lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" wenzelm@12897: -- {* Equivalent version *} wenzelm@12897: by blast 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. *} wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: haftmann@32139: subsection {* Complement *} haftmann@32135: haftmann@43817: lemma Compl_UN [simp]: "- (\x\A. B x) = (\x\A. -B x)" wenzelm@12897: by blast wenzelm@12897: haftmann@43817: lemma Compl_INT [simp]: "- (\x\A. B x) = (\x\A. -B x)" wenzelm@12897: by blast wenzelm@12897: wenzelm@12897: haftmann@32139: subsection {* 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@43852: "\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@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 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@43831: "\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@35828: lemma ball_simps [simp,no_atp]: haftmann@43852: "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" haftmann@43852: "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" haftmann@43852: "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" haftmann@43852: "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" haftmann@43852: "\P. (\x\{}. P x) \ True" haftmann@43852: "\P. (\x\UNIV. P x) \ (\x. P x)" haftmann@43852: "\a B P. (\x\insert a B. P x) \ (P a \ (\x\B. 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)" haftmann@43852: "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" haftmann@43852: "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" haftmann@43852: "\A P. (\ (\x\A. P x)) \ (\x\A. \ P x)" wenzelm@12897: by auto wenzelm@12897: blanchet@35828: lemma bex_simps [simp,no_atp]: haftmann@43852: "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" haftmann@43852: "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" haftmann@43852: "\P. (\x\{}. P x) \ False" haftmann@43852: "\P. (\x\UNIV. P x) \ (\x. P x)" haftmann@43852: "\a B P. (\x\insert a B. P x) \ (P a | (\x\B. 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)" haftmann@43852: "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" haftmann@43852: "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" haftmann@43852: "\A P. (\(\x\A. P x)) \ (\x\A. \ P x)" wenzelm@12897: by auto wenzelm@12897: 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@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) = (\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: paulson@13860: haftmann@32135: no_notation haftmann@32135: less_eq (infix "\" 50) and haftmann@32135: less (infix "\" 50) and haftmann@41082: bot ("\") and haftmann@41082: top ("\") and haftmann@32135: inf (infixl "\" 70) and haftmann@32135: sup (infixl "\" 65) and haftmann@32135: Inf ("\_" [900] 900) and haftmann@41082: Sup ("\_" [900] 900) haftmann@32135: haftmann@41080: no_syntax (xsymbols) haftmann@41082: "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@41082: "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@41080: "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) haftmann@41080: "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) haftmann@41080: 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