diff --git a/src/analysis/convex/cone.lean b/src/analysis/convex/cone.lean index 66c018ce84c98..b123863a1754e 100644 --- a/src/analysis/convex/cone.lean +++ b/src/analysis/convex/cone.lean @@ -74,35 +74,24 @@ variables [ordered_semiring π•œ] [add_comm_monoid E] section has_smul variables [has_smul π•œ E] (S T : convex_cone π•œ E) -instance : has_coe (convex_cone π•œ E) (set E) := ⟨convex_cone.carrier⟩ - -instance : has_mem E (convex_cone π•œ E) := ⟨λ m S, m ∈ S.carrier⟩ - -instance : has_le (convex_cone π•œ E) := ⟨λ S T, (S : set E) βŠ† T⟩ - -instance : has_lt (convex_cone π•œ E) := ⟨λ S T, (S : set E) βŠ‚ T⟩ - -@[simp, norm_cast] lemma mem_coe {x : E} : x ∈ (S : set E) ↔ x ∈ S := iff.rfl +instance : set_like (convex_cone π•œ E) E := +{ coe := carrier, + coe_injective' := Ξ» S T h, by cases S; cases T; congr' } @[simp] lemma coe_mk {s : set E} {h₁ hβ‚‚} : ↑(@mk π•œ _ _ _ _ s h₁ hβ‚‚) = s := rfl @[simp] lemma mem_mk {s : set E} {h₁ hβ‚‚ x} : x ∈ @mk π•œ _ _ _ _ s h₁ hβ‚‚ ↔ x ∈ s := iff.rfl -/-- Two `convex_cone`s are equal if the underlying sets are equal. -/ -theorem ext' {S T : convex_cone π•œ E} (h : (S : set E) = T) : S = T := -by cases S; cases T; congr' - -/-- Two `convex_cone`s are equal if and only if the underlying sets are equal. -/ -protected theorem ext'_iff {S T : convex_cone π•œ E} : (S : set E) = T ↔ S = T := -⟨ext', Ξ» h, h β–Έ rfl⟩ - /-- Two `convex_cone`s are equal if they have the same elements. -/ -@[ext] theorem ext {S T : convex_cone π•œ E} (h : βˆ€ x, x ∈ S ↔ x ∈ T) : S = T := ext' $ set.ext h +@[ext] theorem ext {S T : convex_cone π•œ E} (h : βˆ€ x, x ∈ S ↔ x ∈ T) : S = T := set_like.ext h lemma smul_mem {c : π•œ} {x : E} (hc : 0 < c) (hx : x ∈ S) : c β€’ x ∈ S := S.smul_mem' hc hx lemma add_mem ⦃x⦄ (hx : x ∈ S) ⦃y⦄ (hy : y ∈ S) : x + y ∈ S := S.add_mem' hx hy +instance : add_mem_class (convex_cone π•œ E) E := +{ add_mem := Ξ» c a b ha hb, add_mem c ha hb } + instance : has_inf (convex_cone π•œ E) := ⟨λ S T, ⟨S ∩ T, Ξ» c hc x hx, ⟨S.smul_mem hc hx.1, T.smul_mem hc hx.2⟩, Ξ» x hx y hy, ⟨S.add_mem hx.1 hy.1, T.add_mem hx.2 hy.2⟩⟩⟩ @@ -161,7 +150,7 @@ instance : complete_lattice (convex_cone π•œ E) := Sup_le := Ξ» s p hs x hx, mem_Inf.1 hx p hs, le_Inf := Ξ» s a ha x hx, mem_Inf.2 $ Ξ» t ht, ha t ht hx, Inf_le := Ξ» s a ha x hx, mem_Inf.1 hx _ ha, - .. partial_order.lift (coe : convex_cone π•œ E β†’ set E) (Ξ» a b, ext') } + .. set_like.partial_order } instance : inhabited (convex_cone π•œ E) := ⟨βŠ₯⟩ @@ -204,9 +193,10 @@ def map (f : E β†’β‚—[π•œ] F) (S : convex_cone π•œ E) : convex_cone π•œ F := lemma map_map (g : F β†’β‚—[π•œ] G) (f : E β†’β‚—[π•œ] F) (S : convex_cone π•œ E) : (S.map f).map g = S.map (g.comp f) := -ext' $ image_image g f S +set_like.coe_injective $ image_image g f S -@[simp] lemma map_id (S : convex_cone π•œ E) : S.map linear_map.id = S := ext' $ image_id _ +@[simp] lemma map_id (S : convex_cone π•œ E) : S.map linear_map.id = S := +set_like.coe_injective $ image_id _ /-- The preimage of a convex cone under a `π•œ`-linear map is a convex cone. -/ def comap (f : E β†’β‚—[π•œ] F) (S : convex_cone π•œ F) : convex_cone π•œ E := @@ -214,11 +204,12 @@ def comap (f : E β†’β‚—[π•œ] F) (S : convex_cone π•œ F) : convex_cone π•œ E : smul_mem' := Ξ» c hc x hx, by { rw [mem_preimage, f.map_smul c], exact S.smul_mem hc hx }, add_mem' := Ξ» x hx y hy, by { rw [mem_preimage, f.map_add], exact S.add_mem hx hy } } -@[simp] lemma comap_id (S : convex_cone π•œ E) : S.comap linear_map.id = S := ext' preimage_id +@[simp] lemma comap_id (S : convex_cone π•œ E) : S.comap linear_map.id = S := +set_like.coe_injective preimage_id lemma comap_comap (g : F β†’β‚—[π•œ] G) (f : E β†’β‚—[π•œ] F) (S : convex_cone π•œ G) : (S.comap g).comap f = S.comap (g.comp f) := -ext' $ preimage_comp.symm +set_like.coe_injective $ preimage_comp.symm @[simp] lemma mem_comap {f : E β†’β‚—[π•œ] F} {S : convex_cone π•œ F} {x : E} : x ∈ S.comap f ↔ f x ∈ S := iff.rfl @@ -430,7 +421,7 @@ end lemma convex_hull_to_cone_eq_Inf (s : set E) : (convex_convex_hull π•œ s).to_cone _ = Inf {t : convex_cone π•œ E | s βŠ† t} := -(convex_hull_to_cone_is_least s).is_glb.Inf_eq.symm +eq.symm $ is_glb.Inf_eq $ is_least.is_glb $ convex_hull_to_cone_is_least s end cone_from_convex @@ -620,8 +611,7 @@ lemma mem_inner_dual_cone (y : H) (s : set H) : y ∈ s.inner_dual_cone ↔ βˆ€ x ∈ s, 0 ≀ βŸͺ x, y ⟫ := by refl @[simp] lemma inner_dual_cone_empty : (βˆ… : set H).inner_dual_cone = ⊀ := -convex_cone.ext' (eq_univ_of_forall - (Ξ» x y hy, false.elim (set.not_mem_empty _ hy))) +eq_top_iff.mpr $ Ξ» x hy y, false.elim lemma inner_dual_cone_le_inner_dual_cone (h : t βŠ† s) : s.inner_dual_cone ≀ t.inner_dual_cone := @@ -639,7 +629,7 @@ begin { refine set.mem_Inter.2 (Ξ» i, set.mem_Inter.2 (Ξ» hi _, _)), rintro ⟨ ⟩, exact hx i hi }, - { simp only [set.mem_Inter, convex_cone.mem_coe, mem_inner_dual_cone, + { simp only [set.mem_Inter, set_like.mem_coe, mem_inner_dual_cone, set.mem_singleton_iff, forall_eq, imp_self] } end