# HG changeset patch # User haftmann # Date 1311052454 -7200 # Node ID 8064210028b73e7da98f4ab9df2b8635c782efdf # Parent b28745c3ddcec0b080313033d046c178c1a73b09# Parent 3ab6c30d256dac49c5fe6769b55cbefc98d9ab2c merged diff -r b28745c3ddce -r 8064210028b7 NEWS --- a/NEWS Tue Jul 19 00:16:18 2011 +0200 +++ b/NEWS Tue Jul 19 07:14:14 2011 +0200 @@ -71,6 +71,7 @@ le_SUPI ~> le_SUP_I le_SUPI2 ~> le_SUP_I2 le_INFI ~> le_INF_I + INF_subset ~> INF_superset_mono INFI_bool_eq ~> INF_bool_eq SUPR_bool_eq ~> SUP_bool_eq INFI_apply ~> INF_apply diff -r b28745c3ddce -r 8064210028b7 src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Tue Jul 19 00:16:18 2011 +0200 +++ b/src/HOL/Complete_Lattice.thy Tue Jul 19 07:14:14 2011 +0200 @@ -88,6 +88,12 @@ lemma Sup_le_iff: "\A \ b \ (\a\A. a \ b)" by (auto intro: Sup_least dest: Sup_upper) +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_mono: assumes "\b. b \ B \ \a\A. a \ b" shows "\A \ \B" @@ -134,10 +140,10 @@ ultimately show ?thesis by (rule Sup_upper2) qed -lemma Inf_inter_less_eq: "\A \ \B \ \(A \ B)" +lemma less_eq_Inf_inter: "\A \ \B \ \(A \ B)" by (auto intro: Inf_greatest Inf_lower) -lemma Sup_inter_greater_eq: "\(A \ B) \ \A \ \B " +lemma Sup_inter_less_eq: "\(A \ B) \ \A \ \B " by (auto intro: Sup_least Sup_upper) lemma Inf_union_distrib: "\(A \ B) = \A \ \B" @@ -177,12 +183,6 @@ 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 INF_def: "INFI A f = \ (f ` A)" @@ -269,6 +269,12 @@ lemma SUP_const [simp]: "A \ {} \ (\i\A. f) = f" by (auto intro: antisym SUP_leI le_SUP_I) +lemma INF_top: "(\x\A. \) = \" + by (cases "A = {}") (simp_all add: INF_empty) + +lemma SUP_bot: "(\x\A. \) = \" + by (cases "A = {}") (simp_all add: SUP_empty) + 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) @@ -285,13 +291,13 @@ "(\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_subset: - "A \ B \ INFI B f \ INFI A f" - by (intro INF_mono) auto +lemma INF_superset_mono: + "B \ A \ INFI A f \ INFI B f" + by (rule INF_mono) auto -lemma SUP_subset: +lemma SUP_subset_mono: "A \ B \ SUPR A f \ SUPR B f" - by (intro SUP_mono) auto + by (rule SUP_mono) auto lemma INF_commute: "(\i\A. \j\B. f i j) = (\j\B. \i\A. f i j)" by (iprover intro: INF_leI le_INF_I order_trans antisym) @@ -365,13 +371,13 @@ "(\b. A b) = A True \ A False" by (simp add: UNIV_bool SUP_empty SUP_insert sup_commute) -lemma INF_anti_mono: - "B \ A \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\B. g x)" +lemma INF_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) + by (rule INF_mono) auto -lemma SUP_anti_mono: - "B \ A \ (\x. x \ A \ g x \ f x) \ (\x\B. g x) \ (\x\B. f x)" +lemma SUP_mono': + "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" -- {* The last inclusion is POSITIVE! *} by (blast intro: SUP_mono dest: subsetD) @@ -397,14 +403,6 @@ shows "a \ (\i\A. f i) \ (\x\A. a \ f x)" unfolding SUP_def less_Sup_iff by auto --- "FIXME move" - -lemma image_ident [simp]: "(%x. x) ` Y = Y" - by blast - -lemma vimage_ident [simp]: "(%x. x) -` Y = Y" - by blast - class complete_boolean_algebra = boolean_algebra + complete_lattice begin @@ -562,7 +560,7 @@ by (fact Inf_insert) lemma Inter_Un_subset: "\A \ \B \ \(A \ B)" - by (fact Inf_inter_less_eq) + by (fact less_eq_Inf_inter) lemma Inter_Un_distrib: "\(A \ B) = \A \ \B" by (fact Inf_union_distrib) @@ -573,7 +571,7 @@ by (fact Inf_top_conv)+ lemma Inter_anti_mono: "B \ A \ \A \ \B" - by (fact Inf_anti_mono) + by (fact Inf_superset_mono) subsection {* Intersections of families *} @@ -688,9 +686,9 @@ by (fact INF_UNIV_bool_expand) lemma INT_anti_mono: - "B \ A \ (\x. x \ A \ f x \ g x) \ (\x\B. f x) \ (\x\B. g x)" + "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_anti_mono) + by (fact INF_mono') lemma Pow_INT_eq: "Pow (\x\A. B x) = (\x\A. Pow (B x))" by blast @@ -732,36 +730,36 @@ by auto lemma Union_upper: "B \ A \ B \ \A" - by (iprover intro: subsetI UnionI) + by (fact Sup_upper) lemma Union_least: "(\X. X \ A \ X \ C) \ \A \ C" - by (iprover intro: subsetI elim: UnionE dest: subsetD) + by (fact Sup_least) lemma Un_eq_Union: "A \ B = \{A, B}" by blast lemma Union_empty [simp]: "\{} = {}" - by blast + by (fact Sup_empty) lemma Union_UNIV [simp]: "\UNIV = UNIV" - by blast + by (fact Sup_UNIV) lemma Union_insert [simp]: "\insert a B = a \ \B" - by blast + by (fact Sup_insert) lemma Union_Un_distrib [simp]: "\(A \ B) = \A \ \B" - by blast + by (fact Sup_union_distrib) lemma Union_Int_subset: "\(A \ B) \ \A \ \B" - by blast + by (fact Sup_inter_less_eq) lemma Union_empty_conv [simp,no_atp]: "(\A = {}) \ (\x\A. x = {})" - by blast + by (fact Sup_bot_conv) lemma empty_Union_conv [simp,no_atp]: "({} = \A) \ (\x\A. x = {})" - by blast + by (fact Sup_bot_conv) -lemma Union_disjoint: "(\C \ A = {}) \ (\B\C. B \ A = {})" +lemma Union_disjoint: "(\C \ A = {}) \ (\B\C. B \ A = {})" -- "FIXME generalize" by blast lemma subset_Pow_Union: "A \ Pow (\A)" @@ -771,7 +769,7 @@ by blast lemma Union_mono: "A \ B \ \A \ \B" - by blast + by (fact Sup_subset_mono) subsection {* Unions of families *} @@ -843,12 +841,12 @@ by (unfold UNION_def) blast lemma UN_cong [cong]: - "A = B \ (\x. x \ B \ C x = D x) \ (\x\A. C x) = (\x\B. D x)" - by (simp add: UNION_def) + "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 (simp add: UNION_def simp_implies_def) + "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 @@ -857,7 +855,7 @@ by (fact le_SUP_I) lemma UN_least: "(\x. x \ A \ B x \ C) \ (\x\A. B x) \ C" - by (iprover intro: subsetI elim: UN_E dest: subsetD) + by (fact SUP_leI) lemma Collect_bex_eq [no_atp]: "{x. \y\A. P x y} = (\y\A. {x. P x y})" by blast @@ -865,58 +863,58 @@ lemma UN_insert_distrib: "u \ A \ (\x\A. insert a (B x)) = insert a (\x\A. B x)" by blast -lemma UN_empty [simp,no_atp]: "(\x\{}. B x) = {}" - by blast +lemma UN_empty [simp, no_atp]: "(\x\{}. B x) = {}" + by (fact SUP_empty) lemma UN_empty2 [simp]: "(\x\A. {}) = {}" - by blast + by (fact SUP_bot) lemma UN_singleton [simp]: "(\x\A. {x}) = A" by blast lemma UN_absorb: "k \ I \ A k \ (\i\I. A i) = (\i\I. A i)" - by auto + by (fact SUP_absorb) lemma UN_insert [simp]: "(\x\insert a A. B x) = B a \ UNION A B" - by blast + by (fact SUP_insert) lemma UN_Un[simp]: "(\i \ A \ B. M i) = (\i\A. M i) \ (\i\B. M i)" - by blast + by (fact SUP_union) -lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" +lemma UN_UN_flatten: "(\x \ (\y\A. B y). C x) = (\y\A. \x\B y. C x)" -- "FIXME generalize" by blast lemma UN_subset_iff: "((\i\I. A i) \ B) = (\i\I. A i \ B)" by (fact SUP_le_iff) -lemma image_Union: "f ` \S = (\x\S. f ` x)" - by blast - lemma UN_constant [simp]: "(\y\A. c) = (if A = {} then {} else c)" - by auto + by (fact SUP_constant) lemma UN_eq: "(\x\A. B x) = \({Y. \x\A. Y = B x})" + by (fact SUP_eq) + +lemma image_Union: "f ` \S = (\x\S. f ` x)" -- "FIXME generalize" by blast lemma UNION_empty_conv[simp]: "{} = (\x\A. B x) \ (\x\A. B x = {})" "(\x\A. B x) = {} \ (\x\A. B x = {})" -by blast+ + 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)" +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)" +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 (auto intro: bool_contrapos) + by (fact SUP_UNIV_bool_expand) lemma UN_Pow_subset: "(\x\A. Pow (B x)) \ Pow (\x\A. B x)" by blast @@ -924,7 +922,7 @@ lemma UN_mono: "A \ B \ (\x. x \ A \ f x \ g x) \ (\x\A. f x) \ (\x\B. g x)" - by (blast dest: subsetD) + by (fact SUP_mono') lemma vimage_Union: "f -` (\A) = (\X\A. f -` X)" by blast @@ -1085,6 +1083,7 @@ lemmas (in complete_lattice) le_SUPI = le_SUP_I lemmas (in complete_lattice) le_SUPI2 = le_SUP_I2 lemmas (in complete_lattice) le_INFI = le_INF_I +lemmas (in complete_lattice) INF_subset = INF_superset_mono lemmas INFI_apply = INF_apply lemmas SUPR_apply = SUP_apply diff -r b28745c3ddce -r 8064210028b7 src/HOL/Set.thy --- a/src/HOL/Set.thy Tue Jul 19 00:16:18 2011 +0200 +++ b/src/HOL/Set.thy Tue Jul 19 07:14:14 2011 +0200 @@ -880,6 +880,9 @@ \medskip Range of a function -- just a translation for image! *} +lemma image_ident [simp]: "(%x. x) ` Y = Y" + by blast + lemma range_eqI: "b = f x ==> b \ range f" by simp @@ -1163,6 +1166,12 @@ lemma image_cong: "M = N ==> (!!x. x \ N ==> f x = g x) ==> f`M = g`N" by (simp add: image_def) +lemma image_Int_subset: "f`(A Int B) <= f`A Int f`B" +by blast + +lemma image_diff_subset: "f`A - f`B <= f`(A - B)" +by blast + text {* \medskip @{text range}. *} @@ -1673,11 +1682,8 @@ "(\ w. w \ S \ f w = g w) \ f -` y \ S = g -` y \ S" by auto -lemma image_Int_subset: "f`(A Int B) <= f`A Int f`B" -by blast - -lemma image_diff_subset: "f`A - f`B <= f`(A - B)" -by blast +lemma vimage_ident [simp]: "(%x. x) -` Y = Y" + by blast subsubsection {* Getting the Contents of a Singleton Set *}