# HG changeset patch # User Andreas Lochbihler # Date 1311667399 -7200 # Node ID 24a3ddc79a5c76fe5e5f69abc5ea8c57a5472aac # Parent 892030194015e17d0f6790c44435293318979dab# Parent 3d204d261903b2591e146a03fc36cb35a84b967c merged diff -r 892030194015 -r 24a3ddc79a5c NEWS --- a/NEWS Mon Jul 25 16:55:48 2011 +0200 +++ b/NEWS Tue Jul 26 10:03:19 2011 +0200 @@ -79,6 +79,10 @@ SUPR_apply ~> SUP_apply INCOMPATIBILITY. +* Theorem collections ball_simps and bex_simps do not contain theorems referring +to UNION any longer; these have been moved to collection UN_ball_bex_simps. +INCOMPATIBILITY. + * Archimedian_Field.thy: floor now is defined as parameter of a separate type class floor_ceiling. diff -r 892030194015 -r 24a3ddc79a5c src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Mon Jul 25 16:55:48 2011 +0200 +++ b/src/HOL/Complete_Lattice.thy Tue Jul 26 10:03:19 2011 +0200 @@ -476,6 +476,14 @@ from dual.Sup_eq_top_iff show ?thesis . qed +lemma INF_eq_bot_iff: + "(\i\A. f i) = \ \ (\x>\. \i\A. f i < x)" + unfolding INF_def Inf_eq_bot_iff by auto + +lemma SUP_eq_top_iff: + "(\i\A. f i) = \ \ (\x<\. \i\A. x < f i)" + unfolding SUP_def Sup_eq_top_iff by auto + end @@ -929,7 +937,7 @@ 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)" -- "FIXME generalize" +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)" @@ -1067,44 +1075,13 @@ "\A B f. (\x\f`A. B x) = (\a\A. B (f a))" by auto -lemma ball_simps [simp,no_atp]: - "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" - "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" - "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" - "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" - "\P. (\x\{}. P x) \ True" - "\P. (\x\UNIV. P x) \ (\x. P x)" - "\a B P. (\x\insert a B. P x) \ (P a \ (\x\B. P x))" +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)" - "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" - "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" - "\A P. (\ (\x\A. P x)) \ (\x\A. \ P x)" - by auto - -lemma bex_simps [simp,no_atp]: - "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" - "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" - "\P. (\x\{}. P x) \ False" - "\P. (\x\UNIV. P x) \ (\x. P x)" - "\a B P. (\x\insert a B. P x) \ (P a | (\x\B. 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)" - "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" - "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" - "\A P. (\(\x\A. P x)) \ (\x\A. \ P x)" by auto -lemma (in complete_linorder) INF_eq_bot_iff: - fixes f :: "'b \ 'a" - shows "(\i\A. f i) = \ \ (\x>\. \i\A. f i < x)" - unfolding INF_def Inf_eq_bot_iff by auto - -lemma (in complete_linorder) SUP_eq_top_iff: - fixes f :: "'b \ 'a" - shows "(\i\A. f i) = \ \ (\x<\. \i\A. x < f i)" - unfolding SUP_def Sup_eq_top_iff by auto - text {* \medskip Maxiscoping: pulling out big Unions and Intersections. *} diff -r 892030194015 -r 24a3ddc79a5c src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Mon Jul 25 16:55:48 2011 +0200 +++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy Tue Jul 26 10:03:19 2011 +0200 @@ -2539,7 +2539,7 @@ fixes s :: "'a::real_normed_vector set" assumes "open s" shows "open(convex hull s)" - unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(10) + unfolding open_contains_cball convex_hull_explicit unfolding mem_Collect_eq ball_simps(8) proof(rule, rule) fix a assume "\sa u. finite sa \ sa \ s \ (\x\sa. 0 \ u x) \ setsum u sa = 1 \ (\v\sa. u v *\<^sub>R v) = a" then obtain t u where obt:"finite t" "t\s" "\x\t. 0 \ u x" "setsum u t = 1" "(\v\t. u v *\<^sub>R v) = a" by auto diff -r 892030194015 -r 24a3ddc79a5c src/HOL/Multivariate_Analysis/Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Mon Jul 25 16:55:48 2011 +0200 +++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy Tue Jul 26 10:03:19 2011 +0200 @@ -2505,7 +2505,7 @@ by auto from inj_on_image_set_diff[OF "2.prems"(3), of "insert a b " "{a}", symmetric] have tha: "f ` insert a b - f ` {a} = f ` (insert a b - {a})" by blast - from "2.prems"(2)[unfolded dependent_def bex_simps(10), rule_format, of "f a"] + from "2.prems"(2) [unfolded dependent_def bex_simps(8), rule_format, of "f a"] have "f a \ span (f ` b)" using tha using "2.hyps"(2) "2.prems"(3) by auto diff -r 892030194015 -r 24a3ddc79a5c src/HOL/Set.thy --- a/src/HOL/Set.thy Mon Jul 25 16:55:48 2011 +0200 +++ b/src/HOL/Set.thy Tue Jul 26 10:03:19 2011 +0200 @@ -1540,6 +1540,30 @@ lemma distinct_lemma: "f x \ f y ==> x \ y" by iprover +lemma ball_simps [simp, no_atp]: + "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" + "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" + "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" + "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" + "\P. (\x\{}. P x) \ True" + "\P. (\x\UNIV. P x) \ (\x. P x)" + "\a B P. (\x\insert a B. P x) \ (P a \ (\x\B. P x))" + "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" + "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" + "\A P. (\ (\x\A. P x)) \ (\x\A. \ P x)" + by auto + +lemma bex_simps [simp, no_atp]: + "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" + "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" + "\P. (\x\{}. P x) \ False" + "\P. (\x\UNIV. P x) \ (\x. P x)" + "\a B P. (\x\insert a B. P x) \ (P a | (\x\B. P x))" + "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" + "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" + "\A P. (\(\x\A. P x)) \ (\x\A. \ P x)" + by auto + subsubsection {* Monotonicity of various operations *}