Skip to content

Commit 8b8aaa2

Browse files
committed
chore: rename lemmas containing "of_open" to match the naming convention (#8229)
Mostly, this means replacing "of_open" by "of_isOpen". A few lemmas names were misleading and are corrected differently. [Zulip discussion](https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Naming.20convention/near/402702125).
1 parent 42649c4 commit 8b8aaa2

File tree

47 files changed

+188
-183
lines changed

Some content is hidden

Large Commits have some content hidden by default. Use the searchbox below for content that may be hidden.

47 files changed

+188
-183
lines changed

Mathlib/AlgebraicGeometry/OpenImmersion.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -267,7 +267,7 @@ theorem affineBasisCover_is_basis (X : Scheme) :
267267
TopologicalSpace.IsTopologicalBasis
268268
{x : Set X |
269269
∃ a : X.affineBasisCover.J, x = Set.range (X.affineBasisCover.map a).1.base} := by
270-
apply TopologicalSpace.isTopologicalBasis_of_open_of_nhds
270+
apply TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds
271271
· rintro _ ⟨a, rfl⟩
272272
exact IsOpenImmersion.open_range (X.affineBasisCover.map a)
273273
· rintro a U haU hU

Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -808,7 +808,7 @@ theorem basicOpen_pow (f : R) (n : ℕ) (hn : 0 < n) : basicOpen (f ^ n) = basic
808808
theorem isTopologicalBasis_basic_opens :
809809
TopologicalSpace.IsTopologicalBasis
810810
(Set.range fun r : R => (basicOpen r : Set (PrimeSpectrum R))) := by
811-
apply TopologicalSpace.isTopologicalBasis_of_open_of_nhds
811+
apply TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds
812812
· rintro _ ⟨r, rfl⟩
813813
exact isOpen_basicOpen
814814
· rintro p U hp ⟨s, hs⟩

Mathlib/AlgebraicGeometry/ProjectiveSpectrum/Topology.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -453,7 +453,7 @@ theorem basicOpen_eq_union_of_projection (f : A) :
453453
theorem isTopologicalBasis_basic_opens :
454454
TopologicalSpace.IsTopologicalBasis
455455
(Set.range fun r : A => (basicOpen 𝒜 r : Set (ProjectiveSpectrum 𝒜))) := by
456-
apply TopologicalSpace.isTopologicalBasis_of_open_of_nhds
456+
apply TopologicalSpace.isTopologicalBasis_of_isOpen_of_nhds
457457
· rintro _ ⟨r, rfl⟩
458458
exact isOpen_basicOpen 𝒜
459459
· rintro p U hp ⟨s, hs⟩

Mathlib/Analysis/Calculus/ContDiff/Basic.lean

Lines changed: 14 additions & 14 deletions
Original file line numberDiff line numberDiff line change
@@ -1833,7 +1833,7 @@ theorem LocalHomeomorph.contDiffAt_symm [CompleteSpace E] (f : LocalHomeomorph E
18331833
obtain ⟨t, htu, ht, htf⟩ := mem_nhds_iff.mp (Filter.inter_mem hu h_nhds)
18341834
use f.target ∩ f.symm ⁻¹' t
18351835
refine' ⟨IsOpen.mem_nhds _ _, _⟩
1836-
· exact f.preimage_open_of_open_symm ht
1836+
· exact f.isOpen_inter_preimage_symm ht
18371837
· exact mem_inter ha (mem_preimage.mpr htf)
18381838
intro x hx
18391839
obtain ⟨hxu, e, he⟩ := htu hx.2
@@ -1937,11 +1937,11 @@ theorem contDiffOn_succ_iff_derivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 s₂)
19371937

19381938
/-- A function is `C^(n + 1)` on an open domain if and only if it is
19391939
differentiable there, and its derivative (formulated with `deriv`) is `C^n`. -/
1940-
theorem contDiffOn_succ_iff_deriv_of_open {n : ℕ} (hs : IsOpen s₂) :
1940+
theorem contDiffOn_succ_iff_deriv_of_isOpen {n : ℕ} (hs : IsOpen s₂) :
19411941
ContDiffOn 𝕜 (n + 1 : ℕ) f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 n (deriv f₂) s₂ := by
19421942
rw [contDiffOn_succ_iff_derivWithin hs.uniqueDiffOn]
1943-
exact Iff.rfl.and (contDiffOn_congr fun _ => derivWithin_of_open hs)
1944-
#align cont_diff_on_succ_iff_deriv_of_open contDiffOn_succ_iff_deriv_of_open
1943+
exact Iff.rfl.and (contDiffOn_congr fun _ => derivWithin_of_isOpen hs)
1944+
#align cont_diff_on_succ_iff_deriv_of_open contDiffOn_succ_iff_deriv_of_isOpen
19451945

19461946
/-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable
19471947
there, and its derivative (formulated with `derivWithin`) is `C^∞`. -/
@@ -1961,11 +1961,11 @@ theorem contDiffOn_top_iff_derivWithin (hs : UniqueDiffOn 𝕜 s₂) :
19611961

19621962
/-- A function is `C^∞` on an open domain if and only if it is differentiable
19631963
there, and its derivative (formulated with `deriv`) is `C^∞`. -/
1964-
theorem contDiffOn_top_iff_deriv_of_open (hs : IsOpen s₂) :
1964+
theorem contDiffOn_top_iff_deriv_of_isOpen (hs : IsOpen s₂) :
19651965
ContDiffOn 𝕜 ∞ f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 ∞ (deriv f₂) s₂ := by
19661966
rw [contDiffOn_top_iff_derivWithin hs.uniqueDiffOn]
1967-
exact Iff.rfl.and <| contDiffOn_congr fun _ => derivWithin_of_open hs
1968-
#align cont_diff_on_top_iff_deriv_of_open contDiffOn_top_iff_deriv_of_open
1967+
exact Iff.rfl.and <| contDiffOn_congr fun _ => derivWithin_of_isOpen hs
1968+
#align cont_diff_on_top_iff_deriv_of_open contDiffOn_top_iff_deriv_of_isOpen
19691969

19701970
protected theorem ContDiffOn.derivWithin (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : UniqueDiffOn 𝕜 s₂)
19711971
(hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (derivWithin f₂ s₂) s₂ := by
@@ -1978,26 +1978,26 @@ protected theorem ContDiffOn.derivWithin (hf : ContDiffOn 𝕜 n f₂ s₂) (hs
19781978
exact ((contDiffOn_succ_iff_derivWithin hs).1 (hf.of_le hmn)).2
19791979
#align cont_diff_on.deriv_within ContDiffOn.derivWithin
19801980

1981-
theorem ContDiffOn.deriv_of_open (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) (hmn : m + 1 ≤ n) :
1981+
theorem ContDiffOn.deriv_of_isOpen (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) (hmn : m + 1 ≤ n) :
19821982
ContDiffOn 𝕜 m (deriv f₂) s₂ :=
1983-
(hf.derivWithin hs.uniqueDiffOn hmn).congr fun _ hx => (derivWithin_of_open hs hx).symm
1984-
#align cont_diff_on.deriv_of_open ContDiffOn.deriv_of_open
1983+
(hf.derivWithin hs.uniqueDiffOn hmn).congr fun _ hx => (derivWithin_of_isOpen hs hx).symm
1984+
#align cont_diff_on.deriv_of_open ContDiffOn.deriv_of_isOpen
19851985

19861986
theorem ContDiffOn.continuousOn_derivWithin (h : ContDiffOn 𝕜 n f₂ s₂) (hs : UniqueDiffOn 𝕜 s₂)
19871987
(hn : 1 ≤ n) : ContinuousOn (derivWithin f₂ s₂) s₂ :=
19881988
((contDiffOn_succ_iff_derivWithin hs).1 (h.of_le hn)).2.continuousOn
19891989
#align cont_diff_on.continuous_on_deriv_within ContDiffOn.continuousOn_derivWithin
19901990

1991-
theorem ContDiffOn.continuousOn_deriv_of_open (h : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂)
1991+
theorem ContDiffOn.continuousOn_deriv_of_isOpen (h : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂)
19921992
(hn : 1 ≤ n) : ContinuousOn (deriv f₂) s₂ :=
1993-
((contDiffOn_succ_iff_deriv_of_open hs).1 (h.of_le hn)).2.continuousOn
1994-
#align cont_diff_on.continuous_on_deriv_of_open ContDiffOn.continuousOn_deriv_of_open
1993+
((contDiffOn_succ_iff_deriv_of_isOpen hs).1 (h.of_le hn)).2.continuousOn
1994+
#align cont_diff_on.continuous_on_deriv_of_open ContDiffOn.continuousOn_deriv_of_isOpen
19951995

19961996
/-- A function is `C^(n + 1)` if and only if it is differentiable,
19971997
and its derivative (formulated in terms of `deriv`) is `C^n`. -/
19981998
theorem contDiff_succ_iff_deriv {n : ℕ} :
19991999
ContDiff 𝕜 (n + 1 : ℕ) f₂ ↔ Differentiable 𝕜 f₂ ∧ ContDiff 𝕜 n (deriv f₂) := by
2000-
simp only [← contDiffOn_univ, contDiffOn_succ_iff_deriv_of_open, isOpen_univ,
2000+
simp only [← contDiffOn_univ, contDiffOn_succ_iff_deriv_of_isOpen, isOpen_univ,
20012001
differentiableOn_univ]
20022002
#align cont_diff_succ_iff_deriv contDiff_succ_iff_deriv
20032003

Mathlib/Analysis/Calculus/ContDiff/Defs.lean

Lines changed: 9 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -1156,12 +1156,12 @@ theorem contDiffOn_succ_iff_has_fderiv_within {n : ℕ} (hs : UniqueDiffOn 𝕜
11561156

11571157
/-- A function is `C^(n + 1)` on an open domain if and only if it is
11581158
differentiable there, and its derivative (expressed with `fderiv`) is `C^n`. -/
1159-
theorem contDiffOn_succ_iff_fderiv_of_open {n : ℕ} (hs : IsOpen s) :
1159+
theorem contDiffOn_succ_iff_fderiv_of_isOpen {n : ℕ} (hs : IsOpen s) :
11601160
ContDiffOn 𝕜 (n + 1 : ℕ) f s ↔
11611161
DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 n (fun y => fderiv 𝕜 f y) s := by
11621162
rw [contDiffOn_succ_iff_fderivWithin hs.uniqueDiffOn]
11631163
exact Iff.rfl.and (contDiffOn_congr fun x hx ↦ fderivWithin_of_isOpen hs hx)
1164-
#align cont_diff_on_succ_iff_fderiv_of_open contDiffOn_succ_iff_fderiv_of_open
1164+
#align cont_diff_on_succ_iff_fderiv_of_open contDiffOn_succ_iff_fderiv_of_isOpen
11651165

11661166
/-- A function is `C^∞` on a domain with unique derivatives if and only if it is differentiable
11671167
there, and its derivative (expressed with `fderivWithin`) is `C^∞`. -/
@@ -1182,11 +1182,11 @@ theorem contDiffOn_top_iff_fderivWithin (hs : UniqueDiffOn 𝕜 s) :
11821182

11831183
/-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its
11841184
derivative (expressed with `fderiv`) is `C^∞`. -/
1185-
theorem contDiffOn_top_iff_fderiv_of_open (hs : IsOpen s) :
1185+
theorem contDiffOn_top_iff_fderiv_of_isOpen (hs : IsOpen s) :
11861186
ContDiffOn 𝕜 ∞ f s ↔ DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 ∞ (fun y => fderiv 𝕜 f y) s := by
11871187
rw [contDiffOn_top_iff_fderivWithin hs.uniqueDiffOn]
11881188
exact Iff.rfl.and <| contDiffOn_congr fun x hx ↦ fderivWithin_of_isOpen hs hx
1189-
#align cont_diff_on_top_iff_fderiv_of_open contDiffOn_top_iff_fderiv_of_open
1189+
#align cont_diff_on_top_iff_fderiv_of_open contDiffOn_top_iff_fderiv_of_isOpen
11901190

11911191
protected theorem ContDiffOn.fderivWithin (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s)
11921192
(hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun y => fderivWithin 𝕜 f s y) s := by
@@ -1199,20 +1199,20 @@ protected theorem ContDiffOn.fderivWithin (hf : ContDiffOn 𝕜 n f s) (hs : Uni
11991199
exact ((contDiffOn_succ_iff_fderivWithin hs).1 (hf.of_le hmn)).2
12001200
#align cont_diff_on.fderiv_within ContDiffOn.fderivWithin
12011201

1202-
theorem ContDiffOn.fderiv_of_open (hf : ContDiffOn 𝕜 n f s) (hs : IsOpen s) (hmn : m + 1 ≤ n) :
1202+
theorem ContDiffOn.fderiv_of_isOpen (hf : ContDiffOn 𝕜 n f s) (hs : IsOpen s) (hmn : m + 1 ≤ n) :
12031203
ContDiffOn 𝕜 m (fun y => fderiv 𝕜 f y) s :=
12041204
(hf.fderivWithin hs.uniqueDiffOn hmn).congr fun _ hx => (fderivWithin_of_isOpen hs hx).symm
1205-
#align cont_diff_on.fderiv_of_open ContDiffOn.fderiv_of_open
1205+
#align cont_diff_on.fderiv_of_open ContDiffOn.fderiv_of_isOpen
12061206

12071207
theorem ContDiffOn.continuousOn_fderivWithin (h : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s)
12081208
(hn : 1 ≤ n) : ContinuousOn (fun x => fderivWithin 𝕜 f s x) s :=
12091209
((contDiffOn_succ_iff_fderivWithin hs).1 (h.of_le hn)).2.continuousOn
12101210
#align cont_diff_on.continuous_on_fderiv_within ContDiffOn.continuousOn_fderivWithin
12111211

1212-
theorem ContDiffOn.continuousOn_fderiv_of_open (h : ContDiffOn 𝕜 n f s) (hs : IsOpen s)
1212+
theorem ContDiffOn.continuousOn_fderiv_of_isOpen (h : ContDiffOn 𝕜 n f s) (hs : IsOpen s)
12131213
(hn : 1 ≤ n) : ContinuousOn (fun x => fderiv 𝕜 f x) s :=
1214-
((contDiffOn_succ_iff_fderiv_of_open hs).1 (h.of_le hn)).2.continuousOn
1215-
#align cont_diff_on.continuous_on_fderiv_of_open ContDiffOn.continuousOn_fderiv_of_open
1214+
((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 (h.of_le hn)).2.continuousOn
1215+
#align cont_diff_on.continuous_on_fderiv_of_open ContDiffOn.continuousOn_fderiv_of_isOpen
12161216

12171217
/-! ### Functions with a Taylor series on the whole space -/
12181218

Mathlib/Analysis/Calculus/Deriv/Basic.lean

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -531,9 +531,9 @@ theorem derivWithin_inter (ht : t ∈ 𝓝 x) : derivWithin f (s ∩ t) x = deri
531531
theorem derivWithin_of_mem_nhds (h : s ∈ 𝓝 x) : derivWithin f s x = deriv f x := by
532532
simp only [derivWithin, deriv, fderivWithin_of_mem_nhds h]
533533

534-
theorem derivWithin_of_open (hs : IsOpen s) (hx : x ∈ s) : derivWithin f s x = deriv f x :=
534+
theorem derivWithin_of_isOpen (hs : IsOpen s) (hx : x ∈ s) : derivWithin f s x = deriv f x :=
535535
derivWithin_of_mem_nhds (hs.mem_nhds hx)
536-
#align deriv_within_of_open derivWithin_of_open
536+
#align deriv_within_of_open derivWithin_of_isOpen
537537

538538
theorem deriv_mem_iff {f : 𝕜 → F} {s : Set F} {x : 𝕜} :
539539
deriv f x ∈ s ↔

Mathlib/Analysis/Calculus/LineDeriv/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -278,7 +278,7 @@ theorem lineDerivWithin_of_mem_nhds (h : s ∈ 𝓝 x) :
278278
apply (Continuous.continuousAt _).preimage_mem_nhds (by simpa using h)
279279
continuity
280280

281-
theorem lineDerivWithin_of_open (hs : IsOpen s) (hx : x ∈ s) :
281+
theorem lineDerivWithin_of_isOpen (hs : IsOpen s) (hx : x ∈ s) :
282282
lineDerivWithin 𝕜 f s x v = lineDeriv 𝕜 f x v :=
283283
lineDerivWithin_of_mem_nhds (hs.mem_nhds hx)
284284

Mathlib/Analysis/Complex/AbsMax.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -240,7 +240,7 @@ theorem norm_eqOn_of_isPreconnected_of_isMaxOn {f : E → F} {U : Set E} {c : E}
240240
using isOpen_setOf_mem_nhds_and_isMaxOn_norm hd
241241
have hVne : (U ∩ V).Nonempty := ⟨c, hcU, hcU, hm⟩
242242
set W := U ∩ {z | ‖f z‖ ≠ ‖f c‖}
243-
have hWo : IsOpen W := hd.continuousOn.norm.preimage_open_of_open ho isOpen_ne
243+
have hWo : IsOpen W := hd.continuousOn.norm.isOpen_inter_preimage ho isOpen_ne
244244
have hdVW : Disjoint V W := disjoint_left.mpr fun x hxV hxW => hxW.2 (hV x hxV)
245245
have hUVW : U ⊆ V ∪ W := fun x hx =>
246246
(eq_or_ne ‖f x‖ ‖f c‖).imp (fun h => ⟨hx, fun y hy => (hm hy).out.trans_eq h.symm⟩)

Mathlib/Analysis/Convex/Gauge.lean

Lines changed: 9 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -359,24 +359,24 @@ theorem interior_subset_gauge_lt_one (s : Set E) : interior s ⊆ { x | gauge s
359359
exact (gauge_le_of_mem hr₀.le hxr).trans_lt hr₁
360360
#align interior_subset_gauge_lt_one interior_subset_gauge_lt_one
361361

362-
theorem gauge_lt_one_eq_self_of_open (hs₁ : Convex ℝ s) (hs₀ : (0 : E) ∈ s) (hs₂ : IsOpen s) :
362+
theorem gauge_lt_one_eq_self_of_isOpen (hs₁ : Convex ℝ s) (hs₀ : (0 : E) ∈ s) (hs₂ : IsOpen s) :
363363
{ x | gauge s x < 1 } = s := by
364364
refine' (gauge_lt_one_subset_self hs₁ ‹_› <| absorbent_nhds_zero <| hs₂.mem_nhds hs₀).antisymm _
365365
convert interior_subset_gauge_lt_one s
366366
exact hs₂.interior_eq.symm
367-
#align gauge_lt_one_eq_self_of_open gauge_lt_one_eq_self_of_open
367+
#align gauge_lt_one_eq_self_of_open gauge_lt_one_eq_self_of_isOpen
368368

369369
-- porting note: droped unneeded assumptions
370-
theorem gauge_lt_one_of_mem_of_open (hs₂ : IsOpen s) {x : E} (hx : x ∈ s) :
370+
theorem gauge_lt_one_of_mem_of_isOpen (hs₂ : IsOpen s) {x : E} (hx : x ∈ s) :
371371
gauge s x < 1 :=
372372
interior_subset_gauge_lt_one s <| by rwa [hs₂.interior_eq]
373-
#align gauge_lt_one_of_mem_of_open gauge_lt_one_of_mem_of_openₓ
373+
#align gauge_lt_one_of_mem_of_open gauge_lt_one_of_mem_of_isOpenₓ
374374

375375
-- porting note: droped unneeded assumptions
376376
theorem gauge_lt_of_mem_smul (x : E) (ε : ℝ) (hε : 0 < ε) (hs₂ : IsOpen s) (hx : x ∈ ε • s) :
377377
gauge s x < ε := by
378378
have : ε⁻¹ • x ∈ s := by rwa [← mem_smul_set_iff_inv_smul_mem₀ hε.ne']
379-
have h_gauge_lt := gauge_lt_one_of_mem_of_open hs₂ this
379+
have h_gauge_lt := gauge_lt_one_of_mem_of_isOpen hs₂ this
380380
rwa [gauge_smul_of_nonneg (inv_nonneg.2 hε.le), smul_eq_mul, inv_mul_lt_iff hε, mul_one]
381381
at h_gauge_lt
382382
#align gauge_lt_of_mem_smul gauge_lt_of_mem_smulₓ
@@ -480,14 +480,14 @@ def gaugeSeminorm (hs₀ : Balanced 𝕜 s) (hs₁ : Convex ℝ s) (hs₂ : Abso
480480
variable {hs₀ : Balanced 𝕜 s} {hs₁ : Convex ℝ s} {hs₂ : Absorbent ℝ s} [TopologicalSpace E]
481481
[ContinuousSMul ℝ E]
482482

483-
theorem gaugeSeminorm_lt_one_of_open (hs : IsOpen s) {x : E} (hx : x ∈ s) :
483+
theorem gaugeSeminorm_lt_one_of_isOpen (hs : IsOpen s) {x : E} (hx : x ∈ s) :
484484
gaugeSeminorm hs₀ hs₁ hs₂ x < 1 :=
485-
gauge_lt_one_of_mem_of_open hs hx
486-
#align gauge_seminorm_lt_one_of_open gaugeSeminorm_lt_one_of_open
485+
gauge_lt_one_of_mem_of_isOpen hs hx
486+
#align gauge_seminorm_lt_one_of_open gaugeSeminorm_lt_one_of_isOpen
487487

488488
theorem gaugeSeminorm_ball_one (hs : IsOpen s) : (gaugeSeminorm hs₀ hs₁ hs₂).ball 0 1 = s := by
489489
rw [Seminorm.ball_zero_eq]
490-
exact gauge_lt_one_eq_self_of_open hs₁ hs₂.zero_mem hs
490+
exact gauge_lt_one_eq_self_of_isOpen hs₁ hs₂.zero_mem hs
491491
#align gauge_seminorm_ball_one gaugeSeminorm_ball_one
492492

493493
end IsROrC

Mathlib/Analysis/Convex/Strict.lean

Lines changed: 4 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -109,12 +109,13 @@ protected theorem StrictConvex.convex (hs : StrictConvex 𝕜 s) : Convex 𝕜 s
109109
#align strict_convex.convex StrictConvex.convex
110110

111111
/-- An open convex set is strictly convex. -/
112-
protected theorem Convex.strictConvex_of_open (h : IsOpen s) (hs : Convex 𝕜 s) : StrictConvex 𝕜 s :=
112+
protected theorem Convex.strictConvex_of_isOpen (h : IsOpen s) (hs : Convex 𝕜 s) :
113+
StrictConvex 𝕜 s :=
113114
fun _ hx _ hy _ _ _ ha hb hab => h.interior_eq.symm ▸ hs hx hy ha.le hb.le hab
114-
#align convex.strict_convex_of_open Convex.strictConvex_of_open
115+
#align convex.strict_convex_of_open Convex.strictConvex_of_isOpen
115116

116117
theorem IsOpen.strictConvex_iff (h : IsOpen s) : StrictConvex 𝕜 s ↔ Convex 𝕜 s :=
117-
⟨StrictConvex.convex, Convex.strictConvex_of_open h⟩
118+
⟨StrictConvex.convex, Convex.strictConvex_of_isOpen h⟩
118119
#align is_open.strict_convex_iff IsOpen.strictConvex_iff
119120

120121
theorem strictConvex_singleton (c : E) : StrictConvex 𝕜 ({c} : Set E) :=

0 commit comments

Comments
 (0)