Skip to content

Commit 1bbe6b2

Browse files
committed
feat(Analysis.Seminorm): some results about the order on seminorms (#5807)
1 parent 0b6307b commit 1bbe6b2

File tree

1 file changed

+26
-0
lines changed

1 file changed

+26
-0
lines changed

Mathlib/Analysis/Seminorm.lean

Lines changed: 26 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -600,6 +600,10 @@ protected theorem bddAbove_iff {s : Set <| Seminorm 𝕜 E} :
600600
le_ciSup ⟨q x, forall_range_iff.mpr fun i : s => hq (mem_image_of_mem _ i.2) x⟩ ⟨p, hp⟩⟩⟩
601601
#align seminorm.bdd_above_iff Seminorm.bddAbove_iff
602602

603+
protected theorem bddAbove_range_iff {p : ι → Seminorm 𝕜 E} :
604+
BddAbove (range p) ↔ ∀ x, BddAbove (range fun i ↦ p i x) := by
605+
rw [Seminorm.bddAbove_iff, ← range_comp, bddAbove_range_pi]; rfl
606+
603607
protected theorem coe_sSup_eq {s : Set <| Seminorm 𝕜 E} (hs : BddAbove s) :
604608
↑(sSup s) = ⨆ p : s, ((p : Seminorm 𝕜 E) : E → ℝ) :=
605609
Seminorm.coe_sSup_eq' (Seminorm.bddAbove_iff.mp hs)
@@ -611,6 +615,19 @@ protected theorem coe_iSup_eq {ι : Type _} {p : ι → Seminorm 𝕜 E} (hp : B
611615
exact iSup_range' (fun p : Seminorm 𝕜 E => (p : E → ℝ)) p
612616
#align seminorm.coe_supr_eq Seminorm.coe_iSup_eq
613617

618+
protected theorem sSup_apply {s : Set (Seminorm 𝕜 E)} (hp : BddAbove s) {x : E} :
619+
(sSup s) x = ⨆ p : s, (p : E → ℝ) x := by
620+
rw [Seminorm.coe_sSup_eq hp, iSup_apply]
621+
622+
protected theorem iSup_apply {ι : Type _} {p : ι → Seminorm 𝕜 E}
623+
(hp : BddAbove (range p)) {x : E} : (⨆ i, p i) x = ⨆ i, p i x := by
624+
rw [Seminorm.coe_iSup_eq hp, iSup_apply]
625+
626+
protected theorem sSup_empty : sSup (∅ : Set (Seminorm 𝕜 E)) = ⊥ := by
627+
ext
628+
rw [Seminorm.sSup_apply bddAbove_empty, Real.ciSup_empty]
629+
rfl
630+
614631
private theorem Seminorm.isLUB_sSup (s : Set (Seminorm 𝕜 E)) (hs₁ : BddAbove s) (hs₂ : s.Nonempty) :
615632
IsLUB s (sSup s) := by
616633
refine' ⟨fun p hp x => _, fun p hp x => _⟩ <;> haveI : Nonempty ↑s := hs₂.coe_sort <;>
@@ -961,6 +978,15 @@ section NormedField
961978
variable [NormedField 𝕜] [AddCommGroup E] [Module 𝕜 E] (p : Seminorm 𝕜 E) {A B : Set E} {a : 𝕜}
962979
{r : ℝ} {x : E}
963980

981+
theorem closedBall_iSup {p : ι → Seminorm 𝕜 E} (hp : BddAbove (range p)) (e : E) {r : ℝ}
982+
(hr : 0 < r) : closedBall (⨆ i, p i) e r = ⋂ i, closedBall (p i) e r := by
983+
cases isEmpty_or_nonempty ι
984+
· rw [iSup_of_empty', iInter_of_empty, Seminorm.sSup_empty]
985+
exact closedBall_bot _ hr
986+
· ext x
987+
have := Seminorm.bddAbove_range_iff.mp hp (x - e)
988+
simp only [mem_closedBall, mem_iInter, Seminorm.iSup_apply hp, ciSup_le_iff this]
989+
964990
theorem ball_norm_mul_subset {p : Seminorm 𝕜 E} {k : 𝕜} {r : ℝ} :
965991
p.ball 0 (‖k‖ * r) ⊆ k • p.ball 0 r := by
966992
rcases eq_or_ne k 0 with (rfl | hk)

0 commit comments

Comments
 (0)