# HG changeset patch # User haftmann # Date 1315643364 -7200 # Node ID 56101fa00193b0c02f68041ef215916d48204fc1 # Parent 237ba63d6041137cc895d779379d34ed2cbc92bf renamed theory Complete_Lattice to Complete_Lattices, in accordance with Lattices, Orderings etc. diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Sat Sep 10 00:44:25 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1259 +0,0 @@ - (* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel; Florian Haftmann, TU Muenchen *) - -header {* Complete lattices *} - -theory Complete_Lattice -imports Set -begin - -notation - less_eq (infix "\" 50) and - less (infix "\" 50) and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - top ("\") and - bot ("\") - - -subsection {* Syntactic infimum and supremum operations *} - -class Inf = - fixes Inf :: "'a set \ 'a" ("\_" [900] 900) - -class Sup = - fixes Sup :: "'a set \ 'a" ("\_" [900] 900) - -subsection {* Abstract complete lattices *} - -class complete_lattice = bounded_lattice + Inf + Sup + - assumes Inf_lower: "x \ A \ \A \ x" - and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" - assumes Sup_upper: "x \ A \ x \ \A" - and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" -begin - -lemma dual_complete_lattice: - "class.complete_lattice Sup Inf sup (op \) (op >) inf \ \" - by (auto intro!: class.complete_lattice.intro dual_bounded_lattice) - (unfold_locales, (fact bot_least top_greatest - Sup_upper Sup_least Inf_lower Inf_greatest)+) - -definition INFI :: "'b set \ ('b \ 'a) \ 'a" where - INF_def: "INFI A f = \(f ` A)" - -definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where - SUP_def: "SUPR A f = \(f ` A)" - -text {* - Note: must use names @{const INFI} and @{const SUPR} here instead of - @{text INF} and @{text SUP} to allow the following syntax coexist - with the plain constant names. -*} - -end - -syntax - "_INF1" :: "pttrns \ 'b \ 'b" ("(3INF _./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3INF _:_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3SUP _./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3SUP _:_./ _)" [0, 0, 10] 10) - -syntax (xsymbols) - "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - -translations - "INF x y. B" == "INF x. INF y. B" - "INF x. B" == "CONST INFI CONST UNIV (%x. B)" - "INF x. B" == "INF x:CONST UNIV. B" - "INF x:A. B" == "CONST INFI A (%x. B)" - "SUP x y. B" == "SUP x. SUP y. B" - "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)" - "SUP x. B" == "SUP x:CONST UNIV. B" - "SUP x:A. B" == "CONST SUPR A (%x. B)" - -print_translation {* - [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"}, - Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}] -*} -- {* to avoid eta-contraction of body *} - -context complete_lattice -begin - -lemma INF_foundation_dual [no_atp]: - "complete_lattice.SUPR Inf = INFI" -proof (rule ext)+ - interpret dual: complete_lattice Sup Inf sup "op \" "op >" inf \ \ - by (fact dual_complete_lattice) - fix f :: "'b \ 'a" and A - show "complete_lattice.SUPR Inf A f = (\a\A. f a)" - by (simp only: dual.SUP_def INF_def) -qed - -lemma SUP_foundation_dual [no_atp]: - "complete_lattice.INFI Sup = SUPR" -proof (rule ext)+ - interpret dual: complete_lattice Sup Inf sup "op \" "op >" inf \ \ - by (fact dual_complete_lattice) - fix f :: "'b \ 'a" and A - show "complete_lattice.INFI Sup A f = (\a\A. f a)" - by (simp only: dual.INF_def SUP_def) -qed - -lemma INF_lower: "i \ A \ (\i\A. f i) \ f i" - by (auto simp add: INF_def intro: Inf_lower) - -lemma INF_greatest: "(\i. i \ A \ u \ f i) \ u \ (\i\A. f i)" - by (auto simp add: INF_def intro: Inf_greatest) - -lemma SUP_upper: "i \ A \ f i \ (\i\A. f i)" - by (auto simp add: SUP_def intro: Sup_upper) - -lemma SUP_least: "(\i. i \ A \ f i \ u) \ (\i\A. f i) \ u" - by (auto simp add: SUP_def intro: Sup_least) - -lemma Inf_lower2: "u \ A \ u \ v \ \A \ v" - using Inf_lower [of u A] by auto - -lemma INF_lower2: "i \ A \ f i \ u \ (\i\A. f i) \ u" - using INF_lower [of i A f] by auto - -lemma Sup_upper2: "u \ A \ v \ u \ v \ \A" - using Sup_upper [of u A] by auto - -lemma SUP_upper2: "i \ A \ u \ f i \ u \ (\i\A. f i)" - using SUP_upper [of i A f] by auto - -lemma le_Inf_iff (*[simp]*): "b \ \A \ (\a\A. b \ a)" - by (auto intro: Inf_greatest dest: Inf_lower) - -lemma le_INF_iff (*[simp]*): "u \ (\i\A. f i) \ (\i\A. u \ f i)" - by (auto simp add: INF_def le_Inf_iff) - -lemma Sup_le_iff (*[simp]*): "\A \ b \ (\a\A. a \ b)" - by (auto intro: Sup_least dest: Sup_upper) - -lemma SUP_le_iff (*[simp]*): "(\i\A. f i) \ u \ (\i\A. f i \ u)" - by (auto simp add: SUP_def Sup_le_iff) - -lemma Inf_empty [simp]: - "\{} = \" - by (auto intro: antisym Inf_greatest) - -lemma INF_empty [simp]: "(\x\{}. f x) = \" - by (simp add: INF_def) - -lemma Sup_empty [simp]: - "\{} = \" - by (auto intro: antisym Sup_least) - -lemma SUP_empty [simp]: "(\x\{}. f x) = \" - by (simp add: SUP_def) - -lemma Inf_UNIV [simp]: - "\UNIV = \" - by (auto intro!: antisym Inf_lower) - -lemma Sup_UNIV [simp]: - "\UNIV = \" - by (auto intro!: antisym Sup_upper) - -lemma Inf_insert (*[simp]*): "\insert a A = a \ \A" - by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) - -lemma INF_insert: "(\x\insert a A. f x) = f a \ INFI A f" - by (simp add: INF_def Inf_insert) - -lemma Sup_insert (*[simp]*): "\insert a A = a \ \A" - by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) - -lemma SUP_insert: "(\x\insert a A. f x) = f a \ SUPR A f" - by (simp add: SUP_def Sup_insert) - -lemma INF_image (*[simp]*): "(\x\f`A. g x) = (\x\A. g (f x))" - by (simp add: INF_def image_image) - -lemma SUP_image (*[simp]*): "(\x\f`A. g x) = (\x\A. g (f x))" - by (simp add: SUP_def image_image) - -lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Inf_superset_mono: "B \ A \ \A \ \B" - by (auto intro: Inf_greatest Inf_lower) - -lemma Sup_subset_mono: "A \ B \ \A \ \B" - by (auto intro: Sup_least Sup_upper) - -lemma INF_cong: - "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (simp add: INF_def image_def) - -lemma SUP_cong: - "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (simp add: SUP_def image_def) - -lemma Inf_mono: - assumes "\b. b \ B \ \a\A. a \ b" - shows "\A \ \B" -proof (rule Inf_greatest) - fix b assume "b \ B" - with assms obtain a where "a \ A" and "a \ b" by blast - from `a \ A` have "\A \ a" by (rule Inf_lower) - with `a \ b` show "\A \ b" by auto -qed - -lemma INF_mono: - "(\m. m \ B \ \n\A. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" - by (force intro!: Inf_mono simp: INF_def) - -lemma Sup_mono: - assumes "\a. a \ A \ \b\B. a \ b" - shows "\A \ \B" -proof (rule Sup_least) - fix a assume "a \ A" - with assms obtain b where "b \ B" and "a \ b" by blast - from `b \ B` have "b \ \B" by (rule Sup_upper) - with `a \ b` show "a \ \B" by auto -qed - -lemma SUP_mono: - "(\n. n \ A \ \m\B. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" - by (force intro!: Sup_mono simp: SUP_def) - -lemma INF_superset_mono: - "B \ A \ (\x. x \ B \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" - -- {* The last inclusion is POSITIVE! *} - by (blast intro: INF_mono dest: subsetD) - -lemma SUP_subset_mono: - "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" - by (blast intro: SUP_mono dest: subsetD) - -lemma Inf_less_eq: - assumes "\v. v \ A \ v \ u" - and "A \ {}" - shows "\A \ u" -proof - - from `A \ {}` obtain v where "v \ A" by blast - moreover with assms have "v \ u" by blast - ultimately show ?thesis by (rule Inf_lower2) -qed - -lemma less_eq_Sup: - assumes "\v. v \ A \ u \ v" - and "A \ {}" - shows "u \ \A" -proof - - from `A \ {}` obtain v where "v \ A" by blast - moreover with assms have "u \ v" by blast - ultimately show ?thesis by (rule Sup_upper2) -qed - -lemma less_eq_Inf_inter: "\A \ \B \ \(A \ B)" - by (auto intro: Inf_greatest Inf_lower) - -lemma Sup_inter_less_eq: "\(A \ B) \ \A \ \B " - by (auto intro: Sup_least Sup_upper) - -lemma Inf_union_distrib: "\(A \ B) = \A \ \B" - by (rule antisym) (auto intro: Inf_greatest Inf_lower le_infI1 le_infI2) - -lemma INF_union: - "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" - by (auto intro!: antisym INF_mono intro: le_infI1 le_infI2 INF_greatest INF_lower) - -lemma Sup_union_distrib: "\(A \ B) = \A \ \B" - by (rule antisym) (auto intro: Sup_least Sup_upper le_supI1 le_supI2) - -lemma SUP_union: - "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" - by (auto intro!: antisym SUP_mono intro: le_supI1 le_supI2 SUP_least SUP_upper) - -lemma INF_inf_distrib: "(\a\A. f a) \ (\a\A. g a) = (\a\A. f a \ g a)" - by (rule antisym) (rule INF_greatest, auto intro: le_infI1 le_infI2 INF_lower INF_mono) - -lemma SUP_sup_distrib: "(\a\A. f a) \ (\a\A. g a) = (\a\A. f a \ g a)" - by (rule antisym) (auto intro: le_supI1 le_supI2 SUP_upper SUP_mono, - rule SUP_least, auto intro: le_supI1 le_supI2 SUP_upper SUP_mono) - -lemma Inf_top_conv (*[simp]*) [no_atp]: - "\A = \ \ (\x\A. x = \)" - "\ = \A \ (\x\A. x = \)" -proof - - show "\A = \ \ (\x\A. x = \)" - proof - assume "\x\A. x = \" - then have "A = {} \ A = {\}" by auto - then show "\A = \" by (auto simp add: Inf_insert) - next - assume "\A = \" - show "\x\A. x = \" - proof (rule ccontr) - assume "\ (\x\A. x = \)" - then obtain x where "x \ A" and "x \ \" by blast - then obtain B where "A = insert x B" by blast - with `\A = \` `x \ \` show False by (simp add: Inf_insert) - qed - qed - then show "\ = \A \ (\x\A. x = \)" by auto -qed - -lemma INF_top_conv (*[simp]*): - "(\x\A. B x) = \ \ (\x\A. B x = \)" - "\ = (\x\A. B x) \ (\x\A. B x = \)" - by (auto simp add: INF_def Inf_top_conv) - -lemma Sup_bot_conv (*[simp]*) [no_atp]: - "\A = \ \ (\x\A. x = \)" (is ?P) - "\ = \A \ (\x\A. x = \)" (is ?Q) -proof - - interpret dual: complete_lattice Sup Inf sup "op \" "op >" inf \ \ - by (fact dual_complete_lattice) - from dual.Inf_top_conv show ?P and ?Q by simp_all -qed - -lemma SUP_bot_conv (*[simp]*): - "(\x\A. B x) = \ \ (\x\A. B x = \)" - "\ = (\x\A. B x) \ (\x\A. B x = \)" - by (auto simp add: SUP_def Sup_bot_conv) - -lemma INF_const [simp]: "A \ {} \ (\i\A. f) = f" - by (auto intro: antisym INF_lower INF_greatest) - -lemma SUP_const [simp]: "A \ {} \ (\i\A. f) = f" - by (auto intro: antisym SUP_upper SUP_least) - -lemma INF_top (*[simp]*): "(\x\A. \) = \" - by (cases "A = {}") (simp_all add: INF_empty) - -lemma SUP_bot (*[simp]*): "(\x\A. \) = \" - by (cases "A = {}") (simp_all add: SUP_empty) - -lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" - by (iprover intro: INF_lower INF_greatest order_trans antisym) - -lemma SUP_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" - by (iprover intro: SUP_upper SUP_least order_trans antisym) - -lemma INF_absorb: - assumes "k \ I" - shows "A k \ (\i\I. A i) = (\i\I. A i)" -proof - - from assms obtain J where "I = insert k J" by blast - then show ?thesis by (simp add: INF_insert) -qed - -lemma SUP_absorb: - assumes "k \ I" - shows "A k \ (\i\I. A i) = (\i\I. A i)" -proof - - from assms obtain J where "I = insert k J" by blast - then show ?thesis by (simp add: SUP_insert) -qed - -lemma INF_constant: - "(\y\A. c) = (if A = {} then \ else c)" - by (simp add: INF_empty) - -lemma SUP_constant: - "(\y\A. c) = (if A = {} then \ else c)" - by (simp add: SUP_empty) - -lemma less_INF_D: - assumes "y < (\i\A. f i)" "i \ A" shows "y < f i" -proof - - note `y < (\i\A. f i)` - also have "(\i\A. f i) \ f i" using `i \ A` - by (rule INF_lower) - finally show "y < f i" . -qed - -lemma SUP_lessD: - assumes "(\i\A. f i) < y" "i \ A" shows "f i < y" -proof - - have "f i \ (\i\A. f i)" using `i \ A` - by (rule SUP_upper) - also note `(\i\A. f i) < y` - finally show "f i < y" . -qed - -lemma INF_UNIV_bool_expand: - "(\b. A b) = A True \ A False" - by (simp add: UNIV_bool INF_empty INF_insert inf_commute) - -lemma SUP_UNIV_bool_expand: - "(\b. A b) = A True \ A False" - by (simp add: UNIV_bool SUP_empty SUP_insert sup_commute) - -end - -class complete_distrib_lattice = complete_lattice + - assumes sup_Inf: "a \ \B = (\b\B. a \ b)" - assumes inf_Sup: "a \ \B = (\b\B. a \ b)" -begin - -lemma sup_INF: - "a \ (\b\B. f b) = (\b\B. a \ f b)" - by (simp add: INF_def sup_Inf image_image) - -lemma inf_SUP: - "a \ (\b\B. f b) = (\b\B. a \ f b)" - by (simp add: SUP_def inf_Sup image_image) - -lemma dual_complete_distrib_lattice: - "class.complete_distrib_lattice Sup Inf sup (op \) (op >) inf \ \" - apply (rule class.complete_distrib_lattice.intro) - apply (fact dual_complete_lattice) - apply (rule class.complete_distrib_lattice_axioms.intro) - apply (simp_all only: INF_foundation_dual SUP_foundation_dual inf_Sup sup_Inf) - done - -subclass distrib_lattice proof - fix a b c - from sup_Inf have "a \ \{b, c} = (\d\{b, c}. a \ d)" . - then show "a \ b \ c = (a \ b) \ (a \ c)" by (simp add: INF_def Inf_insert) -qed - -lemma Inf_sup: - "\B \ a = (\b\B. b \ a)" - by (simp add: sup_Inf sup_commute) - -lemma Sup_inf: - "\B \ a = (\b\B. b \ a)" - by (simp add: inf_Sup inf_commute) - -lemma INF_sup: - "(\b\B. f b) \ a = (\b\B. f b \ a)" - by (simp add: sup_INF sup_commute) - -lemma SUP_inf: - "(\b\B. f b) \ a = (\b\B. f b \ a)" - by (simp add: inf_SUP inf_commute) - -lemma Inf_sup_eq_top_iff: - "(\B \ a = \) \ (\b\B. b \ a = \)" - by (simp only: Inf_sup INF_top_conv) - -lemma Sup_inf_eq_bot_iff: - "(\B \ a = \) \ (\b\B. b \ a = \)" - by (simp only: Sup_inf SUP_bot_conv) - -lemma INF_sup_distrib2: - "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" - by (subst INF_commute) (simp add: sup_INF INF_sup) - -lemma SUP_inf_distrib2: - "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" - by (subst SUP_commute) (simp add: inf_SUP SUP_inf) - -end - -class complete_boolean_algebra = boolean_algebra + complete_distrib_lattice -begin - -lemma dual_complete_boolean_algebra: - "class.complete_boolean_algebra Sup Inf sup (op \) (op >) inf \ \ (\x y. x \ - y) uminus" - by (rule class.complete_boolean_algebra.intro, rule dual_complete_distrib_lattice, rule dual_boolean_algebra) - -lemma uminus_Inf: - "- (\A) = \(uminus ` A)" -proof (rule antisym) - show "- \A \ \(uminus ` A)" - by (rule compl_le_swap2, rule Inf_greatest, rule compl_le_swap2, rule Sup_upper) simp - show "\(uminus ` A) \ - \A" - by (rule Sup_least, rule compl_le_swap1, rule Inf_lower) auto -qed - -lemma uminus_INF: "- (\x\A. B x) = (\x\A. - B x)" - by (simp add: INF_def SUP_def uminus_Inf image_image) - -lemma uminus_Sup: - "- (\A) = \(uminus ` A)" -proof - - have "\A = - \(uminus ` A)" by (simp add: image_image uminus_Inf) - then show ?thesis by simp -qed - -lemma uminus_SUP: "- (\x\A. B x) = (\x\A. - B x)" - by (simp add: INF_def SUP_def uminus_Sup image_image) - -end - -class complete_linorder = linorder + complete_lattice -begin - -lemma dual_complete_linorder: - "class.complete_linorder Sup Inf sup (op \) (op >) inf \ \" - by (rule class.complete_linorder.intro, rule dual_complete_lattice, rule dual_linorder) - -lemma Inf_less_iff (*[simp]*): - "\S \ a \ (\x\S. x \ a)" - unfolding not_le [symmetric] le_Inf_iff by auto - -lemma INF_less_iff (*[simp]*): - "(\i\A. f i) \ a \ (\x\A. f x \ a)" - unfolding INF_def Inf_less_iff by auto - -lemma less_Sup_iff (*[simp]*): - "a \ \S \ (\x\S. a \ x)" - unfolding not_le [symmetric] Sup_le_iff by auto - -lemma less_SUP_iff (*[simp]*): - "a \ (\i\A. f i) \ (\x\A. a \ f x)" - unfolding SUP_def less_Sup_iff by auto - -lemma Sup_eq_top_iff (*[simp]*): - "\A = \ \ (\x<\. \i\A. x < i)" -proof - assume *: "\A = \" - show "(\x<\. \i\A. x < i)" unfolding * [symmetric] - proof (intro allI impI) - fix x assume "x < \A" then show "\i\A. x < i" - unfolding less_Sup_iff by auto - qed -next - assume *: "\x<\. \i\A. x < i" - show "\A = \" - proof (rule ccontr) - assume "\A \ \" - with top_greatest [of "\A"] - have "\A < \" unfolding le_less by auto - then have "\A < \A" - using * unfolding less_Sup_iff by auto - then show False by auto - qed -qed - -lemma SUP_eq_top_iff (*[simp]*): - "(\i\A. f i) = \ \ (\x<\. \i\A. x < f i)" - unfolding SUP_def Sup_eq_top_iff by auto - -lemma Inf_eq_bot_iff (*[simp]*): - "\A = \ \ (\x>\. \i\A. i < x)" -proof - - interpret dual: complete_linorder Sup Inf sup "op \" "op >" inf \ \ - by (fact dual_complete_linorder) - from dual.Sup_eq_top_iff show ?thesis . -qed - -lemma INF_eq_bot_iff (*[simp]*): - "(\i\A. f i) = \ \ (\x>\. \i\A. f i < x)" - unfolding INF_def Inf_eq_bot_iff by auto - -end - - -subsection {* @{typ bool} and @{typ "_ \ _"} as complete lattice *} - -instantiation bool :: complete_lattice -begin - -definition - [simp]: "\A \ False \ A" - -definition - [simp]: "\A \ True \ A" - -instance proof -qed (auto intro: bool_induct) - -end - -lemma INF_bool_eq [simp]: - "INFI = Ball" -proof (rule ext)+ - fix A :: "'a set" - fix P :: "'a \ bool" - show "(\x\A. P x) \ (\x\A. P x)" - by (auto simp add: INF_def) -qed - -lemma SUP_bool_eq [simp]: - "SUPR = Bex" -proof (rule ext)+ - fix A :: "'a set" - fix P :: "'a \ bool" - show "(\x\A. P x) \ (\x\A. P x)" - by (auto simp add: SUP_def) -qed - -instance bool :: complete_boolean_algebra proof -qed (auto intro: bool_induct) - -instantiation "fun" :: (type, complete_lattice) complete_lattice -begin - -definition - "\A = (\x. \f\A. f x)" - -lemma Inf_apply: - "(\A) x = (\f\A. f x)" - by (simp add: Inf_fun_def) - -definition - "\A = (\x. \f\A. f x)" - -lemma Sup_apply: - "(\A) x = (\f\A. f x)" - by (simp add: Sup_fun_def) - -instance proof -qed (auto simp add: le_fun_def Inf_apply Sup_apply intro: INF_lower INF_greatest SUP_upper SUP_least) - -end - -lemma INF_apply: - "(\y\A. f y) x = (\y\A. f y x)" - by (auto intro: arg_cong [of _ _ Inf] simp add: INF_def Inf_apply) - -lemma SUP_apply: - "(\y\A. f y) x = (\y\A. f y x)" - by (auto intro: arg_cong [of _ _ Sup] simp add: SUP_def Sup_apply) - -instance "fun" :: (type, complete_distrib_lattice) complete_distrib_lattice proof -qed (auto simp add: inf_apply sup_apply Inf_apply Sup_apply INF_def SUP_def inf_Sup sup_Inf image_image) - -instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra .. - - -subsection {* Inter *} - -abbreviation Inter :: "'a set set \ 'a set" where - "Inter S \ \S" - -notation (xsymbols) - Inter ("\_" [90] 90) - -lemma Inter_eq: - "\A = {x. \B \ A. x \ B}" -proof (rule set_eqI) - fix x - have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" - by auto - then show "x \ \A \ x \ {x. \B \ A. x \ B}" - by (simp add: Inf_fun_def) (simp add: mem_def) -qed - -lemma Inter_iff [simp,no_atp]: "A \ \C \ (\X\C. A \ X)" - by (unfold Inter_eq) blast - -lemma InterI [intro!]: "(\X. X \ C \ A \ X) \ A \ \C" - by (simp add: Inter_eq) - -text {* - \medskip A ``destruct'' rule -- every @{term X} in @{term C} - contains @{term A} as an element, but @{prop "A \ X"} can hold when - @{prop "X \ C"} does not! This rule is analogous to @{text spec}. -*} - -lemma InterD [elim, Pure.elim]: "A \ \C \ X \ C \ A \ X" - by auto - -lemma InterE [elim]: "A \ \C \ (X \ C \ R) \ (A \ X \ R) \ R" - -- {* ``Classical'' elimination rule -- does not require proving - @{prop "X \ C"}. *} - by (unfold Inter_eq) blast - -lemma Inter_lower: "B \ A \ \A \ B" - by (fact Inf_lower) - -lemma Inter_subset: - "(\X. X \ A \ X \ B) \ A \ {} \ \A \ B" - by (fact Inf_less_eq) - -lemma Inter_greatest: "(\X. X \ A \ C \ X) \ C \ Inter A" - by (fact Inf_greatest) - -lemma Inter_empty: "\{} = UNIV" - by (fact Inf_empty) (* already simp *) - -lemma Inter_UNIV: "\UNIV = {}" - by (fact Inf_UNIV) (* already simp *) - -lemma Inter_insert [simp]: "\(insert a B) = a \ \B" - by (fact Inf_insert) - -lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" - by (fact less_eq_Inf_inter) - -lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" - by (fact Inf_union_distrib) - -lemma Inter_UNIV_conv [simp, no_atp]: - "\A = UNIV \ (\x\A. x = UNIV)" - "UNIV = \A \ (\x\A. x = UNIV)" - by (fact Inf_top_conv)+ - -lemma Inter_anti_mono: "B \ A \ \A \ \B" - by (fact Inf_superset_mono) - - -subsection {* Intersections of families *} - -abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" where - "INTER \ INFI" - -text {* - Note: must use name @{const INTER} here instead of @{text INT} - to allow the following syntax coexist with the plain constant name. -*} - -syntax - "_INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) - "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 0, 10] 10) - -syntax (xsymbols) - "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) - "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) - -syntax (latex output) - "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) - "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) - -translations - "INT x y. B" == "INT x. INT y. B" - "INT x. B" == "CONST INTER CONST UNIV (%x. B)" - "INT x. B" == "INT x:CONST UNIV. B" - "INT x:A. B" == "CONST INTER A (%x. B)" - -print_translation {* - [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}] -*} -- {* to avoid eta-contraction of body *} - -lemma INTER_eq: - "(\x\A. B x) = {y. \x\A. y \ B x}" - by (auto simp add: INF_def) - -lemma Inter_image_eq [simp]: - "\(B`A) = (\x\A. B x)" - by (rule sym) (fact INF_def) - -lemma INT_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" - by (auto simp add: INF_def image_def) - -lemma INT_I [intro!]: "(\x. x \ A \ b \ B x) \ b \ (\x\A. B x)" - by (auto simp add: INF_def image_def) - -lemma INT_D [elim, Pure.elim]: "b \ (\x\A. B x) \ a \ A \ b \ B a" - by auto - -lemma INT_E [elim]: "b \ (\x\A. B x) \ (b \ B a \ R) \ (a \ A \ R) \ R" - -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a\A"}. *} - by (auto simp add: INF_def image_def) - -lemma INT_cong [cong]: - "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (fact INF_cong) - -lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" - by blast - -lemma Collect_all_eq: "{x. \y. P x y} = (\y. {x. P x y})" - by blast - -lemma INT_lower: "a \ A \ (\x\A. B x) \ B a" - by (fact INF_lower) - -lemma INT_greatest: "(\x. x \ A \ C \ B x) \ C \ (\x\A. B x)" - by (fact INF_greatest) - -lemma INT_empty: "(\x\{}. B x) = UNIV" - by (fact INF_empty) - -lemma INT_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" - by (fact INF_absorb) - -lemma INT_subset_iff: "B \ (\i\I. A i) \ (\i\I. B \ A i)" - by (fact le_INF_iff) - -lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" - by (fact INF_insert) - -lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" - by (fact INF_union) - -lemma INT_insert_distrib: - "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" - by blast - -lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" - by (fact INF_constant) - -lemma INTER_UNIV_conv [simp]: - "(UNIV = (\x\A. B x)) = (\x\A. B x = UNIV)" - "((\x\A. B x) = UNIV) = (\x\A. B x = UNIV)" - by (fact INF_top_conv)+ - -lemma INT_bool_eq: "(\b. A b) = A True \ A False" - by (fact INF_UNIV_bool_expand) - -lemma INT_anti_mono: - "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\A. g x)" - -- {* The last inclusion is POSITIVE! *} - by (fact INF_superset_mono) - -lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" - by blast - -lemma vimage_INT: "f -` (\x\A. B x) = (\x\A. f -` B x)" - by blast - - -subsection {* Union *} - -abbreviation Union :: "'a set set \ 'a set" where - "Union S \ \S" - -notation (xsymbols) - Union ("\_" [90] 90) - -lemma Union_eq: - "\A = {x. \B \ A. x \ B}" -proof (rule set_eqI) - fix x - have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" - by auto - then show "x \ \A \ x \ {x. \B\A. x \ B}" - by (simp add: Sup_fun_def) (simp add: mem_def) -qed - -lemma Union_iff [simp, no_atp]: - "A \ \C \ (\X\C. A\X)" - by (unfold Union_eq) blast - -lemma UnionI [intro]: - "X \ C \ A \ X \ A \ \C" - -- {* The order of the premises presupposes that @{term C} is rigid; - @{term A} may be flexible. *} - by auto - -lemma UnionE [elim!]: - "A \ \C \ (\X. A \ X \ X \ C \ R) \ R" - by auto - -lemma Union_upper: "B \ A \ B \ \A" - by (fact Sup_upper) - -lemma Union_least: "(\X. X \ A \ X \ C) \ \A \ C" - by (fact Sup_least) - -lemma Union_empty [simp]: "\{} = {}" - by (fact Sup_empty) - -lemma Union_UNIV [simp]: "\UNIV = UNIV" - by (fact Sup_UNIV) - -lemma Union_insert [simp]: "\insert a B = a \ \B" - by (fact Sup_insert) - -lemma Union_Un_distrib [simp]: "\(A \ B) = \A \ \B" - by (fact Sup_union_distrib) - -lemma Union_Int_subset: "\(A \ B) \ \A \ \B" - by (fact Sup_inter_less_eq) - -lemma Union_empty_conv [simp,no_atp]: "(\A = {}) \ (\x\A. x = {})" - by (fact Sup_bot_conv) - -lemma empty_Union_conv [simp,no_atp]: "({} = \A) \ (\x\A. x = {})" - by (fact Sup_bot_conv) - -lemma subset_Pow_Union: "A \ Pow (\A)" - by blast - -lemma Union_Pow_eq [simp]: "\(Pow A) = A" - by blast - -lemma Union_mono: "A \ B \ \A \ \B" - by (fact Sup_subset_mono) - - -subsection {* Unions of families *} - -abbreviation UNION :: "'a set \ ('a \ 'b set) \ 'b set" where - "UNION \ SUPR" - -text {* - Note: must use name @{const UNION} here instead of @{text UN} - to allow the following syntax coexist with the plain constant name. -*} - -syntax - "_UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) - "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 0, 10] 10) - -syntax (xsymbols) - "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) - "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) - -syntax (latex output) - "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) - "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) - -translations - "UN x y. B" == "UN x. UN y. B" - "UN x. B" == "CONST UNION CONST UNIV (%x. B)" - "UN x. B" == "UN x:CONST UNIV. B" - "UN x:A. B" == "CONST UNION A (%x. B)" - -text {* - Note the difference between ordinary xsymbol syntax of indexed - unions and intersections (e.g.\ @{text"\a\<^isub>1\A\<^isub>1. B"}) - and their \LaTeX\ rendition: @{term"\a\<^isub>1\A\<^isub>1. B"}. The - former does not make the index expression a subscript of the - union/intersection symbol because this leads to problems with nested - subscripts in Proof General. -*} - -print_translation {* - [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax UNION} @{syntax_const "_UNION"}] -*} -- {* to avoid eta-contraction of body *} - -lemma UNION_eq [no_atp]: - "(\x\A. B x) = {y. \x\A. y \ B x}" - by (auto simp add: SUP_def) - -lemma Union_image_eq [simp]: - "\(B ` A) = (\x\A. B x)" - by (auto simp add: UNION_eq) - -lemma UN_iff [simp]: "(b \ (\x\A. B x)) = (\x\A. b \ B x)" - by (auto simp add: SUP_def image_def) - -lemma UN_I [intro]: "a \ A \ b \ B a \ b \ (\x\A. B x)" - -- {* The order of the premises presupposes that @{term A} is rigid; - @{term b} may be flexible. *} - by auto - -lemma UN_E [elim!]: "b \ (\x\A. B x) \ (\x. x\A \ b \ B x \ R) \ R" - by (auto simp add: SUP_def image_def) - -lemma UN_cong [cong]: - "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (fact SUP_cong) - -lemma strong_UN_cong: - "A = B \ (\x. x \ B =simp=> C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (unfold simp_implies_def) (fact UN_cong) - -lemma image_eq_UN: "f ` A = (\x\A. {f x})" - by blast - -lemma UN_upper: "a \ A \ B a \ (\x\A. B x)" - by (fact SUP_upper) - -lemma UN_least: "(\x. x \ A \ B x \ C) \ (\x\A. B x) \ C" - by (fact SUP_least) - -lemma Collect_bex_eq [no_atp]: "{x. \y\A. P x y} = (\y\A. {x. P x y})" - by blast - -lemma UN_insert_distrib: "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" - by blast - -lemma UN_empty [no_atp]: "(\x\{}. B x) = {}" - by (fact SUP_empty) - -lemma UN_empty2 [simp]: "(\x\A. {}) = {}" - by (fact SUP_bot) - -lemma UN_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" - by (fact SUP_absorb) - -lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" - by (fact SUP_insert) - -lemma UN_Un [simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" - by (fact SUP_union) - -lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" - by blast - -lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" - by (fact SUP_le_iff) - -lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" - by (fact SUP_constant) - -lemma image_Union: "f ` \S = (\x\S. f ` x)" - by blast - -lemma UNION_empty_conv[simp]: - "{} = (\x\A. B x) \ (\x\A. B x = {})" - "(\x\A. B x) = {} \ (\x\A. B x = {})" - by (fact SUP_bot_conv)+ - -lemma Collect_ex_eq [no_atp]: "{x. \y. P x y} = (\y. {x. P x y})" - by blast - -lemma ball_UN: "(\z \ UNION A B. P z) \ (\x\A. \z \ B x. P z)" - by blast - -lemma bex_UN: "(\z \ UNION A B. P z) \ (\x\A. \z\B x. P z)" - by blast - -lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" - by (auto simp add: split_if_mem2) - -lemma UN_bool_eq: "(\b. A b) = (A True \ A False)" - by (fact SUP_UNIV_bool_expand) - -lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" - by blast - -lemma UN_mono: - "A \ B \ (\x. x \ A \ f x \ g x) \ - (\x\A. f x) \ (\x\B. g x)" - by (fact SUP_subset_mono) - -lemma vimage_Union: "f -` (\A) = (\X\A. f -` X)" - by blast - -lemma vimage_UN: "f -` (\x\A. B x) = (\x\A. f -` B x)" - by blast - -lemma vimage_eq_UN: "f -` B = (\y\B. f -` {y})" - -- {* NOT suitable for rewriting *} - by blast - -lemma image_UN: "f ` UNION A B = (\x\A. f ` B x)" - by blast - - -subsection {* Distributive laws *} - -lemma Int_Union: "A \ \B = (\C\B. A \ C)" - by (fact inf_Sup) - -lemma Un_Inter: "A \ \B = (\C\B. A \ C)" - by (fact sup_Inf) - -lemma Int_Union2: "\B \ A = (\C\B. C \ A)" - by (fact Sup_inf) - -lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" - by (rule sym) (rule INF_inf_distrib) - -lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" - by (rule sym) (rule SUP_sup_distrib) - -lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" - by (simp only: INT_Int_distrib INF_def) - -lemma Un_Union_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" - -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *} - -- {* Union of a family of unions *} - by (simp only: UN_Un_distrib SUP_def) - -lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" - by (fact sup_INF) - -lemma Int_UN_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" - -- {* Halmos, Naive Set Theory, page 35. *} - by (fact inf_SUP) - -lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" - by (fact SUP_inf_distrib2) - -lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" - by (fact INF_sup_distrib2) - -lemma Union_disjoint: "(\C \ A = {}) \ (\B\C. B \ A = {})" - by (fact Sup_inf_eq_bot_iff) - - -subsection {* Complement *} - -lemma Compl_INT [simp]: "- (\x\A. B x) = (\x\A. -B x)" - by (fact uminus_INF) - -lemma Compl_UN [simp]: "- (\x\A. B x) = (\x\A. -B x)" - by (fact uminus_SUP) - - -subsection {* Miniscoping and maxiscoping *} - -text {* \medskip Miniscoping: pushing in quantifiers and big Unions - and Intersections. *} - -lemma UN_simps [simp]: - "\a B C. (\x\C. insert a (B x)) = (if C={} then {} else insert a (\x\C. B x))" - "\A B C. (\x\C. A x \ B) = ((if C={} then {} else (\x\C. A x) \ B))" - "\A B C. (\x\C. A \ B x) = ((if C={} then {} else A \ (\x\C. B x)))" - "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \ B)" - "\A B C. (\x\C. A \ B x) = (A \(\x\C. B x))" - "\A B C. (\x\C. A x - B) = ((\x\C. A x) - B)" - "\A B C. (\x\C. A - B x) = (A - (\x\C. B x))" - "\A B. (\x\\A. B x) = (\y\A. \x\y. B x)" - "\A B C. (\z\UNION A B. C z) = (\x\A. \z\B x. C z)" - "\A B f. (\x\f`A. B x) = (\a\A. B (f a))" - by auto - -lemma INT_simps [simp]: - "\A B C. (\x\C. A x \ B) = (if C={} then UNIV else (\x\C. A x) \ B)" - "\A B C. (\x\C. A \ B x) = (if C={} then UNIV else A \(\x\C. B x))" - "\A B C. (\x\C. A x - B) = (if C={} then UNIV else (\x\C. A x) - B)" - "\A B C. (\x\C. A - B x) = (if C={} then UNIV else A - (\x\C. B x))" - "\a B C. (\x\C. insert a (B x)) = insert a (\x\C. B x)" - "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \ B)" - "\A B C. (\x\C. A \ B x) = (A \ (\x\C. B x))" - "\A B. (\x\\A. B x) = (\y\A. \x\y. B x)" - "\A B C. (\z\UNION A B. C z) = (\x\A. \z\B x. C z)" - "\A B f. (\x\f`A. B x) = (\a\A. B (f a))" - by auto - -lemma UN_ball_bex_simps [simp, no_atp]: - "\A P. (\x\\A. P x) \ (\y\A. \x\y. P x)" - "\A B P. (\x\UNION A B. P x) = (\a\A. \x\ B a. P x)" - "\A P. (\x\\A. P x) \ (\y\A. \x\y. P x)" - "\A B P. (\x\UNION A B. P x) \ (\a\A. \x\B a. P x)" - by auto - - -text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *} - -lemma UN_extend_simps: - "\a B C. insert a (\x\C. B x) = (if C={} then {a} else (\x\C. insert a (B x)))" - "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" - "\A B C. A \ (\x\C. B x) = (if C={} then A else (\x\C. A \ B x))" - "\A B C. ((\x\C. A x) \ B) = (\x\C. A x \ B)" - "\A B C. (A \ (\x\C. B x)) = (\x\C. A \ B x)" - "\A B C. ((\x\C. A x) - B) = (\x\C. A x - B)" - "\A B C. (A - (\x\C. B x)) = (\x\C. A - B x)" - "\A B. (\y\A. \x\y. B x) = (\x\\A. B x)" - "\A B C. (\x\A. \z\B x. C z) = (\z\UNION A B. C z)" - "\A B f. (\a\A. B (f a)) = (\x\f`A. B x)" - by auto - -lemma INT_extend_simps: - "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" - "\A B C. A \ (\x\C. B x) = (if C={} then A else (\x\C. A \ B x))" - "\A B C. (\x\C. A x) - B = (if C={} then UNIV - B else (\x\C. A x - B))" - "\A B C. A - (\x\C. B x) = (if C={} then A else (\x\C. A - B x))" - "\a B C. insert a (\x\C. B x) = (\x\C. insert a (B x))" - "\A B C. ((\x\C. A x) \ B) = (\x\C. A x \ B)" - "\A B C. A \ (\x\C. B x) = (\x\C. A \ B x)" - "\A B. (\y\A. \x\y. B x) = (\x\\A. B x)" - "\A B C. (\x\A. \z\B x. C z) = (\z\UNION A B. C z)" - "\A B f. (\a\A. B (f a)) = (\x\f`A. B x)" - by auto - - -text {* Legacy names *} - -lemma (in complete_lattice) Inf_singleton [simp]: - "\{a} = a" - by (simp add: Inf_insert) - -lemma (in complete_lattice) Sup_singleton [simp]: - "\{a} = a" - by (simp add: Sup_insert) - -lemma (in complete_lattice) Inf_binary: - "\{a, b} = a \ b" - by (simp add: Inf_insert) - -lemma (in complete_lattice) Sup_binary: - "\{a, b} = a \ b" - by (simp add: Sup_insert) - -lemmas (in complete_lattice) INFI_def = INF_def -lemmas (in complete_lattice) SUPR_def = SUP_def -lemmas (in complete_lattice) INF_leI = INF_lower -lemmas (in complete_lattice) INF_leI2 = INF_lower2 -lemmas (in complete_lattice) le_INFI = INF_greatest -lemmas (in complete_lattice) le_SUPI = SUP_upper -lemmas (in complete_lattice) le_SUPI2 = SUP_upper2 -lemmas (in complete_lattice) SUP_leI = SUP_least -lemmas (in complete_lattice) less_INFD = less_INF_D - -lemmas INFI_apply = INF_apply -lemmas SUPR_apply = SUP_apply - -text {* Grep and put to news from here *} - -lemma (in complete_lattice) INF_eq: - "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by (simp add: INF_def image_def) - -lemma (in complete_lattice) SUP_eq: - "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by (simp add: SUP_def image_def) - -lemma (in complete_lattice) INF_subset: - "B \ A \ INFI A f \ INFI B f" - by (rule INF_superset_mono) auto - -lemma (in complete_lattice) INF_UNIV_range: - "(\x. f x) = \range f" - by (fact INF_def) - -lemma (in complete_lattice) SUP_UNIV_range: - "(\x. f x) = \range f" - by (fact SUP_def) - -lemma Int_eq_Inter: "A \ B = \{A, B}" - by (simp add: Inf_insert) - -lemma INTER_eq_Inter_image: - "(\x\A. B x) = \(B`A)" - by (fact INF_def) - -lemma Inter_def: - "\S = (\x\S. x)" - by (simp add: INTER_eq_Inter_image image_def) - -lemmas INTER_def = INTER_eq -lemmas UNION_def = UNION_eq - -lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by (fact INF_eq) - -lemma Un_eq_Union: "A \ B = \{A, B}" - by blast - -lemma UNION_eq_Union_image: - "(\x\A. B x) = \(B ` A)" - by (fact SUP_def) - -lemma Union_def: - "\S = (\x\S. x)" - by (simp add: UNION_eq_Union_image image_def) - -lemma UN_singleton [simp]: "(\x\A. {x}) = A" - by blast - -lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by (fact SUP_eq) - - -text {* Finally *} - -no_notation - less_eq (infix "\" 50) and - less (infix "\" 50) and - bot ("\") and - top ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) - -no_syntax (xsymbols) - "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - -lemmas mem_simps = - insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff - mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff - -- {* Each of these has ALREADY been added @{text "[simp]"} above. *} - -end diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Complete_Lattices.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Complete_Lattices.thy Sat Sep 10 10:29:24 2011 +0200 @@ -0,0 +1,1259 @@ + (* Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel; Florian Haftmann, TU Muenchen *) + +header {* Complete lattices *} + +theory Complete_Lattices +imports Set +begin + +notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + inf (infixl "\" 70) and + sup (infixl "\" 65) and + top ("\") and + bot ("\") + + +subsection {* Syntactic infimum and supremum operations *} + +class Inf = + fixes Inf :: "'a set \ 'a" ("\_" [900] 900) + +class Sup = + fixes Sup :: "'a set \ 'a" ("\_" [900] 900) + +subsection {* Abstract complete lattices *} + +class complete_lattice = bounded_lattice + Inf + Sup + + assumes Inf_lower: "x \ A \ \A \ x" + and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" + assumes Sup_upper: "x \ A \ x \ \A" + and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" +begin + +lemma dual_complete_lattice: + "class.complete_lattice Sup Inf sup (op \) (op >) inf \ \" + by (auto intro!: class.complete_lattice.intro dual_bounded_lattice) + (unfold_locales, (fact bot_least top_greatest + Sup_upper Sup_least Inf_lower Inf_greatest)+) + +definition INFI :: "'b set \ ('b \ 'a) \ 'a" where + INF_def: "INFI A f = \(f ` A)" + +definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where + SUP_def: "SUPR A f = \(f ` A)" + +text {* + Note: must use names @{const INFI} and @{const SUPR} here instead of + @{text INF} and @{text SUP} to allow the following syntax coexist + with the plain constant names. +*} + +end + +syntax + "_INF1" :: "pttrns \ 'b \ 'b" ("(3INF _./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3INF _:_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3SUP _./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3SUP _:_./ _)" [0, 0, 10] 10) + +syntax (xsymbols) + "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + +translations + "INF x y. B" == "INF x. INF y. B" + "INF x. B" == "CONST INFI CONST UNIV (%x. B)" + "INF x. B" == "INF x:CONST UNIV. B" + "INF x:A. B" == "CONST INFI A (%x. B)" + "SUP x y. B" == "SUP x. SUP y. B" + "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)" + "SUP x. B" == "SUP x:CONST UNIV. B" + "SUP x:A. B" == "CONST SUPR A (%x. B)" + +print_translation {* + [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INFI} @{syntax_const "_INF"}, + Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax SUPR} @{syntax_const "_SUP"}] +*} -- {* to avoid eta-contraction of body *} + +context complete_lattice +begin + +lemma INF_foundation_dual [no_atp]: + "complete_lattice.SUPR Inf = INFI" +proof (rule ext)+ + interpret dual: complete_lattice Sup Inf sup "op \" "op >" inf \ \ + by (fact dual_complete_lattice) + fix f :: "'b \ 'a" and A + show "complete_lattice.SUPR Inf A f = (\a\A. f a)" + by (simp only: dual.SUP_def INF_def) +qed + +lemma SUP_foundation_dual [no_atp]: + "complete_lattice.INFI Sup = SUPR" +proof (rule ext)+ + interpret dual: complete_lattice Sup Inf sup "op \" "op >" inf \ \ + by (fact dual_complete_lattice) + fix f :: "'b \ 'a" and A + show "complete_lattice.INFI Sup A f = (\a\A. f a)" + by (simp only: dual.INF_def SUP_def) +qed + +lemma INF_lower: "i \ A \ (\i\A. f i) \ f i" + by (auto simp add: INF_def intro: Inf_lower) + +lemma INF_greatest: "(\i. i \ A \ u \ f i) \ u \ (\i\A. f i)" + by (auto simp add: INF_def intro: Inf_greatest) + +lemma SUP_upper: "i \ A \ f i \ (\i\A. f i)" + by (auto simp add: SUP_def intro: Sup_upper) + +lemma SUP_least: "(\i. i \ A \ f i \ u) \ (\i\A. f i) \ u" + by (auto simp add: SUP_def intro: Sup_least) + +lemma Inf_lower2: "u \ A \ u \ v \ \A \ v" + using Inf_lower [of u A] by auto + +lemma INF_lower2: "i \ A \ f i \ u \ (\i\A. f i) \ u" + using INF_lower [of i A f] by auto + +lemma Sup_upper2: "u \ A \ v \ u \ v \ \A" + using Sup_upper [of u A] by auto + +lemma SUP_upper2: "i \ A \ u \ f i \ u \ (\i\A. f i)" + using SUP_upper [of i A f] by auto + +lemma le_Inf_iff (*[simp]*): "b \ \A \ (\a\A. b \ a)" + by (auto intro: Inf_greatest dest: Inf_lower) + +lemma le_INF_iff (*[simp]*): "u \ (\i\A. f i) \ (\i\A. u \ f i)" + by (auto simp add: INF_def le_Inf_iff) + +lemma Sup_le_iff (*[simp]*): "\A \ b \ (\a\A. a \ b)" + by (auto intro: Sup_least dest: Sup_upper) + +lemma SUP_le_iff (*[simp]*): "(\i\A. f i) \ u \ (\i\A. f i \ u)" + by (auto simp add: SUP_def Sup_le_iff) + +lemma Inf_empty [simp]: + "\{} = \" + by (auto intro: antisym Inf_greatest) + +lemma INF_empty [simp]: "(\x\{}. f x) = \" + by (simp add: INF_def) + +lemma Sup_empty [simp]: + "\{} = \" + by (auto intro: antisym Sup_least) + +lemma SUP_empty [simp]: "(\x\{}. f x) = \" + by (simp add: SUP_def) + +lemma Inf_UNIV [simp]: + "\UNIV = \" + by (auto intro!: antisym Inf_lower) + +lemma Sup_UNIV [simp]: + "\UNIV = \" + by (auto intro!: antisym Sup_upper) + +lemma Inf_insert (*[simp]*): "\insert a A = a \ \A" + by (auto intro: le_infI le_infI1 le_infI2 antisym Inf_greatest Inf_lower) + +lemma INF_insert: "(\x\insert a A. f x) = f a \ INFI A f" + by (simp add: INF_def Inf_insert) + +lemma Sup_insert (*[simp]*): "\insert a A = a \ \A" + by (auto intro: le_supI le_supI1 le_supI2 antisym Sup_least Sup_upper) + +lemma SUP_insert: "(\x\insert a A. f x) = f a \ SUPR A f" + by (simp add: SUP_def Sup_insert) + +lemma INF_image (*[simp]*): "(\x\f`A. g x) = (\x\A. g (f x))" + by (simp add: INF_def image_image) + +lemma SUP_image (*[simp]*): "(\x\f`A. g x) = (\x\A. g (f x))" + by (simp add: SUP_def image_image) + +lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" + by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) + +lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" + by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) + +lemma Inf_superset_mono: "B \ A \ \A \ \B" + by (auto intro: Inf_greatest Inf_lower) + +lemma Sup_subset_mono: "A \ B \ \A \ \B" + by (auto intro: Sup_least Sup_upper) + +lemma INF_cong: + "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" + by (simp add: INF_def image_def) + +lemma SUP_cong: + "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" + by (simp add: SUP_def image_def) + +lemma Inf_mono: + assumes "\b. b \ B \ \a\A. a \ b" + shows "\A \ \B" +proof (rule Inf_greatest) + fix b assume "b \ B" + with assms obtain a where "a \ A" and "a \ b" by blast + from `a \ A` have "\A \ a" by (rule Inf_lower) + with `a \ b` show "\A \ b" by auto +qed + +lemma INF_mono: + "(\m. m \ B \ \n\A. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" + by (force intro!: Inf_mono simp: INF_def) + +lemma Sup_mono: + assumes "\a. a \ A \ \b\B. a \ b" + shows "\A \ \B" +proof (rule Sup_least) + fix a assume "a \ A" + with assms obtain b where "b \ B" and "a \ b" by blast + from `b \ B` have "b \ \B" by (rule Sup_upper) + with `a \ b` show "a \ \B" by auto +qed + +lemma SUP_mono: + "(\n. n \ A \ \m\B. f n \ g m) \ (\n\A. f n) \ (\n\B. g n)" + by (force intro!: Sup_mono simp: SUP_def) + +lemma INF_superset_mono: + "B \ A \ (\x. x \ B \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" + -- {* The last inclusion is POSITIVE! *} + by (blast intro: INF_mono dest: subsetD) + +lemma SUP_subset_mono: + "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" + by (blast intro: SUP_mono dest: subsetD) + +lemma Inf_less_eq: + assumes "\v. v \ A \ v \ u" + and "A \ {}" + shows "\A \ u" +proof - + from `A \ {}` obtain v where "v \ A" by blast + moreover with assms have "v \ u" by blast + ultimately show ?thesis by (rule Inf_lower2) +qed + +lemma less_eq_Sup: + assumes "\v. v \ A \ u \ v" + and "A \ {}" + shows "u \ \A" +proof - + from `A \ {}` obtain v where "v \ A" by blast + moreover with assms have "u \ v" by blast + ultimately show ?thesis by (rule Sup_upper2) +qed + +lemma less_eq_Inf_inter: "\A \ \B \ \(A \ B)" + by (auto intro: Inf_greatest Inf_lower) + +lemma Sup_inter_less_eq: "\(A \ B) \ \A \ \B " + by (auto intro: Sup_least Sup_upper) + +lemma Inf_union_distrib: "\(A \ B) = \A \ \B" + by (rule antisym) (auto intro: Inf_greatest Inf_lower le_infI1 le_infI2) + +lemma INF_union: + "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" + by (auto intro!: antisym INF_mono intro: le_infI1 le_infI2 INF_greatest INF_lower) + +lemma Sup_union_distrib: "\(A \ B) = \A \ \B" + by (rule antisym) (auto intro: Sup_least Sup_upper le_supI1 le_supI2) + +lemma SUP_union: + "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" + by (auto intro!: antisym SUP_mono intro: le_supI1 le_supI2 SUP_least SUP_upper) + +lemma INF_inf_distrib: "(\a\A. f a) \ (\a\A. g a) = (\a\A. f a \ g a)" + by (rule antisym) (rule INF_greatest, auto intro: le_infI1 le_infI2 INF_lower INF_mono) + +lemma SUP_sup_distrib: "(\a\A. f a) \ (\a\A. g a) = (\a\A. f a \ g a)" + by (rule antisym) (auto intro: le_supI1 le_supI2 SUP_upper SUP_mono, + rule SUP_least, auto intro: le_supI1 le_supI2 SUP_upper SUP_mono) + +lemma Inf_top_conv (*[simp]*) [no_atp]: + "\A = \ \ (\x\A. x = \)" + "\ = \A \ (\x\A. x = \)" +proof - + show "\A = \ \ (\x\A. x = \)" + proof + assume "\x\A. x = \" + then have "A = {} \ A = {\}" by auto + then show "\A = \" by (auto simp add: Inf_insert) + next + assume "\A = \" + show "\x\A. x = \" + proof (rule ccontr) + assume "\ (\x\A. x = \)" + then obtain x where "x \ A" and "x \ \" by blast + then obtain B where "A = insert x B" by blast + with `\A = \` `x \ \` show False by (simp add: Inf_insert) + qed + qed + then show "\ = \A \ (\x\A. x = \)" by auto +qed + +lemma INF_top_conv (*[simp]*): + "(\x\A. B x) = \ \ (\x\A. B x = \)" + "\ = (\x\A. B x) \ (\x\A. B x = \)" + by (auto simp add: INF_def Inf_top_conv) + +lemma Sup_bot_conv (*[simp]*) [no_atp]: + "\A = \ \ (\x\A. x = \)" (is ?P) + "\ = \A \ (\x\A. x = \)" (is ?Q) +proof - + interpret dual: complete_lattice Sup Inf sup "op \" "op >" inf \ \ + by (fact dual_complete_lattice) + from dual.Inf_top_conv show ?P and ?Q by simp_all +qed + +lemma SUP_bot_conv (*[simp]*): + "(\x\A. B x) = \ \ (\x\A. B x = \)" + "\ = (\x\A. B x) \ (\x\A. B x = \)" + by (auto simp add: SUP_def Sup_bot_conv) + +lemma INF_const [simp]: "A \ {} \ (\i\A. f) = f" + by (auto intro: antisym INF_lower INF_greatest) + +lemma SUP_const [simp]: "A \ {} \ (\i\A. f) = f" + by (auto intro: antisym SUP_upper SUP_least) + +lemma INF_top (*[simp]*): "(\x\A. \) = \" + by (cases "A = {}") (simp_all add: INF_empty) + +lemma SUP_bot (*[simp]*): "(\x\A. \) = \" + by (cases "A = {}") (simp_all add: SUP_empty) + +lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" + by (iprover intro: INF_lower INF_greatest order_trans antisym) + +lemma SUP_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" + by (iprover intro: SUP_upper SUP_least order_trans antisym) + +lemma INF_absorb: + assumes "k \ I" + shows "A k \ (\i\I. A i) = (\i\I. A i)" +proof - + from assms obtain J where "I = insert k J" by blast + then show ?thesis by (simp add: INF_insert) +qed + +lemma SUP_absorb: + assumes "k \ I" + shows "A k \ (\i\I. A i) = (\i\I. A i)" +proof - + from assms obtain J where "I = insert k J" by blast + then show ?thesis by (simp add: SUP_insert) +qed + +lemma INF_constant: + "(\y\A. c) = (if A = {} then \ else c)" + by (simp add: INF_empty) + +lemma SUP_constant: + "(\y\A. c) = (if A = {} then \ else c)" + by (simp add: SUP_empty) + +lemma less_INF_D: + assumes "y < (\i\A. f i)" "i \ A" shows "y < f i" +proof - + note `y < (\i\A. f i)` + also have "(\i\A. f i) \ f i" using `i \ A` + by (rule INF_lower) + finally show "y < f i" . +qed + +lemma SUP_lessD: + assumes "(\i\A. f i) < y" "i \ A" shows "f i < y" +proof - + have "f i \ (\i\A. f i)" using `i \ A` + by (rule SUP_upper) + also note `(\i\A. f i) < y` + finally show "f i < y" . +qed + +lemma INF_UNIV_bool_expand: + "(\b. A b) = A True \ A False" + by (simp add: UNIV_bool INF_empty INF_insert inf_commute) + +lemma SUP_UNIV_bool_expand: + "(\b. A b) = A True \ A False" + by (simp add: UNIV_bool SUP_empty SUP_insert sup_commute) + +end + +class complete_distrib_lattice = complete_lattice + + assumes sup_Inf: "a \ \B = (\b\B. a \ b)" + assumes inf_Sup: "a \ \B = (\b\B. a \ b)" +begin + +lemma sup_INF: + "a \ (\b\B. f b) = (\b\B. a \ f b)" + by (simp add: INF_def sup_Inf image_image) + +lemma inf_SUP: + "a \ (\b\B. f b) = (\b\B. a \ f b)" + by (simp add: SUP_def inf_Sup image_image) + +lemma dual_complete_distrib_lattice: + "class.complete_distrib_lattice Sup Inf sup (op \) (op >) inf \ \" + apply (rule class.complete_distrib_lattice.intro) + apply (fact dual_complete_lattice) + apply (rule class.complete_distrib_lattice_axioms.intro) + apply (simp_all only: INF_foundation_dual SUP_foundation_dual inf_Sup sup_Inf) + done + +subclass distrib_lattice proof + fix a b c + from sup_Inf have "a \ \{b, c} = (\d\{b, c}. a \ d)" . + then show "a \ b \ c = (a \ b) \ (a \ c)" by (simp add: INF_def Inf_insert) +qed + +lemma Inf_sup: + "\B \ a = (\b\B. b \ a)" + by (simp add: sup_Inf sup_commute) + +lemma Sup_inf: + "\B \ a = (\b\B. b \ a)" + by (simp add: inf_Sup inf_commute) + +lemma INF_sup: + "(\b\B. f b) \ a = (\b\B. f b \ a)" + by (simp add: sup_INF sup_commute) + +lemma SUP_inf: + "(\b\B. f b) \ a = (\b\B. f b \ a)" + by (simp add: inf_SUP inf_commute) + +lemma Inf_sup_eq_top_iff: + "(\B \ a = \) \ (\b\B. b \ a = \)" + by (simp only: Inf_sup INF_top_conv) + +lemma Sup_inf_eq_bot_iff: + "(\B \ a = \) \ (\b\B. b \ a = \)" + by (simp only: Sup_inf SUP_bot_conv) + +lemma INF_sup_distrib2: + "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" + by (subst INF_commute) (simp add: sup_INF INF_sup) + +lemma SUP_inf_distrib2: + "(\a\A. f a) \ (\b\B. g b) = (\a\A. \b\B. f a \ g b)" + by (subst SUP_commute) (simp add: inf_SUP SUP_inf) + +end + +class complete_boolean_algebra = boolean_algebra + complete_distrib_lattice +begin + +lemma dual_complete_boolean_algebra: + "class.complete_boolean_algebra Sup Inf sup (op \) (op >) inf \ \ (\x y. x \ - y) uminus" + by (rule class.complete_boolean_algebra.intro, rule dual_complete_distrib_lattice, rule dual_boolean_algebra) + +lemma uminus_Inf: + "- (\A) = \(uminus ` A)" +proof (rule antisym) + show "- \A \ \(uminus ` A)" + by (rule compl_le_swap2, rule Inf_greatest, rule compl_le_swap2, rule Sup_upper) simp + show "\(uminus ` A) \ - \A" + by (rule Sup_least, rule compl_le_swap1, rule Inf_lower) auto +qed + +lemma uminus_INF: "- (\x\A. B x) = (\x\A. - B x)" + by (simp add: INF_def SUP_def uminus_Inf image_image) + +lemma uminus_Sup: + "- (\A) = \(uminus ` A)" +proof - + have "\A = - \(uminus ` A)" by (simp add: image_image uminus_Inf) + then show ?thesis by simp +qed + +lemma uminus_SUP: "- (\x\A. B x) = (\x\A. - B x)" + by (simp add: INF_def SUP_def uminus_Sup image_image) + +end + +class complete_linorder = linorder + complete_lattice +begin + +lemma dual_complete_linorder: + "class.complete_linorder Sup Inf sup (op \) (op >) inf \ \" + by (rule class.complete_linorder.intro, rule dual_complete_lattice, rule dual_linorder) + +lemma Inf_less_iff (*[simp]*): + "\S \ a \ (\x\S. x \ a)" + unfolding not_le [symmetric] le_Inf_iff by auto + +lemma INF_less_iff (*[simp]*): + "(\i\A. f i) \ a \ (\x\A. f x \ a)" + unfolding INF_def Inf_less_iff by auto + +lemma less_Sup_iff (*[simp]*): + "a \ \S \ (\x\S. a \ x)" + unfolding not_le [symmetric] Sup_le_iff by auto + +lemma less_SUP_iff (*[simp]*): + "a \ (\i\A. f i) \ (\x\A. a \ f x)" + unfolding SUP_def less_Sup_iff by auto + +lemma Sup_eq_top_iff (*[simp]*): + "\A = \ \ (\x<\. \i\A. x < i)" +proof + assume *: "\A = \" + show "(\x<\. \i\A. x < i)" unfolding * [symmetric] + proof (intro allI impI) + fix x assume "x < \A" then show "\i\A. x < i" + unfolding less_Sup_iff by auto + qed +next + assume *: "\x<\. \i\A. x < i" + show "\A = \" + proof (rule ccontr) + assume "\A \ \" + with top_greatest [of "\A"] + have "\A < \" unfolding le_less by auto + then have "\A < \A" + using * unfolding less_Sup_iff by auto + then show False by auto + qed +qed + +lemma SUP_eq_top_iff (*[simp]*): + "(\i\A. f i) = \ \ (\x<\. \i\A. x < f i)" + unfolding SUP_def Sup_eq_top_iff by auto + +lemma Inf_eq_bot_iff (*[simp]*): + "\A = \ \ (\x>\. \i\A. i < x)" +proof - + interpret dual: complete_linorder Sup Inf sup "op \" "op >" inf \ \ + by (fact dual_complete_linorder) + from dual.Sup_eq_top_iff show ?thesis . +qed + +lemma INF_eq_bot_iff (*[simp]*): + "(\i\A. f i) = \ \ (\x>\. \i\A. f i < x)" + unfolding INF_def Inf_eq_bot_iff by auto + +end + + +subsection {* @{typ bool} and @{typ "_ \ _"} as complete lattice *} + +instantiation bool :: complete_lattice +begin + +definition + [simp]: "\A \ False \ A" + +definition + [simp]: "\A \ True \ A" + +instance proof +qed (auto intro: bool_induct) + +end + +lemma INF_bool_eq [simp]: + "INFI = Ball" +proof (rule ext)+ + fix A :: "'a set" + fix P :: "'a \ bool" + show "(\x\A. P x) \ (\x\A. P x)" + by (auto simp add: INF_def) +qed + +lemma SUP_bool_eq [simp]: + "SUPR = Bex" +proof (rule ext)+ + fix A :: "'a set" + fix P :: "'a \ bool" + show "(\x\A. P x) \ (\x\A. P x)" + by (auto simp add: SUP_def) +qed + +instance bool :: complete_boolean_algebra proof +qed (auto intro: bool_induct) + +instantiation "fun" :: (type, complete_lattice) complete_lattice +begin + +definition + "\A = (\x. \f\A. f x)" + +lemma Inf_apply: + "(\A) x = (\f\A. f x)" + by (simp add: Inf_fun_def) + +definition + "\A = (\x. \f\A. f x)" + +lemma Sup_apply: + "(\A) x = (\f\A. f x)" + by (simp add: Sup_fun_def) + +instance proof +qed (auto simp add: le_fun_def Inf_apply Sup_apply intro: INF_lower INF_greatest SUP_upper SUP_least) + +end + +lemma INF_apply: + "(\y\A. f y) x = (\y\A. f y x)" + by (auto intro: arg_cong [of _ _ Inf] simp add: INF_def Inf_apply) + +lemma SUP_apply: + "(\y\A. f y) x = (\y\A. f y x)" + by (auto intro: arg_cong [of _ _ Sup] simp add: SUP_def Sup_apply) + +instance "fun" :: (type, complete_distrib_lattice) complete_distrib_lattice proof +qed (auto simp add: inf_apply sup_apply Inf_apply Sup_apply INF_def SUP_def inf_Sup sup_Inf image_image) + +instance "fun" :: (type, complete_boolean_algebra) complete_boolean_algebra .. + + +subsection {* Inter *} + +abbreviation Inter :: "'a set set \ 'a set" where + "Inter S \ \S" + +notation (xsymbols) + Inter ("\_" [90] 90) + +lemma Inter_eq: + "\A = {x. \B \ A. x \ B}" +proof (rule set_eqI) + fix x + have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" + by auto + then show "x \ \A \ x \ {x. \B \ A. x \ B}" + by (simp add: Inf_fun_def) (simp add: mem_def) +qed + +lemma Inter_iff [simp,no_atp]: "A \ \C \ (\X\C. A \ X)" + by (unfold Inter_eq) blast + +lemma InterI [intro!]: "(\X. X \ C \ A \ X) \ A \ \C" + by (simp add: Inter_eq) + +text {* + \medskip A ``destruct'' rule -- every @{term X} in @{term C} + contains @{term A} as an element, but @{prop "A \ X"} can hold when + @{prop "X \ C"} does not! This rule is analogous to @{text spec}. +*} + +lemma InterD [elim, Pure.elim]: "A \ \C \ X \ C \ A \ X" + by auto + +lemma InterE [elim]: "A \ \C \ (X \ C \ R) \ (A \ X \ R) \ R" + -- {* ``Classical'' elimination rule -- does not require proving + @{prop "X \ C"}. *} + by (unfold Inter_eq) blast + +lemma Inter_lower: "B \ A \ \A \ B" + by (fact Inf_lower) + +lemma Inter_subset: + "(\X. X \ A \ X \ B) \ A \ {} \ \A \ B" + by (fact Inf_less_eq) + +lemma Inter_greatest: "(\X. X \ A \ C \ X) \ C \ Inter A" + by (fact Inf_greatest) + +lemma Inter_empty: "\{} = UNIV" + by (fact Inf_empty) (* already simp *) + +lemma Inter_UNIV: "\UNIV = {}" + by (fact Inf_UNIV) (* already simp *) + +lemma Inter_insert [simp]: "\(insert a B) = a \ \B" + by (fact Inf_insert) + +lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" + by (fact less_eq_Inf_inter) + +lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" + by (fact Inf_union_distrib) + +lemma Inter_UNIV_conv [simp, no_atp]: + "\A = UNIV \ (\x\A. x = UNIV)" + "UNIV = \A \ (\x\A. x = UNIV)" + by (fact Inf_top_conv)+ + +lemma Inter_anti_mono: "B \ A \ \A \ \B" + by (fact Inf_superset_mono) + + +subsection {* Intersections of families *} + +abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" where + "INTER \ INFI" + +text {* + Note: must use name @{const INTER} here instead of @{text INT} + to allow the following syntax coexist with the plain constant name. +*} + +syntax + "_INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) + "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 0, 10] 10) + +syntax (xsymbols) + "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) + "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) + +syntax (latex output) + "_INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) + "_INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) + +translations + "INT x y. B" == "INT x. INT y. B" + "INT x. B" == "CONST INTER CONST UNIV (%x. B)" + "INT x. B" == "INT x:CONST UNIV. B" + "INT x:A. B" == "CONST INTER A (%x. B)" + +print_translation {* + [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax INTER} @{syntax_const "_INTER"}] +*} -- {* to avoid eta-contraction of body *} + +lemma INTER_eq: + "(\x\A. B x) = {y. \x\A. y \ B x}" + by (auto simp add: INF_def) + +lemma Inter_image_eq [simp]: + "\(B`A) = (\x\A. B x)" + by (rule sym) (fact INF_def) + +lemma INT_iff [simp]: "b \ (\x\A. B x) \ (\x\A. b \ B x)" + by (auto simp add: INF_def image_def) + +lemma INT_I [intro!]: "(\x. x \ A \ b \ B x) \ b \ (\x\A. B x)" + by (auto simp add: INF_def image_def) + +lemma INT_D [elim, Pure.elim]: "b \ (\x\A. B x) \ a \ A \ b \ B a" + by auto + +lemma INT_E [elim]: "b \ (\x\A. B x) \ (b \ B a \ R) \ (a \ A \ R) \ R" + -- {* "Classical" elimination -- by the Excluded Middle on @{prop "a\A"}. *} + by (auto simp add: INF_def image_def) + +lemma INT_cong [cong]: + "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" + by (fact INF_cong) + +lemma Collect_ball_eq: "{x. \y\A. P x y} = (\y\A. {x. P x y})" + by blast + +lemma Collect_all_eq: "{x. \y. P x y} = (\y. {x. P x y})" + by blast + +lemma INT_lower: "a \ A \ (\x\A. B x) \ B a" + by (fact INF_lower) + +lemma INT_greatest: "(\x. x \ A \ C \ B x) \ C \ (\x\A. B x)" + by (fact INF_greatest) + +lemma INT_empty: "(\x\{}. B x) = UNIV" + by (fact INF_empty) + +lemma INT_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" + by (fact INF_absorb) + +lemma INT_subset_iff: "B \ (\i\I. A i) \ (\i\I. B \ A i)" + by (fact le_INF_iff) + +lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" + by (fact INF_insert) + +lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" + by (fact INF_union) + +lemma INT_insert_distrib: + "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" + by blast + +lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" + by (fact INF_constant) + +lemma INTER_UNIV_conv [simp]: + "(UNIV = (\x\A. B x)) = (\x\A. B x = UNIV)" + "((\x\A. B x) = UNIV) = (\x\A. B x = UNIV)" + by (fact INF_top_conv)+ + +lemma INT_bool_eq: "(\b. A b) = A True \ A False" + by (fact INF_UNIV_bool_expand) + +lemma INT_anti_mono: + "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\A. g x)" + -- {* The last inclusion is POSITIVE! *} + by (fact INF_superset_mono) + +lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" + by blast + +lemma vimage_INT: "f -` (\x\A. B x) = (\x\A. f -` B x)" + by blast + + +subsection {* Union *} + +abbreviation Union :: "'a set set \ 'a set" where + "Union S \ \S" + +notation (xsymbols) + Union ("\_" [90] 90) + +lemma Union_eq: + "\A = {x. \B \ A. x \ B}" +proof (rule set_eqI) + fix x + have "(\Q\{P. \B\A. P \ x \ B}. Q) \ (\B\A. x \ B)" + by auto + then show "x \ \A \ x \ {x. \B\A. x \ B}" + by (simp add: Sup_fun_def) (simp add: mem_def) +qed + +lemma Union_iff [simp, no_atp]: + "A \ \C \ (\X\C. A\X)" + by (unfold Union_eq) blast + +lemma UnionI [intro]: + "X \ C \ A \ X \ A \ \C" + -- {* The order of the premises presupposes that @{term C} is rigid; + @{term A} may be flexible. *} + by auto + +lemma UnionE [elim!]: + "A \ \C \ (\X. A \ X \ X \ C \ R) \ R" + by auto + +lemma Union_upper: "B \ A \ B \ \A" + by (fact Sup_upper) + +lemma Union_least: "(\X. X \ A \ X \ C) \ \A \ C" + by (fact Sup_least) + +lemma Union_empty [simp]: "\{} = {}" + by (fact Sup_empty) + +lemma Union_UNIV [simp]: "\UNIV = UNIV" + by (fact Sup_UNIV) + +lemma Union_insert [simp]: "\insert a B = a \ \B" + by (fact Sup_insert) + +lemma Union_Un_distrib [simp]: "\(A \ B) = \A \ \B" + by (fact Sup_union_distrib) + +lemma Union_Int_subset: "\(A \ B) \ \A \ \B" + by (fact Sup_inter_less_eq) + +lemma Union_empty_conv [simp,no_atp]: "(\A = {}) \ (\x\A. x = {})" + by (fact Sup_bot_conv) + +lemma empty_Union_conv [simp,no_atp]: "({} = \A) \ (\x\A. x = {})" + by (fact Sup_bot_conv) + +lemma subset_Pow_Union: "A \ Pow (\A)" + by blast + +lemma Union_Pow_eq [simp]: "\(Pow A) = A" + by blast + +lemma Union_mono: "A \ B \ \A \ \B" + by (fact Sup_subset_mono) + + +subsection {* Unions of families *} + +abbreviation UNION :: "'a set \ ('a \ 'b set) \ 'b set" where + "UNION \ SUPR" + +text {* + Note: must use name @{const UNION} here instead of @{text UN} + to allow the following syntax coexist with the plain constant name. +*} + +syntax + "_UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) + "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 0, 10] 10) + +syntax (xsymbols) + "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) + "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 0, 10] 10) + +syntax (latex output) + "_UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) + "_UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 0, 10] 10) + +translations + "UN x y. B" == "UN x. UN y. B" + "UN x. B" == "CONST UNION CONST UNIV (%x. B)" + "UN x. B" == "UN x:CONST UNIV. B" + "UN x:A. B" == "CONST UNION A (%x. B)" + +text {* + Note the difference between ordinary xsymbol syntax of indexed + unions and intersections (e.g.\ @{text"\a\<^isub>1\A\<^isub>1. B"}) + and their \LaTeX\ rendition: @{term"\a\<^isub>1\A\<^isub>1. B"}. The + former does not make the index expression a subscript of the + union/intersection symbol because this leads to problems with nested + subscripts in Proof General. +*} + +print_translation {* + [Syntax_Trans.preserve_binder_abs2_tr' @{const_syntax UNION} @{syntax_const "_UNION"}] +*} -- {* to avoid eta-contraction of body *} + +lemma UNION_eq [no_atp]: + "(\x\A. B x) = {y. \x\A. y \ B x}" + by (auto simp add: SUP_def) + +lemma Union_image_eq [simp]: + "\(B ` A) = (\x\A. B x)" + by (auto simp add: UNION_eq) + +lemma UN_iff [simp]: "(b \ (\x\A. B x)) = (\x\A. b \ B x)" + by (auto simp add: SUP_def image_def) + +lemma UN_I [intro]: "a \ A \ b \ B a \ b \ (\x\A. B x)" + -- {* The order of the premises presupposes that @{term A} is rigid; + @{term b} may be flexible. *} + by auto + +lemma UN_E [elim!]: "b \ (\x\A. B x) \ (\x. x\A \ b \ B x \ R) \ R" + by (auto simp add: SUP_def image_def) + +lemma UN_cong [cong]: + "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" + by (fact SUP_cong) + +lemma strong_UN_cong: + "A = B \ (\x. x \ B =simp=> C x = D x) \ (\x\A. C x) = (\x\B. D x)" + by (unfold simp_implies_def) (fact UN_cong) + +lemma image_eq_UN: "f ` A = (\x\A. {f x})" + by blast + +lemma UN_upper: "a \ A \ B a \ (\x\A. B x)" + by (fact SUP_upper) + +lemma UN_least: "(\x. x \ A \ B x \ C) \ (\x\A. B x) \ C" + by (fact SUP_least) + +lemma Collect_bex_eq [no_atp]: "{x. \y\A. P x y} = (\y\A. {x. P x y})" + by blast + +lemma UN_insert_distrib: "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" + by blast + +lemma UN_empty [no_atp]: "(\x\{}. B x) = {}" + by (fact SUP_empty) + +lemma UN_empty2 [simp]: "(\x\A. {}) = {}" + by (fact SUP_bot) + +lemma UN_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" + by (fact SUP_absorb) + +lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" + by (fact SUP_insert) + +lemma UN_Un [simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" + by (fact SUP_union) + +lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" + by blast + +lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" + by (fact SUP_le_iff) + +lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" + by (fact SUP_constant) + +lemma image_Union: "f ` \S = (\x\S. f ` x)" + by blast + +lemma UNION_empty_conv[simp]: + "{} = (\x\A. B x) \ (\x\A. B x = {})" + "(\x\A. B x) = {} \ (\x\A. B x = {})" + by (fact SUP_bot_conv)+ + +lemma Collect_ex_eq [no_atp]: "{x. \y. P x y} = (\y. {x. P x y})" + by blast + +lemma ball_UN: "(\z \ UNION A B. P z) \ (\x\A. \z \ B x. P z)" + by blast + +lemma bex_UN: "(\z \ UNION A B. P z) \ (\x\A. \z\B x. P z)" + by blast + +lemma Un_eq_UN: "A \ B = (\b. if b then A else B)" + by (auto simp add: split_if_mem2) + +lemma UN_bool_eq: "(\b. A b) = (A True \ A False)" + by (fact SUP_UNIV_bool_expand) + +lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" + by blast + +lemma UN_mono: + "A \ B \ (\x. x \ A \ f x \ g x) \ + (\x\A. f x) \ (\x\B. g x)" + by (fact SUP_subset_mono) + +lemma vimage_Union: "f -` (\A) = (\X\A. f -` X)" + by blast + +lemma vimage_UN: "f -` (\x\A. B x) = (\x\A. f -` B x)" + by blast + +lemma vimage_eq_UN: "f -` B = (\y\B. f -` {y})" + -- {* NOT suitable for rewriting *} + by blast + +lemma image_UN: "f ` UNION A B = (\x\A. f ` B x)" + by blast + + +subsection {* Distributive laws *} + +lemma Int_Union: "A \ \B = (\C\B. A \ C)" + by (fact inf_Sup) + +lemma Un_Inter: "A \ \B = (\C\B. A \ C)" + by (fact sup_Inf) + +lemma Int_Union2: "\B \ A = (\C\B. C \ A)" + by (fact Sup_inf) + +lemma INT_Int_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" + by (rule sym) (rule INF_inf_distrib) + +lemma UN_Un_distrib: "(\i\I. A i \ B i) = (\i\I. A i) \ (\i\I. B i)" + by (rule sym) (rule SUP_sup_distrib) + +lemma Int_Inter_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" + by (simp only: INT_Int_distrib INF_def) + +lemma Un_Union_image: "(\x\C. A x \ B x) = \(A ` C) \ \(B ` C)" + -- {* Devlin, Fundamentals of Contemporary Set Theory, page 12, exercise 5: *} + -- {* Union of a family of unions *} + by (simp only: UN_Un_distrib SUP_def) + +lemma Un_INT_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" + by (fact sup_INF) + +lemma Int_UN_distrib: "B \ (\i\I. A i) = (\i\I. B \ A i)" + -- {* Halmos, Naive Set Theory, page 35. *} + by (fact inf_SUP) + +lemma Int_UN_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" + by (fact SUP_inf_distrib2) + +lemma Un_INT_distrib2: "(\i\I. A i) \ (\j\J. B j) = (\i\I. \j\J. A i \ B j)" + by (fact INF_sup_distrib2) + +lemma Union_disjoint: "(\C \ A = {}) \ (\B\C. B \ A = {})" + by (fact Sup_inf_eq_bot_iff) + + +subsection {* Complement *} + +lemma Compl_INT [simp]: "- (\x\A. B x) = (\x\A. -B x)" + by (fact uminus_INF) + +lemma Compl_UN [simp]: "- (\x\A. B x) = (\x\A. -B x)" + by (fact uminus_SUP) + + +subsection {* Miniscoping and maxiscoping *} + +text {* \medskip Miniscoping: pushing in quantifiers and big Unions + and Intersections. *} + +lemma UN_simps [simp]: + "\a B C. (\x\C. insert a (B x)) = (if C={} then {} else insert a (\x\C. B x))" + "\A B C. (\x\C. A x \ B) = ((if C={} then {} else (\x\C. A x) \ B))" + "\A B C. (\x\C. A \ B x) = ((if C={} then {} else A \ (\x\C. B x)))" + "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \ B)" + "\A B C. (\x\C. A \ B x) = (A \(\x\C. B x))" + "\A B C. (\x\C. A x - B) = ((\x\C. A x) - B)" + "\A B C. (\x\C. A - B x) = (A - (\x\C. B x))" + "\A B. (\x\\A. B x) = (\y\A. \x\y. B x)" + "\A B C. (\z\UNION A B. C z) = (\x\A. \z\B x. C z)" + "\A B f. (\x\f`A. B x) = (\a\A. B (f a))" + by auto + +lemma INT_simps [simp]: + "\A B C. (\x\C. A x \ B) = (if C={} then UNIV else (\x\C. A x) \ B)" + "\A B C. (\x\C. A \ B x) = (if C={} then UNIV else A \(\x\C. B x))" + "\A B C. (\x\C. A x - B) = (if C={} then UNIV else (\x\C. A x) - B)" + "\A B C. (\x\C. A - B x) = (if C={} then UNIV else A - (\x\C. B x))" + "\a B C. (\x\C. insert a (B x)) = insert a (\x\C. B x)" + "\A B C. (\x\C. A x \ B) = ((\x\C. A x) \ B)" + "\A B C. (\x\C. A \ B x) = (A \ (\x\C. B x))" + "\A B. (\x\\A. B x) = (\y\A. \x\y. B x)" + "\A B C. (\z\UNION A B. C z) = (\x\A. \z\B x. C z)" + "\A B f. (\x\f`A. B x) = (\a\A. B (f a))" + by auto + +lemma UN_ball_bex_simps [simp, no_atp]: + "\A P. (\x\\A. P x) \ (\y\A. \x\y. P x)" + "\A B P. (\x\UNION A B. P x) = (\a\A. \x\ B a. P x)" + "\A P. (\x\\A. P x) \ (\y\A. \x\y. P x)" + "\A B P. (\x\UNION A B. P x) \ (\a\A. \x\B a. P x)" + by auto + + +text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *} + +lemma UN_extend_simps: + "\a B C. insert a (\x\C. B x) = (if C={} then {a} else (\x\C. insert a (B x)))" + "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" + "\A B C. A \ (\x\C. B x) = (if C={} then A else (\x\C. A \ B x))" + "\A B C. ((\x\C. A x) \ B) = (\x\C. A x \ B)" + "\A B C. (A \ (\x\C. B x)) = (\x\C. A \ B x)" + "\A B C. ((\x\C. A x) - B) = (\x\C. A x - B)" + "\A B C. (A - (\x\C. B x)) = (\x\C. A - B x)" + "\A B. (\y\A. \x\y. B x) = (\x\\A. B x)" + "\A B C. (\x\A. \z\B x. C z) = (\z\UNION A B. C z)" + "\A B f. (\a\A. B (f a)) = (\x\f`A. B x)" + by auto + +lemma INT_extend_simps: + "\A B C. (\x\C. A x) \ B = (if C={} then B else (\x\C. A x \ B))" + "\A B C. A \ (\x\C. B x) = (if C={} then A else (\x\C. A \ B x))" + "\A B C. (\x\C. A x) - B = (if C={} then UNIV - B else (\x\C. A x - B))" + "\A B C. A - (\x\C. B x) = (if C={} then A else (\x\C. A - B x))" + "\a B C. insert a (\x\C. B x) = (\x\C. insert a (B x))" + "\A B C. ((\x\C. A x) \ B) = (\x\C. A x \ B)" + "\A B C. A \ (\x\C. B x) = (\x\C. A \ B x)" + "\A B. (\y\A. \x\y. B x) = (\x\\A. B x)" + "\A B C. (\x\A. \z\B x. C z) = (\z\UNION A B. C z)" + "\A B f. (\a\A. B (f a)) = (\x\f`A. B x)" + by auto + + +text {* Legacy names *} + +lemma (in complete_lattice) Inf_singleton [simp]: + "\{a} = a" + by (simp add: Inf_insert) + +lemma (in complete_lattice) Sup_singleton [simp]: + "\{a} = a" + by (simp add: Sup_insert) + +lemma (in complete_lattice) Inf_binary: + "\{a, b} = a \ b" + by (simp add: Inf_insert) + +lemma (in complete_lattice) Sup_binary: + "\{a, b} = a \ b" + by (simp add: Sup_insert) + +lemmas (in complete_lattice) INFI_def = INF_def +lemmas (in complete_lattice) SUPR_def = SUP_def +lemmas (in complete_lattice) INF_leI = INF_lower +lemmas (in complete_lattice) INF_leI2 = INF_lower2 +lemmas (in complete_lattice) le_INFI = INF_greatest +lemmas (in complete_lattice) le_SUPI = SUP_upper +lemmas (in complete_lattice) le_SUPI2 = SUP_upper2 +lemmas (in complete_lattice) SUP_leI = SUP_least +lemmas (in complete_lattice) less_INFD = less_INF_D + +lemmas INFI_apply = INF_apply +lemmas SUPR_apply = SUP_apply + +text {* Grep and put to news from here *} + +lemma (in complete_lattice) INF_eq: + "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (simp add: INF_def image_def) + +lemma (in complete_lattice) SUP_eq: + "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (simp add: SUP_def image_def) + +lemma (in complete_lattice) INF_subset: + "B \ A \ INFI A f \ INFI B f" + by (rule INF_superset_mono) auto + +lemma (in complete_lattice) INF_UNIV_range: + "(\x. f x) = \range f" + by (fact INF_def) + +lemma (in complete_lattice) SUP_UNIV_range: + "(\x. f x) = \range f" + by (fact SUP_def) + +lemma Int_eq_Inter: "A \ B = \{A, B}" + by (simp add: Inf_insert) + +lemma INTER_eq_Inter_image: + "(\x\A. B x) = \(B`A)" + by (fact INF_def) + +lemma Inter_def: + "\S = (\x\S. x)" + by (simp add: INTER_eq_Inter_image image_def) + +lemmas INTER_def = INTER_eq +lemmas UNION_def = UNION_eq + +lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (fact INF_eq) + +lemma Un_eq_Union: "A \ B = \{A, B}" + by blast + +lemma UNION_eq_Union_image: + "(\x\A. B x) = \(B ` A)" + by (fact SUP_def) + +lemma Union_def: + "\S = (\x\S. x)" + by (simp add: UNION_eq_Union_image image_def) + +lemma UN_singleton [simp]: "(\x\A. {x}) = A" + by blast + +lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (fact SUP_eq) + + +text {* Finally *} + +no_notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + bot ("\") and + top ("\") and + inf (infixl "\" 70) and + sup (infixl "\" 65) and + Inf ("\_" [900] 900) and + Sup ("\_" [900] 900) + +no_syntax (xsymbols) + "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) + "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) + +lemmas mem_simps = + insert_iff empty_iff Un_iff Int_iff Compl_iff Diff_iff + mem_Collect_eq UN_iff Union_iff INT_iff Inter_iff + -- {* Each of these has ALREADY been added @{text "[simp]"} above. *} + +end diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Sat Sep 10 00:44:25 2011 +0200 +++ b/src/HOL/Fun.thy Sat Sep 10 10:29:24 2011 +0200 @@ -6,7 +6,7 @@ header {* Notions about functions *} theory Fun -imports Complete_Lattice +imports Complete_Lattices uses ("Tools/enriched_type.ML") begin diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Import/Generate-HOLLight/GenHOLLight.thy --- a/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Sat Sep 10 00:44:25 2011 +0200 +++ b/src/HOL/Import/Generate-HOLLight/GenHOLLight.thy Sat Sep 10 10:29:24 2011 +0200 @@ -146,9 +146,9 @@ GSPEC > Set.Collect SETSPEC > HOLLightCompat.SETSPEC UNION > Lattices.sup_class.sup :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" - UNIONS > Complete_Lattice.Sup_class.Sup :: "(('a \ bool) \ bool) \ 'a \ bool" + UNIONS > Complete_Lattices.Sup_class.Sup :: "(('a \ bool) \ bool) \ 'a \ bool" INTER > Lattices.inf_class.inf :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" - INTERS > Complete_Lattice.Inf_class.Inf :: "(('a \ bool) \ bool) \ 'a \ bool" + INTERS > Complete_Lattices.Inf_class.Inf :: "(('a \ bool) \ bool) \ 'a \ bool" DIFF > Groups.minus_class.minus :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" SUBSET > Orderings.ord_class.less_eq :: "('a \ bool) \ ('a \ bool) \ bool" PSUBSET > Orderings.ord_class.less :: "('a \ bool) \ ('a \ bool) \ bool" diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Import/HOLLight/hollight.imp --- a/src/HOL/Import/HOLLight/hollight.imp Sat Sep 10 00:44:25 2011 +0200 +++ b/src/HOL/Import/HOLLight/hollight.imp Sat Sep 10 10:29:24 2011 +0200 @@ -266,7 +266,7 @@ "ZBOT" > "HOLLight.hollight.ZBOT" "WF" > "Wellfounded.wfP" "UNIV" > "Orderings.top_class.top" :: "'a => bool" - "UNIONS" > "Complete_Lattice.Sup_class.Sup" :: "(('a => bool) => bool) => 'a => bool" + "UNIONS" > "Complete_Lattices.Sup_class.Sup" :: "(('a => bool) => bool) => 'a => bool" "UNION" > "Lattices.sup_class.sup" :: "('a => bool) => ('a => bool) => 'a => bool" "UNCURRY" > "HOLLight.hollight.UNCURRY" "TL" > "List.tl" @@ -316,7 +316,7 @@ "ITLIST2" > "HOLLightList.fold2" "ITLIST" > "List.foldr" "ISO" > "HOLLight.hollight.ISO" - "INTERS" > "Complete_Lattice.Inf_class.Inf" :: "(('a => bool) => bool) => 'a => bool" + "INTERS" > "Complete_Lattices.Inf_class.Inf" :: "(('a => bool) => bool) => 'a => bool" "INTER" > "Lattices.inf_class.inf" :: "('a => bool) => ('a => bool) => 'a => bool" "INSERT" > "Set.insert" "INR" > "Sum_Type.Inr" @@ -584,15 +584,15 @@ "UNION_COMM" > "Lattices.lattice_class.inf_sup_aci_5" "UNION_ASSOC" > "Lattices.lattice_class.inf_sup_aci_6" "UNION_ACI" > "HOLLight.hollight.UNION_ACI" - "UNIONS_UNION" > "Complete_Lattice.Union_Un_distrib" + "UNIONS_UNION" > "Complete_Lattices.Union_Un_distrib" "UNIONS_SUBSET" > "HOLLight.hollight.UNIONS_SUBSET" "UNIONS_INTERS" > "HOLLight.hollight.UNIONS_INTERS" - "UNIONS_INSERT" > "Complete_Lattice.Union_insert" + "UNIONS_INSERT" > "Complete_Lattices.Union_insert" "UNIONS_IMAGE" > "HOLLight.hollight.UNIONS_IMAGE" "UNIONS_GSPEC" > "HOLLight.hollight.UNIONS_GSPEC" - "UNIONS_2" > "Complete_Lattice.Un_eq_Union" - "UNIONS_1" > "Complete_Lattice.complete_lattice_class.Sup_singleton" - "UNIONS_0" > "Complete_Lattice.Union_empty" + "UNIONS_2" > "Complete_Lattices.Un_eq_Union" + "UNIONS_1" > "Complete_Lattices.complete_lattice_class.Sup_singleton" + "UNIONS_0" > "Complete_Lattices.Union_empty" "UNCURRY_def" > "HOLLight.hollight.UNCURRY_def" "TYDEF_recspace" > "HOLLight.hollight.TYDEF_recspace" "TYDEF_real" > "HOLLight.hollight.TYDEF_real" @@ -779,7 +779,7 @@ "SUB_0" > "HOLLight.hollight.SUB_0" "SUBSET_UNIV" > "Orderings.top_class.top_greatest" "SUBSET_UNION_ABSORPTION" > "Lattices.semilattice_sup_class.le_iff_sup" - "SUBSET_UNIONS" > "Complete_Lattice.Union_mono" + "SUBSET_UNIONS" > "Complete_Lattices.Union_mono" "SUBSET_UNION" > "HOLLight.hollight.SUBSET_UNION" "SUBSET_TRANS" > "Orderings.order_trans_rules_23" "SUBSET_RESTRICT" > "HOLLight.hollight.SUBSET_RESTRICT" @@ -1544,7 +1544,7 @@ "ISO_FUN" > "HOLLight.hollight.ISO_FUN" "IN_UNIV" > "Set.UNIV_I" "IN_UNIONS" > "HOLLight.hollight.IN_UNIONS" - "IN_UNION" > "Complete_Lattice.mem_simps_3" + "IN_UNION" > "Complete_Lattices.mem_simps_3" "IN_SUPPORT" > "HOLLight.hollight.IN_SUPPORT" "IN_SING" > "Set.singleton_iff" "IN_SET_OF_LIST" > "HOLLightList.IN_SET_OF_LIST" @@ -1552,13 +1552,13 @@ "IN_NUMSEG_0" > "HOLLight.hollight.IN_NUMSEG_0" "IN_NUMSEG" > "SetInterval.ord_class.atLeastAtMost_iff" "IN_INTERS" > "HOLLight.hollight.IN_INTERS" - "IN_INTER" > "Complete_Lattice.mem_simps_4" - "IN_INSERT" > "Complete_Lattice.mem_simps_1" + "IN_INTER" > "Complete_Lattices.mem_simps_4" + "IN_INSERT" > "Complete_Lattices.mem_simps_1" "IN_IMAGE" > "HOLLight.hollight.IN_IMAGE" "IN_ELIM_THM" > "HOLLight.hollight.IN_ELIM_THM" "IN_ELIM_PAIR_THM" > "HOLLight.hollight.IN_ELIM_PAIR_THM" "IN_DISJOINT" > "HOLLight.hollight.IN_DISJOINT" - "IN_DIFF" > "Complete_Lattice.mem_simps_6" + "IN_DIFF" > "Complete_Lattices.mem_simps_6" "IN_DELETE_EQ" > "HOLLight.hollight.IN_DELETE_EQ" "IN_DELETE" > "HOLLight.hollight.IN_DELETE" "IN_CROSS" > "HOLLight.hollight.IN_CROSS" @@ -1594,12 +1594,12 @@ "INTER_ASSOC" > "Lattices.lattice_class.inf_sup_aci_2" "INTER_ACI" > "HOLLight.hollight.INTER_ACI" "INTERS_UNIONS" > "HOLLight.hollight.INTERS_UNIONS" - "INTERS_INSERT" > "Complete_Lattice.Inter_insert" + "INTERS_INSERT" > "Complete_Lattices.Inter_insert" "INTERS_IMAGE" > "HOLLight.hollight.INTERS_IMAGE" "INTERS_GSPEC" > "HOLLight.hollight.INTERS_GSPEC" - "INTERS_2" > "Complete_Lattice.Int_eq_Inter" - "INTERS_1" > "Complete_Lattice.complete_lattice_class.Inf_singleton" - "INTERS_0" > "Complete_Lattice.Inter_empty" + "INTERS_2" > "Complete_Lattices.Int_eq_Inter" + "INTERS_1" > "Complete_Lattices.complete_lattice_class.Inf_singleton" + "INTERS_0" > "Complete_Lattices.Inter_empty" "INSERT_UNIV" > "HOLLight.hollight.INSERT_UNIV" "INSERT_UNION_EQ" > "Set.Un_insert_left" "INSERT_UNION" > "HOLLight.hollight.INSERT_UNION" diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Inductive.thy --- a/src/HOL/Inductive.thy Sat Sep 10 00:44:25 2011 +0200 +++ b/src/HOL/Inductive.thy Sat Sep 10 10:29:24 2011 +0200 @@ -5,7 +5,7 @@ header {* Knaster-Tarski Fixpoint Theorem and inductive definitions *} theory Inductive -imports Complete_Lattice +imports Complete_Lattices uses ("Tools/inductive.ML") "Tools/dseq.ML" diff -r 237ba63d6041 -r 56101fa00193 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Sat Sep 10 00:44:25 2011 +0200 +++ b/src/HOL/IsaMakefile Sat Sep 10 10:29:24 2011 +0200 @@ -171,7 +171,7 @@ $(SRC)/Tools/Metis/metis.ML \ $(SRC)/Tools/rat.ML \ ATP.thy \ - Complete_Lattice.thy \ + Complete_Lattices.thy \ Complete_Partial_Order.thy \ Datatype.thy \ Extraction.thy \ diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Library/Executable_Set.thy --- a/src/HOL/Library/Executable_Set.thy Sat Sep 10 00:44:25 2011 +0200 +++ b/src/HOL/Library/Executable_Set.thy Sat Sep 10 10:29:24 2011 +0200 @@ -199,17 +199,17 @@ by simp definition Inf :: "'a::complete_lattice set \ 'a" where - [simp]: "Inf = Complete_Lattice.Inf" + [simp]: "Inf = Complete_Lattices.Inf" lemma [code_unfold]: - "Complete_Lattice.Inf = Inf" + "Complete_Lattices.Inf = Inf" by simp definition Sup :: "'a::complete_lattice set \ 'a" where - [simp]: "Sup = Complete_Lattice.Sup" + [simp]: "Sup = Complete_Lattices.Sup" lemma [code_unfold]: - "Complete_Lattice.Sup = Sup" + "Complete_Lattices.Sup = Sup" by simp definition Inter :: "'a set set \ 'a set" where diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Library/Lattice_Syntax.thy --- a/src/HOL/Library/Lattice_Syntax.thy Sat Sep 10 00:44:25 2011 +0200 +++ b/src/HOL/Library/Lattice_Syntax.thy Sat Sep 10 10:29:24 2011 +0200 @@ -3,7 +3,7 @@ header {* Pretty syntax for lattice operations *} theory Lattice_Syntax -imports Complete_Lattice +imports Complete_Lattices begin notation diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Main.thy --- a/src/HOL/Main.thy Sat Sep 10 00:44:25 2011 +0200 +++ b/src/HOL/Main.thy Sat Sep 10 10:29:24 2011 +0200 @@ -21,7 +21,7 @@ "Sup A \ (\x\A. x)" by auto -declare Complete_Lattice.Inf_bool_def [simp del] -declare Complete_Lattice.Sup_bool_def [simp del] +declare Complete_Lattices.Inf_bool_def [simp del] +declare Complete_Lattices.Sup_bool_def [simp del] end diff -r 237ba63d6041 -r 56101fa00193 src/HOL/Quotient_Examples/Cset.thy --- a/src/HOL/Quotient_Examples/Cset.thy Sat Sep 10 00:44:25 2011 +0200 +++ b/src/HOL/Quotient_Examples/Cset.thy Sat Sep 10 10:29:24 2011 +0200 @@ -89,7 +89,7 @@ quotient_definition Sup where "Sup :: ('a :: Sup) Cset.set \ 'a" is "Sup_class.Sup :: ('a :: Sup) set \ 'a" quotient_definition UNION where "UNION :: 'a Cset.set \ ('a \ 'b Cset.set) \ 'b Cset.set" -is "Complete_Lattice.UNION :: 'a set \ ('a \ 'b set) \ 'b set" +is "Complete_Lattices.UNION :: 'a set \ ('a \ 'b set) \ 'b set" hide_const (open) is_empty insert remove map filter forall exists card set subset psubset inter union empty UNIV uminus minus Inf Sup UNION