mathlib3 documentation

analysis.convex.star

Star-convex sets #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

This files defines star-convex sets (aka star domains, star-shaped set, radially convex set).

A set is star-convex at x if every segment from x to a point in the set is contained in the set.

This is the prototypical example of a contractible set in homotopy theory (by scaling every point towards x), but has wider uses.

Note that this has nothing to do with star rings, has_star and co.

Main declarations #

Implementation notes #

Instead of saying that a set is star-convex, we say a set is star-convex at a point. This has the advantage of allowing us to talk about convexity as being "everywhere star-convexity" and of making the union of star-convex sets be star-convex.

Incidentally, this choice means we don't need to assume a set is nonempty for it to be star-convex. Concretely, the empty set is star-convex at every point.

TODO #

Balanced sets are star-convex.

The closure of a star-convex set is star-convex.

Star-convex sets are contractible.

A nonempty open star-convex set in ā„^n is diffeomorphic to the entire space.

def star_convex (š•œ : Type u_1) {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] (x : E) (s : set E) :
Prop

Star-convexity of sets. s is star-convex at x if every segment from x to a point in s is contained in s.

Equations
theorem star_convex_iff_segment_subset {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {s : set E} :
star_convex š•œ x s ↔ āˆ€ ⦃y : E⦄, y ∈ s → segment š•œ x y āŠ† s
theorem star_convex.segment_subset {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {s : set E} (h : star_convex š•œ x s) {y : E} (hy : y ∈ s) :
segment š•œ x y āŠ† s
theorem star_convex.open_segment_subset {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {s : set E} (h : star_convex š•œ x s) {y : E} (hy : y ∈ s) :
open_segment š•œ x y āŠ† s
theorem star_convex_iff_pointwise_add_subset {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {s : set E} :
star_convex š•œ x s ↔ āˆ€ ⦃a b : š•œā¦„, 0 ≤ a → 0 ≤ b → a + b = 1 → a • {x} + b • s āŠ† s

Alternative definition of star-convexity, in terms of pointwise set operations.

theorem star_convex_empty {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] (x : E) :
star_convex š•œ x āˆ…
theorem star_convex_univ {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] (x : E) :
theorem star_convex.inter {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {s t : set E} (hs : star_convex š•œ x s) (ht : star_convex š•œ x t) :
star_convex š•œ x (s ∩ t)
theorem star_convex_sInter {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {S : set (set E)} (h : āˆ€ (s : set E), s ∈ S → star_convex š•œ x s) :
star_convex š•œ x (ā‹‚ā‚€ S)
theorem star_convex_Inter {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {ι : Sort u_3} {s : ι → set E} (h : āˆ€ (i : ι), star_convex š•œ x (s i)) :
star_convex š•œ x (ā‹‚ (i : ι), s i)
theorem star_convex.union {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {s t : set E} (hs : star_convex š•œ x s) (ht : star_convex š•œ x t) :
star_convex š•œ x (s ∪ t)
theorem star_convex_Union {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {ι : Sort u_3} {s : ι → set E} (hs : āˆ€ (i : ι), star_convex š•œ x (s i)) :
star_convex š•œ x (ā‹ƒ (i : ι), s i)
theorem star_convex_sUnion {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [has_smul š•œ E] {x : E} {S : set (set E)} (hS : āˆ€ (s : set E), s ∈ S → star_convex š•œ x s) :
star_convex š•œ x (ā‹ƒā‚€ S)
theorem star_convex.prod {š•œ : Type u_1} {E : Type u_2} {F : Type u_3} [ordered_semiring š•œ] [add_comm_monoid E] [add_comm_monoid F] [has_smul š•œ E] [has_smul š•œ F] {x : E} {y : F} {s : set E} {t : set F} (hs : star_convex š•œ x s) (ht : star_convex š•œ y t) :
star_convex š•œ (x, y) (s Ć—Ė¢ t)
theorem star_convex_pi {š•œ : Type u_1} [ordered_semiring š•œ] {ι : Type u_2} {E : ι → Type u_3} [Ī  (i : ι), add_comm_monoid (E i)] [Ī  (i : ι), has_smul š•œ (E i)] {x : Ī  (i : ι), E i} {s : set ι} {t : Ī  (i : ι), set (E i)} (ht : āˆ€ ⦃i : ι⦄, i ∈ s → star_convex š•œ (x i) (t i)) :
star_convex š•œ x (s.pi t)
theorem star_convex.mem {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x : E} {s : set E} (hs : star_convex š•œ x s) (h : s.nonempty) :
x ∈ s
theorem star_convex_iff_forall_pos {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x : E} {s : set E} (hx : x ∈ s) :
star_convex š•œ x s ↔ āˆ€ ⦃y : E⦄, y ∈ s → āˆ€ ⦃a b : š•œā¦„, 0 < a → 0 < b → a + b = 1 → a • x + b • y ∈ s
theorem star_convex_iff_forall_ne_pos {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x : E} {s : set E} (hx : x ∈ s) :
star_convex š•œ x s ↔ āˆ€ ⦃y : E⦄, y ∈ s → x ≠ y → āˆ€ ⦃a b : š•œā¦„, 0 < a → 0 < b → a + b = 1 → a • x + b • y ∈ s
theorem star_convex_iff_open_segment_subset {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x : E} {s : set E} (hx : x ∈ s) :
star_convex š•œ x s ↔ āˆ€ ⦃y : E⦄, y ∈ s → open_segment š•œ x y āŠ† s
theorem star_convex_singleton {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] (x : E) :
star_convex š•œ x {x}
theorem star_convex.linear_image {š•œ : Type u_1} {E : Type u_2} {F : Type u_3} [ordered_semiring š•œ] [add_comm_monoid E] [add_comm_monoid F] [module š•œ E] [module š•œ F] {x : E} {s : set E} (hs : star_convex š•œ x s) (f : E →ₗ[š•œ] F) :
star_convex š•œ (⇑f x) (⇑f '' s)
theorem star_convex.is_linear_image {š•œ : Type u_1} {E : Type u_2} {F : Type u_3} [ordered_semiring š•œ] [add_comm_monoid E] [add_comm_monoid F] [module š•œ E] [module š•œ F] {x : E} {s : set E} (hs : star_convex š•œ x s) {f : E → F} (hf : is_linear_map š•œ f) :
star_convex š•œ (f x) (f '' s)
theorem star_convex.linear_preimage {š•œ : Type u_1} {E : Type u_2} {F : Type u_3} [ordered_semiring š•œ] [add_comm_monoid E] [add_comm_monoid F] [module š•œ E] [module š•œ F] {x : E} {s : set F} (f : E →ₗ[š•œ] F) (hs : star_convex š•œ (⇑f x) s) :
theorem star_convex.is_linear_preimage {š•œ : Type u_1} {E : Type u_2} {F : Type u_3} [ordered_semiring š•œ] [add_comm_monoid E] [add_comm_monoid F] [module š•œ E] [module š•œ F] {x : E} {s : set F} {f : E → F} (hs : star_convex š•œ (f x) s) (hf : is_linear_map š•œ f) :
star_convex š•œ x (f ⁻¹' s)
theorem star_convex.add {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x y : E} {s t : set E} (hs : star_convex š•œ x s) (ht : star_convex š•œ y t) :
star_convex š•œ (x + y) (s + t)
theorem star_convex.add_left {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x : E} {s : set E} (hs : star_convex š•œ x s) (z : E) :
star_convex š•œ (z + x) ((Ī» (x : E), z + x) '' s)
theorem star_convex.add_right {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x : E} {s : set E} (hs : star_convex š•œ x s) (z : E) :
star_convex š•œ (x + z) ((Ī» (x : E), x + z) '' s)
theorem star_convex.preimage_add_right {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x z : E} {s : set E} (hs : star_convex š•œ (z + x) s) :
star_convex š•œ x ((Ī» (x : E), z + x) ⁻¹' s)

The translation of a star-convex set is also star-convex.

theorem star_convex.preimage_add_left {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x z : E} {s : set E} (hs : star_convex š•œ (x + z) s) :
star_convex š•œ x ((Ī» (x : E), x + z) ⁻¹' s)

The translation of a star-convex set is also star-convex.

theorem star_convex.sub' {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [add_comm_group E] [module š•œ E] {x y : E} {s : set (E Ɨ E)} (hs : star_convex š•œ (x, y) s) :
star_convex š•œ (x - y) ((Ī» (x : E Ɨ E), x.fst - x.snd) '' s)
theorem star_convex.smul {š•œ : Type u_1} {E : Type u_2} [ordered_comm_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x : E} {s : set E} (hs : star_convex š•œ x s) (c : š•œ) :
star_convex š•œ (c • x) (c • s)
theorem star_convex.preimage_smul {š•œ : Type u_1} {E : Type u_2} [ordered_comm_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x : E} {s : set E} {c : š•œ} (hs : star_convex š•œ (c • x) s) :
star_convex š•œ x ((Ī» (z : E), c • z) ⁻¹' s)
theorem star_convex.affinity {š•œ : Type u_1} {E : Type u_2} [ordered_comm_semiring š•œ] [add_comm_monoid E] [module š•œ E] {x : E} {s : set E} (hs : star_convex š•œ x s) (z : E) (c : š•œ) :
star_convex š•œ (z + c • x) ((Ī» (x : E), z + c • x) '' s)
theorem star_convex_zero_iff {š•œ : Type u_1} {E : Type u_2} [ordered_ring š•œ] [add_comm_monoid E] [smul_with_zero š•œ E] {s : set E} :
star_convex š•œ 0 s ↔ āˆ€ ⦃x : E⦄, x ∈ s → āˆ€ ⦃a : š•œā¦„, 0 ≤ a → a ≤ 1 → a • x ∈ s
theorem star_convex.add_smul_mem {š•œ : Type u_1} {E : Type u_2} [ordered_ring š•œ] [add_comm_group E] [module š•œ E] {x y : E} {s : set E} (hs : star_convex š•œ x s) (hy : x + y ∈ s) {t : š•œ} (htā‚€ : 0 ≤ t) (ht₁ : t ≤ 1) :
x + t • y ∈ s
theorem star_convex.smul_mem {š•œ : Type u_1} {E : Type u_2} [ordered_ring š•œ] [add_comm_group E] [module š•œ E] {x : E} {s : set E} (hs : star_convex š•œ 0 s) (hx : x ∈ s) {t : š•œ} (htā‚€ : 0 ≤ t) (ht₁ : t ≤ 1) :
theorem star_convex.add_smul_sub_mem {š•œ : Type u_1} {E : Type u_2} [ordered_ring š•œ] [add_comm_group E] [module š•œ E] {x y : E} {s : set E} (hs : star_convex š•œ x s) (hy : y ∈ s) {t : š•œ} (htā‚€ : 0 ≤ t) (ht₁ : t ≤ 1) :
x + t • (y - x) ∈ s
theorem star_convex.affine_preimage {š•œ : Type u_1} {E : Type u_2} {F : Type u_3} [ordered_ring š•œ] [add_comm_group E] [add_comm_group F] [module š•œ E] [module š•œ F] {x : E} (f : E ā†’įµƒ[š•œ] F) {s : set F} (hs : star_convex š•œ (⇑f x) s) :

The preimage of a star-convex set under an affine map is star-convex.

theorem star_convex.affine_image {š•œ : Type u_1} {E : Type u_2} {F : Type u_3} [ordered_ring š•œ] [add_comm_group E] [add_comm_group F] [module š•œ E] [module š•œ F] {x : E} (f : E ā†’įµƒ[š•œ] F) {s : set E} (hs : star_convex š•œ x s) :
star_convex š•œ (⇑f x) (⇑f '' s)

The image of a star-convex set under an affine map is star-convex.

theorem star_convex.neg {š•œ : Type u_1} {E : Type u_2} [ordered_ring š•œ] [add_comm_group E] [module š•œ E] {x : E} {s : set E} (hs : star_convex š•œ x s) :
star_convex š•œ (-x) (-s)
theorem star_convex.sub {š•œ : Type u_1} {E : Type u_2} [ordered_ring š•œ] [add_comm_group E] [module š•œ E] {x y : E} {s t : set E} (hs : star_convex š•œ x s) (ht : star_convex š•œ y t) :
star_convex š•œ (x - y) (s - t)
theorem star_convex_iff_div {š•œ : Type u_1} {E : Type u_2} [linear_ordered_field š•œ] [add_comm_group E] [module š•œ E] {x : E} {s : set E} :
star_convex š•œ x s ↔ āˆ€ ⦃y : E⦄, y ∈ s → āˆ€ ⦃a b : š•œā¦„, 0 ≤ a → 0 ≤ b → 0 < a + b → (a / (a + b)) • x + (b / (a + b)) • y ∈ s

Alternative definition of star-convexity, using division.

theorem star_convex.mem_smul {š•œ : Type u_1} {E : Type u_2} [linear_ordered_field š•œ] [add_comm_group E] [module š•œ E] {x : E} {s : set E} (hs : star_convex š•œ 0 s) (hx : x ∈ s) {t : š•œ} (ht : 1 ≤ t) :

Star-convex sets in an ordered space #

Relates star_convex and set.ord_connected.

theorem set.ord_connected.star_convex {š•œ : Type u_1} {E : Type u_2} [ordered_semiring š•œ] [ordered_add_comm_monoid E] [module š•œ E] [ordered_smul š•œ E] {x : E} {s : set E} (hs : s.ord_connected) (hx : x ∈ s) (h : āˆ€ (y : E), y ∈ s → x ≤ y ∨ y ≤ x) :
star_convex š•œ x s
theorem star_convex_iff_ord_connected {š•œ : Type u_1} [linear_ordered_field š•œ] {x : š•œ} {s : set š•œ} (hx : x ∈ s) :
theorem star_convex.ord_connected {š•œ : Type u_1} [linear_ordered_field š•œ] {x : š•œ} {s : set š•œ} (hx : x ∈ s) :

Alias of the forward direction of star_convex_iff_ord_connected.