Documentation

Mathlib.Analysis.Convex.Star

Star-convex sets #

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, 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 StarConvex (π•œ : Type u_1) {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] (x : E) (s : Set E) :

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
Instances For
    theorem starConvex_iff_segment_subset {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {s : Set E} :
    StarConvex π•œ x s ↔ βˆ€ ⦃y : E⦄, y ∈ s β†’ segment π•œ x y βŠ† s
    theorem StarConvex.segment_subset {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {s : Set E} (h : StarConvex π•œ x s) {y : E} (hy : y ∈ s) :
    segment π•œ x y βŠ† s
    theorem StarConvex.openSegment_subset {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {s : Set E} (h : StarConvex π•œ x s) {y : E} (hy : y ∈ s) :
    openSegment π•œ x y βŠ† s
    theorem starConvex_iff_pointwise_add_subset {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {s : Set E} :
    StarConvex π•œ 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 starConvex_empty {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] (x : E) :
    StarConvex π•œ x βˆ…
    theorem starConvex_univ {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] (x : E) :
    StarConvex π•œ x Set.univ
    theorem StarConvex.inter {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {s t : Set E} (hs : StarConvex π•œ x s) (ht : StarConvex π•œ x t) :
    StarConvex π•œ x (s ∩ t)
    theorem starConvex_sInter {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {S : Set (Set E)} (h : βˆ€ s ∈ S, StarConvex π•œ x s) :
    StarConvex π•œ x (β‹‚β‚€ S)
    theorem starConvex_iInter {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {ΞΉ : Sort u_4} {s : ΞΉ β†’ Set E} (h : βˆ€ (i : ΞΉ), StarConvex π•œ x (s i)) :
    StarConvex π•œ x (β‹‚ (i : ΞΉ), s i)
    theorem starConvex_iInterβ‚‚ {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {ΞΉ : Sort u_4} {ΞΊ : ΞΉ β†’ Sort u_5} {s : (i : ΞΉ) β†’ ΞΊ i β†’ Set E} (h : βˆ€ (i : ΞΉ) (j : ΞΊ i), StarConvex π•œ x (s i j)) :
    StarConvex π•œ x (β‹‚ (i : ΞΉ), β‹‚ (j : ΞΊ i), s i j)
    theorem StarConvex.union {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {s t : Set E} (hs : StarConvex π•œ x s) (ht : StarConvex π•œ x t) :
    StarConvex π•œ x (s βˆͺ t)
    theorem starConvex_iUnion {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {ΞΉ : Sort u_4} {s : ΞΉ β†’ Set E} (hs : βˆ€ (i : ΞΉ), StarConvex π•œ x (s i)) :
    StarConvex π•œ x (⋃ (i : ΞΉ), s i)
    theorem starConvex_iUnionβ‚‚ {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {ΞΉ : Sort u_4} {ΞΊ : ΞΉ β†’ Sort u_5} {s : (i : ΞΉ) β†’ ΞΊ i β†’ Set E} (h : βˆ€ (i : ΞΉ) (j : ΞΊ i), StarConvex π•œ x (s i j)) :
    StarConvex π•œ x (⋃ (i : ΞΉ), ⋃ (j : ΞΊ i), s i j)
    theorem starConvex_sUnion {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [SMul π•œ E] {x : E} {S : Set (Set E)} (hS : βˆ€ s ∈ S, StarConvex π•œ x s) :
    StarConvex π•œ x (⋃₀ S)
    theorem StarConvex.prod {π•œ : Type u_1} {E : Type u_2} {F : Type u_3} [OrderedSemiring π•œ] [AddCommMonoid E] [AddCommMonoid F] [SMul π•œ E] [SMul π•œ F] {x : E} {y : F} {s : Set E} {t : Set F} (hs : StarConvex π•œ x s) (ht : StarConvex π•œ y t) :
    StarConvex π•œ (x, y) (s Γ—Λ’ t)
    theorem starConvex_pi {π•œ : Type u_1} [OrderedSemiring π•œ] {ΞΉ : Type u_4} {E : ΞΉ β†’ Type u_5} [(i : ΞΉ) β†’ AddCommMonoid (E i)] [(i : ΞΉ) β†’ SMul π•œ (E i)] {x : (i : ΞΉ) β†’ E i} {s : Set ΞΉ} {t : (i : ΞΉ) β†’ Set (E i)} (ht : βˆ€ ⦃i : ι⦄, i ∈ s β†’ StarConvex π•œ (x i) (t i)) :
    StarConvex π•œ x (s.pi t)
    theorem StarConvex.mem {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x : E} {s : Set E} (hs : StarConvex π•œ x s) (h : s.Nonempty) :
    x ∈ s
    theorem starConvex_iff_forall_pos {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x : E} {s : Set E} (hx : x ∈ s) :
    StarConvex π•œ x s ↔ βˆ€ ⦃y : E⦄, y ∈ s β†’ βˆ€ ⦃a b : π•œβ¦„, 0 < a β†’ 0 < b β†’ a + b = 1 β†’ a β€’ x + b β€’ y ∈ s
    theorem starConvex_iff_forall_ne_pos {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x : E} {s : Set E} (hx : x ∈ s) :
    StarConvex π•œ x s ↔ βˆ€ ⦃y : E⦄, y ∈ s β†’ x β‰  y β†’ βˆ€ ⦃a b : π•œβ¦„, 0 < a β†’ 0 < b β†’ a + b = 1 β†’ a β€’ x + b β€’ y ∈ s
    theorem starConvex_iff_openSegment_subset {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x : E} {s : Set E} (hx : x ∈ s) :
    StarConvex π•œ x s ↔ βˆ€ ⦃y : E⦄, y ∈ s β†’ openSegment π•œ x y βŠ† s
    theorem starConvex_singleton {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] (x : E) :
    StarConvex π•œ x {x}
    theorem StarConvex.linear_image {π•œ : Type u_1} {E : Type u_2} {F : Type u_3} [OrderedSemiring π•œ] [AddCommMonoid E] [AddCommMonoid F] [Module π•œ E] [Module π•œ F] {x : E} {s : Set E} (hs : StarConvex π•œ x s) (f : E β†’β‚—[π•œ] F) :
    StarConvex π•œ (f x) (⇑f '' s)
    theorem StarConvex.is_linear_image {π•œ : Type u_1} {E : Type u_2} {F : Type u_3} [OrderedSemiring π•œ] [AddCommMonoid E] [AddCommMonoid F] [Module π•œ E] [Module π•œ F] {x : E} {s : Set E} (hs : StarConvex π•œ x s) {f : E β†’ F} (hf : IsLinearMap π•œ f) :
    StarConvex π•œ (f x) (f '' s)
    theorem StarConvex.linear_preimage {π•œ : Type u_1} {E : Type u_2} {F : Type u_3} [OrderedSemiring π•œ] [AddCommMonoid E] [AddCommMonoid F] [Module π•œ E] [Module π•œ F] {x : E} {s : Set F} (f : E β†’β‚—[π•œ] F) (hs : StarConvex π•œ (f x) s) :
    StarConvex π•œ x (⇑f ⁻¹' s)
    theorem StarConvex.is_linear_preimage {π•œ : Type u_1} {E : Type u_2} {F : Type u_3} [OrderedSemiring π•œ] [AddCommMonoid E] [AddCommMonoid F] [Module π•œ E] [Module π•œ F] {x : E} {s : Set F} {f : E β†’ F} (hs : StarConvex π•œ (f x) s) (hf : IsLinearMap π•œ f) :
    StarConvex π•œ x (f ⁻¹' s)
    theorem StarConvex.add {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x y : E} {s t : Set E} (hs : StarConvex π•œ x s) (ht : StarConvex π•œ y t) :
    StarConvex π•œ (x + y) (s + t)
    theorem StarConvex.add_left {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x : E} {s : Set E} (hs : StarConvex π•œ x s) (z : E) :
    StarConvex π•œ (z + x) ((fun (x : E) => z + x) '' s)
    theorem StarConvex.add_right {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x : E} {s : Set E} (hs : StarConvex π•œ x s) (z : E) :
    StarConvex π•œ (x + z) ((fun (x : E) => x + z) '' s)
    theorem StarConvex.preimage_add_right {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x z : E} {s : Set E} (hs : StarConvex π•œ (z + x) s) :
    StarConvex π•œ x ((fun (x : E) => z + x) ⁻¹' s)

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

    theorem StarConvex.preimage_add_left {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x z : E} {s : Set E} (hs : StarConvex π•œ (x + z) s) :
    StarConvex π•œ x ((fun (x : E) => x + z) ⁻¹' s)

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

    theorem StarConvex.sub' {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [AddCommGroup E] [Module π•œ E] {x y : E} {s : Set (E Γ— E)} (hs : StarConvex π•œ (x, y) s) :
    StarConvex π•œ (x - y) ((fun (x : E Γ— E) => x.1 - x.2) '' s)
    theorem StarConvex.smul {π•œ : Type u_1} {E : Type u_2} [OrderedCommSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x : E} {s : Set E} (hs : StarConvex π•œ x s) (c : π•œ) :
    StarConvex π•œ (c β€’ x) (c β€’ s)
    theorem StarConvex.zero_smul {π•œ : Type u_1} {E : Type u_2} [OrderedCommSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {s : Set E} (hs : StarConvex π•œ 0 s) (c : π•œ) :
    StarConvex π•œ 0 (c β€’ s)
    theorem StarConvex.preimage_smul {π•œ : Type u_1} {E : Type u_2} [OrderedCommSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x : E} {s : Set E} {c : π•œ} (hs : StarConvex π•œ (c β€’ x) s) :
    StarConvex π•œ x ((fun (z : E) => c β€’ z) ⁻¹' s)
    theorem StarConvex.affinity {π•œ : Type u_1} {E : Type u_2} [OrderedCommSemiring π•œ] [AddCommMonoid E] [Module π•œ E] {x : E} {s : Set E} (hs : StarConvex π•œ x s) (z : E) (c : π•œ) :
    StarConvex π•œ (z + c β€’ x) ((fun (x : E) => z + c β€’ x) '' s)
    theorem starConvex_zero_iff {π•œ : Type u_1} {E : Type u_2} [OrderedRing π•œ] [AddCommMonoid E] [SMulWithZero π•œ E] {s : Set E} :
    StarConvex π•œ 0 s ↔ βˆ€ ⦃x : E⦄, x ∈ s β†’ βˆ€ ⦃a : π•œβ¦„, 0 ≀ a β†’ a ≀ 1 β†’ a β€’ x ∈ s
    theorem StarConvex.add_smul_mem {π•œ : Type u_1} {E : Type u_2} [OrderedRing π•œ] [AddCommGroup E] [Module π•œ E] {x y : E} {s : Set E} (hs : StarConvex π•œ x s) (hy : x + y ∈ s) {t : π•œ} (htβ‚€ : 0 ≀ t) (ht₁ : t ≀ 1) :
    x + t β€’ y ∈ s
    theorem StarConvex.smul_mem {π•œ : Type u_1} {E : Type u_2} [OrderedRing π•œ] [AddCommGroup E] [Module π•œ E] {x : E} {s : Set E} (hs : StarConvex π•œ 0 s) (hx : x ∈ s) {t : π•œ} (htβ‚€ : 0 ≀ t) (ht₁ : t ≀ 1) :
    theorem StarConvex.add_smul_sub_mem {π•œ : Type u_1} {E : Type u_2} [OrderedRing π•œ] [AddCommGroup E] [Module π•œ E] {x y : E} {s : Set E} (hs : StarConvex π•œ x s) (hy : y ∈ s) {t : π•œ} (htβ‚€ : 0 ≀ t) (ht₁ : t ≀ 1) :
    x + t β€’ (y - x) ∈ s
    theorem StarConvex.affine_preimage {π•œ : Type u_1} {E : Type u_2} {F : Type u_3} [OrderedRing π•œ] [AddCommGroup E] [AddCommGroup F] [Module π•œ E] [Module π•œ F] {x : E} (f : E →ᡃ[π•œ] F) {s : Set F} (hs : StarConvex π•œ (f x) s) :
    StarConvex π•œ x (⇑f ⁻¹' s)

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

    theorem StarConvex.affine_image {π•œ : Type u_1} {E : Type u_2} {F : Type u_3} [OrderedRing π•œ] [AddCommGroup E] [AddCommGroup F] [Module π•œ E] [Module π•œ F] {x : E} (f : E →ᡃ[π•œ] F) {s : Set E} (hs : StarConvex π•œ x s) :
    StarConvex π•œ (f x) (⇑f '' s)

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

    theorem StarConvex.neg {π•œ : Type u_1} {E : Type u_2} [OrderedRing π•œ] [AddCommGroup E] [Module π•œ E] {x : E} {s : Set E} (hs : StarConvex π•œ x s) :
    StarConvex π•œ (-x) (-s)
    theorem StarConvex.sub {π•œ : Type u_1} {E : Type u_2} [OrderedRing π•œ] [AddCommGroup E] [Module π•œ E] {x y : E} {s t : Set E} (hs : StarConvex π•œ x s) (ht : StarConvex π•œ y t) :
    StarConvex π•œ (x - y) (s - t)
    theorem starConvex_compl_Iic {π•œ : Type u_1} {E : Type u_2} [OrderedRing π•œ] [OrderedAddCommGroup E] [Module π•œ E] [OrderedSMul π•œ E] {x y : E} (h : x < y) :
    StarConvex π•œ y (Set.Iic x)ᢜ

    If x < y, then (Set.Iic x)ᢜ is star convex at y.

    theorem starConvex_compl_Ici {π•œ : Type u_1} {E : Type u_2} [OrderedRing π•œ] [OrderedAddCommGroup E] [Module π•œ E] [OrderedSMul π•œ E] {x y : E} (h : x < y) :
    StarConvex π•œ x (Set.Ici y)ᢜ

    If x < y, then (Set.Ici y)ᢜ is star convex at x.

    theorem starConvex_iff_div {π•œ : Type u_1} {E : Type u_2} [LinearOrderedField π•œ] [AddCommGroup E] [Module π•œ E] {x : E} {s : Set E} :
    StarConvex π•œ 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 StarConvex.mem_smul {π•œ : Type u_1} {E : Type u_2} [LinearOrderedField π•œ] [AddCommGroup E] [Module π•œ E] {x : E} {s : Set E} (hs : StarConvex π•œ 0 s) (hx : x ∈ s) {t : π•œ} (ht : 1 ≀ t) :

    Star-convex sets in an ordered space #

    Relates starConvex and Set.ordConnected.

    theorem Set.OrdConnected.starConvex {π•œ : Type u_1} {E : Type u_2} [OrderedSemiring π•œ] [OrderedAddCommMonoid E] [Module π•œ E] [OrderedSMul π•œ E] {x : E} {s : Set E} (hs : s.OrdConnected) (hx : x ∈ s) (h : βˆ€ y ∈ s, x ≀ y ∨ y ≀ x) :
    StarConvex π•œ x s

    If s is an order-connected set in an ordered module over an ordered semiring and all elements of s are comparable with x ∈ s, then s is StarConvex at x.

    theorem starConvex_iff_ordConnected {π•œ : Type u_1} [LinearOrderedField π•œ] {x : π•œ} {s : Set π•œ} (hx : x ∈ s) :
    theorem StarConvex.ordConnected {π•œ : Type u_1} [LinearOrderedField π•œ] {x : π•œ} {s : Set π•œ} (hx : x ∈ s) :
    StarConvex π•œ x s β†’ s.OrdConnected

    Alias of the forward direction of starConvex_iff_ordConnected.