Skip to content

Commit

Permalink
chore(analysis/locally_convex/with_seminorms): fix namespaces for dot…
Browse files Browse the repository at this point in the history
…-notation (#16771)

Some lemmas where in the wrong namespaces.
  • Loading branch information
mcdoll committed Oct 3, 2022
1 parent ebe6bda commit d246177
Show file tree
Hide file tree
Showing 3 changed files with 10 additions and 11 deletions.
2 changes: 1 addition & 1 deletion src/analysis/locally_convex/weak_dual.lean
Expand Up @@ -135,6 +135,6 @@ variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [add_comm_group
variables [nonempty ι] [normed_space ℝ 𝕜] [module ℝ E] [is_scalar_tower ℝ 𝕜 E]

instance {B : E →ₗ[𝕜] F →ₗ[𝕜] 𝕜} : locally_convex_space ℝ (weak_bilin B) :=
seminorm_family.to_locally_convex_space (B.weak_bilin_with_seminorms)
(B.weak_bilin_with_seminorms).to_locally_convex_space

end locally_convex
6 changes: 3 additions & 3 deletions src/analysis/locally_convex/with_seminorms.lean
Expand Up @@ -271,7 +271,7 @@ variables [normed_field 𝕜] [add_comm_group E] [module 𝕜 E] [nonempty ι]
structure with_seminorms (p : seminorm_family 𝕜 E ι) [t : topological_space E] : Prop :=
(topology_eq_with_seminorms : t = p.module_filter_basis.topology)

lemma seminorm_family.with_seminorms_eq {p : seminorm_family 𝕜 E ι} [t : topological_space E]
lemma with_seminorms.with_seminorms_eq {p : seminorm_family 𝕜 E ι} [t : topological_space E]
(hp : with_seminorms p) : t = p.module_filter_basis.topology := hp.1

variables [topological_space E]
Expand Down Expand Up @@ -493,7 +493,7 @@ variables [nonempty ι] [normed_field 𝕜] [normed_space ℝ 𝕜]
[add_comm_group E] [module 𝕜 E] [module ℝ E] [is_scalar_tower ℝ 𝕜 E] [topological_space E]
[topological_add_group E]

lemma seminorm_family.to_locally_convex_space {p : seminorm_family 𝕜 E ι} (hp : with_seminorms p) :
lemma with_seminorms.to_locally_convex_space {p : seminorm_family 𝕜 E ι} (hp : with_seminorms p) :
locally_convex_space ℝ E :=
begin
apply of_basis_zero ℝ E id (λ s, s ∈ p.basis_sets),
Expand All @@ -516,7 +516,7 @@ variables (𝕜) [normed_field 𝕜] [normed_space ℝ 𝕜] [seminormed_add_com
slightly weaker instance version. -/
lemma normed_space.to_locally_convex_space' [normed_space 𝕜 E] [module ℝ E]
[is_scalar_tower ℝ 𝕜 E] : locally_convex_space ℝ E :=
seminorm_family.to_locally_convex_space (norm_with_seminorms 𝕜 E)
(norm_with_seminorms 𝕜 E).to_locally_convex_space

/-- See `normed_space.to_locally_convex_space'` for a slightly stronger version which is not an
instance. -/
Expand Down
13 changes: 6 additions & 7 deletions src/analysis/schwartz_space.lean
Expand Up @@ -22,7 +22,7 @@ natural numbers `k` and `n` we have uniform bounds `∥x∥^k * ∥iterated_fder
This approach completely avoids using partial derivatives as well as polynomials.
We construct the topology on the Schwartz space by a family of seminorms, which are the best
constants in the above estimates, which is by abstract theory from
`seminorm_family.module_filter_basis` and `seminorm_family.to_locally_convex_space` turns the
`seminorm_family.module_filter_basis` and `with_seminorms.to_locally_convex_space` turns the
Schwartz space into a locally convex topological vector space.
## Main definitions
Expand Down Expand Up @@ -375,22 +375,21 @@ variables {𝕜 E F}

instance : has_continuous_smul 𝕜 𝓢(E, F) :=
begin
rw seminorm_family.with_seminorms_eq (schwartz_with_seminorms 𝕜 E F),
rw (schwartz_with_seminorms 𝕜 E F).with_seminorms_eq,
exact (schwartz_seminorm_family 𝕜 E F).module_filter_basis.has_continuous_smul,
end

instance : topological_add_group 𝓢(E, F) :=
(schwartz_seminorm_family ℝ E F).module_filter_basis.to_add_group_filter_basis
.is_topological_add_group
(schwartz_seminorm_family ℝ E F).add_group_filter_basis.is_topological_add_group

instance : uniform_space 𝓢(E, F) :=
(schwartz_seminorm_family ℝ E F).module_filter_basis.to_add_group_filter_basis.uniform_space
(schwartz_seminorm_family ℝ E F).add_group_filter_basis.uniform_space

instance : uniform_add_group 𝓢(E, F) :=
(schwartz_seminorm_family ℝ E F).module_filter_basis.to_add_group_filter_basis.uniform_add_group
(schwartz_seminorm_family ℝ E F).add_group_filter_basis.uniform_add_group

instance : locally_convex_space ℝ 𝓢(E, F) :=
seminorm_family.to_locally_convex_space (schwartz_with_seminorms ℝ E F)
(schwartz_with_seminorms ℝ E F).to_locally_convex_space

end topology

Expand Down

0 comments on commit d246177

Please sign in to comment.