# HG changeset patch # User haftmann # Date 1310927013 -7200 # Node ID 9684251c7ec1d5511342d9782b7e77c2c4a68110 # Parent 771014555553d94e5d47af05811f923523154f41 more lemmas about Sup diff -r 771014555553 -r 9684251c7ec1 src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Sun Jul 17 19:55:17 2011 +0200 +++ b/src/HOL/Complete_Lattice.thy Sun Jul 17 20:23:33 2011 +0200 @@ -33,7 +33,7 @@ begin lemma dual_complete_lattice: - "class.complete_lattice Sup Inf (op \) (op >) (op \) (op \) \ \" + "class.complete_lattice Sup Inf (op \) (op >) sup 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)+) @@ -109,10 +109,79 @@ qed lemma Sup_upper2: "u \ A \ v \ u \ v \ \A" - using Sup_upper[of u A] by auto + using Sup_upper [of u A] by auto lemma Inf_lower2: "u \ A \ u \ v \ \A \ v" - using Inf_lower[of u A] by auto + using Inf_lower [of u A] by auto + +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 Inf_inter_less_eq: "\A \ \B \ \(A \ B)" + by (auto intro: Inf_greatest Inf_lower) + +lemma Sup_inter_greater_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 Sup_union_distrib: "\(A \ B) = \A \ \B" + by (rule antisym) (auto intro: Sup_least Sup_upper le_supI1 le_supI2) + +lemma Inf_top_conv [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 + 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 Sup_bot_conv [no_atp]: + "\A = \ \ (\x\A. x = \)" (is ?P) + "\ = \A \ (\x\A. x = \)" (is ?Q) +proof - + interpret dual: complete_lattice Sup Inf "op \" "op >" sup inf \ \ + by (fact dual_complete_lattice) + from dual.Inf_top_conv show ?P and ?Q by simp_all +qed + +lemma Inf_anti_mono: "B \ A \ \A \ \B" + by (auto intro: Inf_greatest Inf_lower) + +lemma Sup_anti_mono: "A \ B \ \A \ \B" + by (auto intro: Sup_least Sup_upper) definition INFI :: "'b set \ ('b \ 'a) \ 'a" where "INFI A f = \ (f ` A)" @@ -187,6 +256,48 @@ lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" by (iprover intro: INF_leI le_INFI order_trans antisym) +lemma (in complete_lattice) INFI_empty: + "(\x\{}. B x) = \" + by (simp add: INFI_def) + +lemma (in complete_lattice) INFI_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 (in complete_lattice) 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 le_INFI INF_leI) + +lemma (in complete_lattice) INF_constant: + "(\y\A. c) = (if A = {} then \ else c)" + by (simp add: INF_empty) + +lemma (in complete_lattice) INF_eq: + "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (simp add: INFI_def image_def) + +lemma (in complete_lattice) INF_top_conv: + "\ = (\x\A. B x) \ (\x\A. B x = \)" + "(\x\A. B x) = \ \ (\x\A. B x = \)" + by (auto simp add: INFI_def Inf_top_conv) + +lemma (in complete_lattice) INFI_UNIV_range: + "(\x\UNIV. f x) = \range f" + by (simp add: INFI_def) + +lemma (in complete_lattice) INF_bool_eq: + "(\b. A b) = A True \ A False" + by (simp add: UNIV_bool INF_empty INF_insert inf_commute) + +lemma (in complete_lattice) INF_anti_mono: + "B \ A \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\B. g x)" + -- {* The last inclusion is POSITIVE! *} + by (blast intro: INF_mono dest: subsetD) + lemma SUP_cong: "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" by (simp add: SUPR_def image_def) @@ -351,16 +462,6 @@ lemma Inter_lower: "B \ A \ \A \ B" by (fact Inf_lower) -lemma (in complete_lattice) 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 Inter_subset: "(\X. X \ A \ X \ B) \ A \ {} \ \A \ B" by (fact Inf_less_eq) @@ -380,48 +481,17 @@ lemma Inter_insert [simp]: "\(insert a B) = a \ \B" by (fact Inf_insert) -lemma (in complete_lattice) Inf_inter_less: "\A \ \B \ \(A \ B)" - by (auto intro: Inf_greatest Inf_lower) - lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" - by (fact Inf_inter_less) - -lemma (in complete_lattice) Inf_union_distrib: "\(A \ B) = \A \ \B" - by (rule antisym) (auto intro: Inf_greatest Inf_lower le_infI1 le_infI2) + by (fact Inf_inter_less_eq) lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" by (fact Inf_union_distrib) -lemma (in complete_lattice) Inf_top_conv [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 - 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 Inter_UNIV_conv [simp,no_atp]: +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 (in complete_lattice) Inf_anti_mono: "B \ A \ \A \ \B" - by (auto intro: Inf_greatest Inf_lower) - lemma Inter_anti_mono: "B \ A \ \A \ \B" by (fact Inf_anti_mono) @@ -498,21 +568,9 @@ lemma INT_greatest: "(\x. x \ A \ C \ B x) \ C \ (\x\A. B x)" by (fact le_INFI) -lemma (in complete_lattice) INFI_empty: - "(\x\{}. B x) = \" - by (simp add: INFI_def) - lemma INT_empty [simp]: "(\x\{}. B x) = UNIV" by (fact INFI_empty) -lemma (in complete_lattice) INFI_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 INT_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" by (fact INFI_absorb) @@ -522,10 +580,6 @@ lemma INT_insert [simp]: "(\x \ insert a A. B x) = B a \ INTER A B" by (fact INF_insert) -lemma (in complete_lattice) 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 le_INFI INF_leI) - lemma INT_Un: "(\i \ A \ B. M i) = (\i \ A. M i) \ (\i\B. M i)" by (fact INF_union) @@ -533,47 +587,21 @@ "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" by blast -lemma (in complete_lattice) INF_constant: - "(\y\A. c) = (if A = {} then \ else c)" - by (simp add: INF_empty) - lemma INT_constant [simp]: "(\y\A. c) = (if A = {} then UNIV else c)" by (fact INF_constant) -lemma (in complete_lattice) INF_eq: - "(\x\A. B x) = \({Y. \x\A. Y = B x})" - by (simp add: INFI_def image_def) - lemma INT_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" -- {* Look: it has an \emph{existential} quantifier *} by (fact INF_eq) -lemma (in complete_lattice) INF_top_conv: - "\ = (\x\A. B x) \ (\x\A. B x = \)" - "(\x\A. B x) = \ \ (\x\A. B x = \)" - by (auto simp add: INFI_def Inf_top_conv) - 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 (in complete_lattice) INFI_UNIV_range: - "(\x\UNIV. f x) = \range f" - by (simp add: INFI_def) - -lemma (in complete_lattice) INF_bool_eq: - "(\b. A b) = A True \ A False" - by (simp add: UNIV_bool INF_empty INF_insert inf_commute) - lemma INT_bool_eq: "(\b. A b) = A True \ A False" by (fact INF_bool_eq) -lemma (in complete_lattice) INF_anti_mono: - "B \ A \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\B. g x)" - -- {* The last inclusion is POSITIVE! *} - by (blast intro: INF_mono dest: subsetD) - lemma INT_anti_mono: "B \ A \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\B. g x)" -- {* The last inclusion is POSITIVE! *}