Skip to content

Commit

Permalink
feat(topology/algebra/algebra): algebra_clm does not need a normed …
Browse files Browse the repository at this point in the history
…space or field (#16690)

This also clean up the variables in the `topology/algebra/algebra.lean` file, to ensure that the type variables come first (as this is a useful convention for when using `@` notation).
  • Loading branch information
eric-wieser committed Sep 29, 2022
1 parent 37cfab4 commit 2e09d79
Show file tree
Hide file tree
Showing 4 changed files with 23 additions and 28 deletions.
19 changes: 0 additions & 19 deletions src/analysis/normed_space/basic.lean
Original file line number Diff line number Diff line change
Expand Up @@ -455,25 +455,6 @@ begin
rw [dist_eq_norm, dist_eq_norm, ← ring_hom.map_sub, norm_algebra_map'],
end

/-- The inclusion of the base field in a normed algebra as a continuous linear map. -/
@[simps]
def algebra_map_clm : 𝕜 →L[𝕜] 𝕜' :=
{ to_fun := algebra_map 𝕜 𝕜',
map_add' := (algebra_map 𝕜 𝕜').map_add,
map_smul' := λ r x, by rw [algebra.id.smul_eq_mul, map_mul, ring_hom.id_apply, algebra.smul_def],
cont :=
have lipschitz_with ∥(1 : 𝕜')∥₊ (algebra_map 𝕜 𝕜') := λ x y, begin
rw [edist_eq_coe_nnnorm_sub, edist_eq_coe_nnnorm_sub, ←map_sub, ←ennreal.coe_mul,
ennreal.coe_le_coe, mul_comm],
exact (nnnorm_algebra_map _ _).le,
end, this.continuous }

lemma algebra_map_clm_coe :
(algebra_map_clm 𝕜 𝕜' : 𝕜 → 𝕜') = (algebra_map 𝕜 𝕜' : 𝕜 → 𝕜') := rfl

lemma algebra_map_clm_to_linear_map :
(algebra_map_clm 𝕜 𝕜').to_linear_map = algebra.linear_map 𝕜 𝕜' := rfl

instance normed_algebra.id : normed_algebra 𝕜 𝕜 :=
{ .. normed_field.to_normed_space,
.. algebra.id 𝕜}
Expand Down
3 changes: 2 additions & 1 deletion src/analysis/normed_space/exponential.lean
Original file line number Diff line number Diff line change
Expand Up @@ -8,6 +8,7 @@ import analysis.analytic.basic
import analysis.complex.basic
import data.nat.choose.cast
import data.finset.noncomm_prod
import topology.algebra.algebra

/-!
# Exponential in a Banach algebra
Expand Down Expand Up @@ -289,7 +290,7 @@ end complete_algebra
lemma algebra_map_exp_comm_of_mem_ball [complete_space 𝕂] (x : 𝕂)
(hx : x ∈ emetric.ball (0 : 𝕂) (exp_series 𝕂 𝕂).radius) :
algebra_map 𝕂 𝔸 (exp 𝕂 x) = exp 𝕂 (algebra_map 𝕂 𝔸 x) :=
map_exp_of_mem_ball _ (algebra_map_clm _ _).continuous _ hx
map_exp_of_mem_ball _ (continuous_algebra_map 𝕂 𝔸) _ hx

end any_field_any_algebra

Expand Down
2 changes: 1 addition & 1 deletion src/analysis/normed_space/spectrum.lean
Original file line number Diff line number Diff line change
Expand Up @@ -71,7 +71,7 @@ not_not.mp $ λ hn, h.not_le $ le_supr₂ k hn
variable [complete_space A]

lemma is_open_resolvent_set (a : A) : is_open (ρ a) :=
units.is_open.preimage ((algebra_map_clm 𝕜 A).continuous.sub continuous_const)
units.is_open.preimage ((continuous_algebra_map 𝕜 A).sub continuous_const)

lemma is_closed (a : A) : is_closed (σ a) :=
(is_open_resolvent_set a).is_closed_compl
Expand Down
27 changes: 20 additions & 7 deletions src/topology/algebra/algebra.lean
Original file line number Diff line number Diff line change
Expand Up @@ -28,11 +28,11 @@ open_locale classical
universes u v w

section topological_algebra
variables (R : Type*) [topological_space R] [comm_semiring R]
variables (A : Type u) [topological_space A]
variables [semiring A]
variables (R : Type*) (A : Type u)
variables [comm_semiring R] [semiring A] [algebra R A]
variables [topological_space R] [topological_space A] [topological_semiring A]

lemma continuous_algebra_map_iff_smul [algebra R A] [topological_semiring A] :
lemma continuous_algebra_map_iff_smul :
continuous (algebra_map R A) ↔ continuous (λ p : R × A, p.1 • p.2) :=
begin
refine ⟨λ h, _, λ h, _⟩,
Expand All @@ -41,15 +41,28 @@ begin
end

@[continuity]
lemma continuous_algebra_map [algebra R A] [topological_semiring A] [has_continuous_smul R A] :
lemma continuous_algebra_map [has_continuous_smul R A] :
continuous (algebra_map R A) :=
(continuous_algebra_map_iff_smul R A).2 continuous_smul

lemma has_continuous_smul_of_algebra_map [algebra R A] [topological_semiring A]
(h : continuous (algebra_map R A)) :
lemma has_continuous_smul_of_algebra_map (h : continuous (algebra_map R A)) :
has_continuous_smul R A :=
⟨(continuous_algebra_map_iff_smul R A).1 h⟩

variables [has_continuous_smul R A]

/-- The inclusion of the base ring in a topological algebra as a continuous linear map. -/
@[simps]
def algebra_map_clm : R →L[R] A :=
{ to_fun := algebra_map R A,
cont := continuous_algebra_map R A,
.. algebra.linear_map R A }

lemma algebra_map_clm_coe : ⇑(algebra_map_clm R A) = algebra_map R A := rfl

lemma algebra_map_clm_to_linear_map :
(algebra_map_clm R A).to_linear_map = algebra.linear_map R A := rfl

end topological_algebra

section topological_algebra
Expand Down

0 comments on commit 2e09d79

Please sign in to comment.