diff -r 2bc2a8599369 -r a03a63b81f44 src/HOL/Analysis/Convex.thy --- a/src/HOL/Analysis/Convex.thy Mon Jan 14 18:33:53 2019 +0000 +++ b/src/HOL/Analysis/Convex.thy Mon Jan 14 18:35:03 2019 +0000 @@ -613,15 +613,18 @@ qed lemma convex_translation: - assumes "convex S" - shows "convex ((\x. a + x) ` S)" + "convex ((+) a ` S)" if "convex S" proof - - have "(\ x\ {a}. \y \ S. {x + y}) = (\x. a + x) ` S" + have "(\ x\ {a}. \y \ S. {x + y}) = (+) a ` S" by auto then show ?thesis - using convex_sums[OF convex_singleton[of a] assms] by auto + using convex_sums [OF convex_singleton [of a] that] by auto qed +lemma convex_translation_subtract: + "convex ((\b. b - a) ` S)" if "convex S" + using convex_translation [of S "- a"] that by (simp cong: image_cong_simp) + lemma convex_affinity: assumes "convex S" shows "convex ((\x. a + c *\<^sub>R x) ` S)" @@ -1106,9 +1109,14 @@ finally show ?thesis . qed (insert assms(2), simp_all) -lemma convex_translation_eq [simp]: "convex ((\x. a + x) ` s) \ convex s" +lemma convex_translation_eq [simp]: + "convex ((+) a ` s) \ convex s" by (metis convex_translation translation_galois) +lemma convex_translation_subtract_eq [simp]: + "convex ((\b. b - a) ` s) \ convex s" + using convex_translation_eq [of "- a"] by (simp cong: image_cong_simp) + lemma convex_linear_image_eq [simp]: fixes f :: "'a::real_vector \ 'b::real_vector" shows "\linear f; inj f\ \ convex (f ` s) \ convex s" @@ -1614,13 +1622,13 @@ qed lemma affine_translation: - fixes a :: "'a::real_vector" - shows "affine S \ affine ((\x. a + x) ` S)" -proof - - have "affine S \ affine ((\x. a + x) ` S)" - using affine_translation_aux[of "-a" "((\x. a + x) ` S)"] - using translation_assoc[of "-a" a S] by auto - then show ?thesis using affine_translation_aux by auto + "affine S \ affine ((+) a ` S)" for a :: "'a::real_vector" +proof + show "affine ((+) a ` S)" if "affine S" + using that translation_assoc [of "- a" a S] + by (auto intro: affine_translation_aux [of "- a" "((+) a ` S)"]) + show "affine S" if "affine ((+) a ` S)" + using that by (rule affine_translation_aux) qed lemma parallel_is_affine: @@ -1700,6 +1708,10 @@ ultimately show ?thesis using subspace_affine by auto qed +lemma affine_diffs_subspace_subtract: + "subspace ((\x. x - a) ` S)" if "affine S" "a \ S" + using that affine_diffs_subspace [of _ a] by simp + lemma parallel_subspace_explicit: assumes "affine S" and "a \ S" @@ -3235,8 +3247,7 @@ qed lemma aff_dim_translation_eq: - fixes a :: "'n::euclidean_space" - shows "aff_dim ((\x. a + x) ` S) = aff_dim S" + "aff_dim ((+) a ` S) = aff_dim S" for a :: "'n::euclidean_space" proof - have "affine_parallel (affine hull S) (affine hull ((\x. a + x) ` S))" unfolding affine_parallel_def @@ -3248,6 +3259,10 @@ using aff_dim_parallel_eq[of S "(\x. a + x) ` S"] by auto qed +lemma aff_dim_translation_eq_subtract: + "aff_dim ((\x. x - a) ` S) = aff_dim S" for a :: "'n::euclidean_space" + using aff_dim_translation_eq [of "- a"] by (simp cong: image_cong_simp) + lemma aff_dim_affine: fixes S L :: "'n::euclidean_space set" assumes "S \ {}" @@ -3306,17 +3321,21 @@ qed lemma aff_dim_eq_dim: - fixes S :: "'n::euclidean_space set" - assumes "a \ affine hull S" - shows "aff_dim S = int (dim ((\x. -a+x) ` S))" + "aff_dim S = int (dim ((+) (- a) ` S))" if "a \ affine hull S" + for S :: "'n::euclidean_space set" proof - - have "0 \ affine hull ((\x. -a+x) ` S)" + have "0 \ affine hull (+) (- a) ` S" unfolding affine_hull_translation - using assms by (simp add: ab_group_add_class.ab_left_minus image_iff) + using that by (simp add: ac_simps) with aff_dim_zero show ?thesis by (metis aff_dim_translation_eq) qed +lemma aff_dim_eq_dim_subtract: + "aff_dim S = int (dim ((\x. x - a) ` S))" if "a \ affine hull S" + for S :: "'n::euclidean_space set" + using aff_dim_eq_dim [of a] that by (simp cong: image_cong_simp) + lemma aff_dim_UNIV [simp]: "aff_dim (UNIV :: 'n::euclidean_space set) = int(DIM('n))" using aff_dim_subspace[of "(UNIV :: 'n::euclidean_space set)"] dim_UNIV[where 'a="'n::euclidean_space"]