diff --git a/src/analysis/convex/strict.lean b/src/analysis/convex/strict.lean index ed7d4e8a92fa1..1d62c3e56e0ad 100644 --- a/src/analysis/convex/strict.lean +++ b/src/analysis/convex/strict.lean @@ -21,7 +21,7 @@ Define strictly convex spaces. open set open_locale convex pointwise -variables {๐•œ E F ฮฒ : Type*} +variables {๐•œ ๐• E F ฮฒ : Type*} open function set open_locale convex @@ -107,13 +107,14 @@ lemma strict_convex_singleton (c : E) : strict_convex ๐•œ ({c} : set E) := pair lemma set.subsingleton.strict_convex (hs : s.subsingleton) : strict_convex ๐•œ s := hs.pairwise _ -lemma strict_convex.linear_image (hs : strict_convex ๐•œ s) (f : E โ†’โ‚—[๐•œ] F) (hf : is_open_map f) : +lemma strict_convex.linear_image [semiring ๐•] [module ๐• E] [module ๐• F] + [linear_map.compatible_smul E F ๐•œ ๐•] (hs : strict_convex ๐•œ s) (f : E โ†’โ‚—[๐•] F) + (hf : is_open_map f) : strict_convex ๐•œ (f '' s) := begin rintro _ โŸจx, hx, rflโŸฉ _ โŸจy, hy, rflโŸฉ hxy a b ha hb hab, - exact hf.image_interior_subset _ - โŸจa โ€ข x + b โ€ข y, hs hx hy (ne_of_apply_ne _ hxy) ha hb hab, - by rw [f.map_add, f.map_smul, f.map_smul]โŸฉ, + refine hf.image_interior_subset _ โŸจa โ€ข x + b โ€ข y, hs hx hy (ne_of_apply_ne _ hxy) ha hb hab, _โŸฉ, + rw [map_add, f.map_smul_of_tower a, f.map_smul_of_tower b] end lemma strict_convex.is_linear_image (hs : strict_convex ๐•œ s) {f : E โ†’ F} (h : is_linear_map ๐•œ f) @@ -183,7 +184,7 @@ end add_comm_monoid section add_cancel_comm_monoid variables [add_cancel_comm_monoid E] [has_continuous_add E] [module ๐•œ E] {s : set E} -/-- The translation of a strict_convex set is also strict_convex. -/ +/-- The translation of a strictly convex set is also strictly convex. -/ lemma strict_convex.preimage_add_right (hs : strict_convex ๐•œ s) (z : E) : strict_convex ๐•œ ((ฮป x, z + x) โปยน' s) := begin @@ -193,7 +194,7 @@ begin rwa [smul_add, smul_add, add_add_add_comm, โ†add_smul, hab, one_smul] at h, end -/-- The translation of a strict_convex set is also strict_convex. -/ +/-- The translation of a strictly convex set is also strictly convex. -/ lemma strict_convex.preimage_add_left (hs : strict_convex ๐•œ s) (z : E) : strict_convex ๐•œ ((ฮป x, x + z) โปยน' s) := by simpa only [add_comm] using hs.preimage_add_right z @@ -201,44 +202,55 @@ by simpa only [add_comm] using hs.preimage_add_right z end add_cancel_comm_monoid section add_comm_group -variables [add_comm_group E] [module ๐•œ E] {s t : set E} - -lemma strict_convex.add_left [has_continuous_add E] (hs : strict_convex ๐•œ s) (z : E) : - strict_convex ๐•œ ((ฮป x, z + x) '' s) := -begin - rintro _ โŸจx, hx, rflโŸฉ _ โŸจy, hy, rflโŸฉ hxy a b ha hb hab, - refine (is_open_map_add_left _).image_interior_subset _ _, - refine โŸจa โ€ข x + b โ€ข y, hs hx hy (ne_of_apply_ne _ hxy) ha hb hab, _โŸฉ, - rw [smul_add, smul_add, add_add_add_comm, โ†add_smul, hab, one_smul], -end +variables [add_comm_group E] [add_comm_group F] [module ๐•œ E] [module ๐•œ F] -lemma strict_convex.add_right [has_continuous_add E] (hs : strict_convex ๐•œ s) (z : E) : - strict_convex ๐•œ ((ฮป x, x + z) '' s) := -by simpa only [add_comm] using hs.add_left z +section continuous_add +variables [has_continuous_add E] {s t : set E} -lemma strict_convex.add [has_continuous_add E] {t : set E} (hs : strict_convex ๐•œ s) - (ht : strict_convex ๐•œ t) : +lemma strict_convex.add (hs : strict_convex ๐•œ s) (ht : strict_convex ๐•œ t) : strict_convex ๐•œ (s + t) := begin rintro _ โŸจv, w, hv, hw, rflโŸฉ _ โŸจx, y, hx, hy, rflโŸฉ h a b ha hb hab, rw [smul_add, smul_add, add_add_add_comm], obtain rfl | hvx := eq_or_ne v x, - { rw convex.combo_self hab, - suffices : v + (a โ€ข w + b โ€ข y) โˆˆ interior ({v} + t), - { exact interior_mono (add_subset_add (singleton_subset_iff.2 hv) (subset.refl _)) this }, - rw singleton_add, + { refine interior_mono (add_subset_add (singleton_subset_iff.2 hv) subset.rfl) _, + rw [convex.combo_self hab, singleton_add], exact (is_open_map_add_left _).image_interior_subset _ (mem_image_of_mem _ $ ht hw hy (ne_of_apply_ne _ h) ha hb hab) }, - obtain rfl | hwy := eq_or_ne w y, - { rw convex.combo_self hab, - suffices : a โ€ข v + b โ€ข x + w โˆˆ interior (s + {w}), - { exact interior_mono (add_subset_add (subset.refl _) (singleton_subset_iff.2 hw)) this }, - rw add_singleton, - exact (is_open_map_add_right _).image_interior_subset _ - (mem_image_of_mem _ $ hs hv hx hvx ha hb hab) }, - exact subset_interior_add (add_mem_add (hs hv hx hvx ha hb hab) $ ht hw hy hwy ha hb hab), + exact subset_interior_add_left (add_mem_add (hs hv hx hvx ha hb hab) $ + ht.convex hw hy ha.le hb.le hab) +end + +lemma strict_convex.add_left (hs : strict_convex ๐•œ s) (z : E) : + strict_convex ๐•œ ((ฮป x, z + x) '' s) := +by simpa only [singleton_add] using (strict_convex_singleton z).add hs + +lemma strict_convex.add_right (hs : strict_convex ๐•œ s) (z : E) : + strict_convex ๐•œ ((ฮป x, x + z) '' s) := +by simpa only [add_comm] using hs.add_left z + +/-- The translation of a strictly convex set is also strictly convex. -/ +lemma strict_convex.vadd (hs : strict_convex ๐•œ s) (x : E) : strict_convex ๐•œ (x +แตฅ s) := +hs.add_left x + +end continuous_add + +section continuous_smul +variables [linear_ordered_field ๐•] [module ๐• E] [has_continuous_const_smul ๐• E] + [linear_map.compatible_smul E E ๐•œ ๐•] {s : set E} {x : E} + +lemma strict_convex.smul (hs : strict_convex ๐•œ s) (c : ๐•) : strict_convex ๐•œ (c โ€ข s) := +begin + obtain rfl | hc := eq_or_ne c 0, + { exact (subsingleton_zero_smul_set _).strict_convex }, + { exact hs.linear_image (linear_map.lsmul _ _ c) (is_open_map_smulโ‚€ hc) } end +lemma strict_convex.affinity [has_continuous_add E] (hs : strict_convex ๐•œ s) (z : E) (c : ๐•) : + strict_convex ๐•œ (z +แตฅ c โ€ข s) := +(hs.smul c).vadd z + +end continuous_smul end add_comm_group end ordered_semiring @@ -308,7 +320,7 @@ begin exact mem_image_of_mem _ โŸจhtโ‚€, htโ‚โŸฉ, end -/-- The preimage of a strict_convex set under an affine map is strict_convex. -/ +/-- The preimage of a strictly convex set under an affine map is strictly convex. -/ lemma strict_convex.affine_preimage {s : set F} (hs : strict_convex ๐•œ s) {f : E โ†’แตƒ[๐•œ] F} (hf : continuous f) (hfinj : injective f) : strict_convex ๐•œ (f โปยน' s) := @@ -319,7 +331,7 @@ begin exact hs hx hy (hfinj.ne hxy) ha hb hab, end -/-- The image of a strict_convex set under an affine map is strict_convex. -/ +/-- The image of a strictly convex set under an affine map is strictly convex. -/ lemma strict_convex.affine_image (hs : strict_convex ๐•œ s) {f : E โ†’แตƒ[๐•œ] F} (hf : is_open_map f) : strict_convex ๐•œ (f '' s) := begin @@ -345,24 +357,7 @@ variables [linear_ordered_field ๐•œ] [topological_space E] section add_comm_group variables [add_comm_group E] [add_comm_group F] [module ๐•œ E] [module ๐•œ F] {s : set E} {x : E} -lemma strict_convex.smul [has_continuous_const_smul ๐•œ E] (hs : strict_convex ๐•œ s) - (c : ๐•œ) : - strict_convex ๐•œ (c โ€ข s) := -begin - obtain rfl | hc := eq_or_ne c 0, - { exact (subsingleton_zero_smul_set _).strict_convex }, - { exact hs.linear_image (linear_map.lsmul _ _ c) (is_open_map_smulโ‚€ hc) } -end - -lemma strict_convex.affinity [has_continuous_add E] [has_continuous_const_smul ๐•œ E] - (hs : strict_convex ๐•œ s) (z : E) (c : ๐•œ) : - strict_convex ๐•œ ((ฮป x, z + c โ€ข x) '' s) := -begin - have h := (hs.smul c).add_left z, - rwa [โ†image_smul, image_image] at h, -end - -/-- Alternative definition of set strict_convexity, using division. -/ +/-- Alternative definition of set strict convexity, using division. -/ lemma strict_convex_iff_div : strict_convex ๐•œ s โ†” s.pairwise (ฮป x y, โˆ€ โฆƒa b : ๐•œโฆ„, 0 < a โ†’ 0 < b โ†’ (a / (a + b)) โ€ข x + (b / (a + b)) โ€ข y โˆˆ interior s) :=