Skip to content

Commit

Permalink
feat(analysis/seminorm): Add has_add and has_scalar nnreal (#11414)
Browse files Browse the repository at this point in the history
Add instances of `has_add` and `has_scalar nnreal` type classes for `seminorm`.

Co-authored-by: Eric Wieser <wieser.eric@gmail.com>



Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
  • Loading branch information
mcdoll and eric-wieser committed Jan 16, 2022
1 parent 3fb5b82 commit d60541c
Showing 1 changed file with 67 additions and 2 deletions.
69 changes: 67 additions & 2 deletions src/analysis/seminorm.lean
Expand Up @@ -66,9 +66,9 @@ Absorbent and balanced sets in a vector space over a normed field.
-/

open normed_field set
open_locale pointwise topological_space
open_locale pointwise topological_space nnreal

variables {𝕜 E ι : Type*}
variables {R 𝕜 E ι : Type*}

section semi_normed_ring
variables [semi_normed_ring 𝕜]
Expand Down Expand Up @@ -278,6 +278,71 @@ variables (p : seminorm 𝕜 E) (c : 𝕜) (x y : E) (r : ℝ)
protected lemma smul : p (c • x) = ∥c∥ * p x := p.smul' _ _
protected lemma triangle : p (x + y) ≤ p x + p y := p.triangle' _ _

/-- Any action on `ℝ` which factors through `ℝ≥0` applies to a seminorm. -/
instance [has_scalar R ℝ] [has_scalar R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] :
has_scalar R (seminorm 𝕜 E) :=
{ smul := λ r p,
{ to_fun := λ x, r • p x,
smul' := λ _ _, begin
simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul],
rw [p.smul, mul_left_comm],
end,
triangle' := λ _ _, begin
simp only [←smul_one_smul ℝ≥0 r (_ : ℝ), nnreal.smul_def, smul_eq_mul],
exact (mul_le_mul_of_nonneg_left (p.triangle _ _) (nnreal.coe_nonneg _)).trans_eq
(mul_add _ _ _),
end } }

lemma coe_smul [has_scalar R ℝ] [has_scalar R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ]
(r : R) (p : seminorm 𝕜 E) : ⇑(r • p) = r • p := rfl

@[simp] lemma smul_apply [has_scalar R ℝ] [has_scalar R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ]
(r : R) (p : seminorm 𝕜 E) (x : E) : (r • p) x = r • p x := rfl

instance : has_add (seminorm 𝕜 E) :=
{ add := λ p q,
{ to_fun := λ x, p x + q x,
smul' := λ a x, by rw [p.smul, q.smul, mul_add],
triangle' := λ _ _, has_le.le.trans_eq (add_le_add (p.triangle _ _) (q.triangle _ _))
(add_add_add_comm _ _ _ _) } }

lemma coe_add (p q : seminorm 𝕜 E) : ⇑(p + q) = p + q := rfl

@[simp] lemma add_apply (p q : seminorm 𝕜 E) (x : E) : (p + q) x = p x + q x := rfl

instance : add_monoid (seminorm 𝕜 E) :=
fun_like.coe_injective.add_monoid_smul _ rfl coe_add (λ p n, coe_smul n p)

instance : ordered_cancel_add_comm_monoid (seminorm 𝕜 E) :=
{ nsmul := (•), -- to avoid introducing a diamond
..seminorm.add_monoid,
..(fun_like.coe_injective.ordered_cancel_add_comm_monoid _ rfl coe_add
: ordered_cancel_add_comm_monoid (seminorm 𝕜 E)) }

instance [monoid R] [mul_action R ℝ] [has_scalar R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] :
mul_action R (seminorm 𝕜 E) :=
fun_like.coe_injective.mul_action _ coe_smul

variables (𝕜 E)

/-- `coe_fn` as an `add_monoid_hom`. Helper definition for showing that `seminorm 𝕜 E` is
a module. -/
@[simps]
def coe_fn_add_monoid_hom : add_monoid_hom (seminorm 𝕜 E) (E → ℝ) := ⟨coe_fn, coe_zero, coe_add⟩

lemma coe_fn_add_monoid_hom_injective : function.injective (coe_fn_add_monoid_hom 𝕜 E) :=
show @function.injective (seminorm 𝕜 E) (E → ℝ) coe_fn, from fun_like.coe_injective

variables {𝕜 E}

instance [monoid R] [distrib_mul_action R ℝ] [has_scalar R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] :
distrib_mul_action R (seminorm 𝕜 E) :=
(coe_fn_add_monoid_hom_injective 𝕜 E).distrib_mul_action _ coe_smul

instance [semiring R] [module R ℝ] [has_scalar R ℝ≥0] [is_scalar_tower R ℝ≥0 ℝ] :
module R (seminorm 𝕜 E) :=
(coe_fn_add_monoid_hom_injective 𝕜 E).module R _ coe_smul

-- TODO: define `has_Sup` too, from the skeleton at
-- https://github.com/leanprover-community/mathlib/pull/11329#issuecomment-1008915345
noncomputable instance : has_sup (seminorm 𝕜 E) :=
Expand Down

0 comments on commit d60541c

Please sign in to comment.