Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit 2e09d79

Browse files
committed
feat(topology/algebra/algebra): algebra_clm does not need a normed 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).
1 parent 37cfab4 commit 2e09d79

File tree

4 files changed

+23
-28
lines changed

4 files changed

+23
-28
lines changed

src/analysis/normed_space/basic.lean

Lines changed: 0 additions & 19 deletions
Original file line numberDiff line numberDiff line change
@@ -455,25 +455,6 @@ begin
455455
rw [dist_eq_norm, dist_eq_norm, ← ring_hom.map_sub, norm_algebra_map'],
456456
end
457457

458-
/-- The inclusion of the base field in a normed algebra as a continuous linear map. -/
459-
@[simps]
460-
def algebra_map_clm : 𝕜 →L[𝕜] 𝕜' :=
461-
{ to_fun := algebra_map 𝕜 𝕜',
462-
map_add' := (algebra_map 𝕜 𝕜').map_add,
463-
map_smul' := λ r x, by rw [algebra.id.smul_eq_mul, map_mul, ring_hom.id_apply, algebra.smul_def],
464-
cont :=
465-
have lipschitz_with ∥(1 : 𝕜')∥₊ (algebra_map 𝕜 𝕜') := λ x y, begin
466-
rw [edist_eq_coe_nnnorm_sub, edist_eq_coe_nnnorm_sub, ←map_sub, ←ennreal.coe_mul,
467-
ennreal.coe_le_coe, mul_comm],
468-
exact (nnnorm_algebra_map _ _).le,
469-
end, this.continuous }
470-
471-
lemma algebra_map_clm_coe :
472-
(algebra_map_clm 𝕜 𝕜' : 𝕜 → 𝕜') = (algebra_map 𝕜 𝕜' : 𝕜 → 𝕜') := rfl
473-
474-
lemma algebra_map_clm_to_linear_map :
475-
(algebra_map_clm 𝕜 𝕜').to_linear_map = algebra.linear_map 𝕜 𝕜' := rfl
476-
477458
instance normed_algebra.id : normed_algebra 𝕜 𝕜 :=
478459
{ .. normed_field.to_normed_space,
479460
.. algebra.id 𝕜}

src/analysis/normed_space/exponential.lean

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -8,6 +8,7 @@ import analysis.analytic.basic
88
import analysis.complex.basic
99
import data.nat.choose.cast
1010
import data.finset.noncomm_prod
11+
import topology.algebra.algebra
1112

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

294295
end any_field_any_algebra
295296

src/analysis/normed_space/spectrum.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -71,7 +71,7 @@ not_not.mp $ λ hn, h.not_le $ le_supr₂ k hn
7171
variable [complete_space A]
7272

7373
lemma is_open_resolvent_set (a : A) : is_open (ρ a) :=
74-
units.is_open.preimage ((algebra_map_clm 𝕜 A).continuous.sub continuous_const)
74+
units.is_open.preimage ((continuous_algebra_map 𝕜 A).sub continuous_const)
7575

7676
lemma is_closed (a : A) : is_closed (σ a) :=
7777
(is_open_resolvent_set a).is_closed_compl

src/topology/algebra/algebra.lean

Lines changed: 20 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -28,11 +28,11 @@ open_locale classical
2828
universes u v w
2929

3030
section topological_algebra
31-
variables (R : Type*) [topological_space R] [comm_semiring R]
32-
variables (A : Type u) [topological_space A]
33-
variables [semiring A]
31+
variables (R : Type*) (A : Type u)
32+
variables [comm_semiring R] [semiring A] [algebra R A]
33+
variables [topological_space R] [topological_space A] [topological_semiring A]
3434

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

4343
@[continuity]
44-
lemma continuous_algebra_map [algebra R A] [topological_semiring A] [has_continuous_smul R A] :
44+
lemma continuous_algebra_map [has_continuous_smul R A] :
4545
continuous (algebra_map R A) :=
4646
(continuous_algebra_map_iff_smul R A).2 continuous_smul
4747

48-
lemma has_continuous_smul_of_algebra_map [algebra R A] [topological_semiring A]
49-
(h : continuous (algebra_map R A)) :
48+
lemma has_continuous_smul_of_algebra_map (h : continuous (algebra_map R A)) :
5049
has_continuous_smul R A :=
5150
⟨(continuous_algebra_map_iff_smul R A).1 h⟩
5251

52+
variables [has_continuous_smul R A]
53+
54+
/-- The inclusion of the base ring in a topological algebra as a continuous linear map. -/
55+
@[simps]
56+
def algebra_map_clm : R →L[R] A :=
57+
{ to_fun := algebra_map R A,
58+
cont := continuous_algebra_map R A,
59+
.. algebra.linear_map R A }
60+
61+
lemma algebra_map_clm_coe : ⇑(algebra_map_clm R A) = algebra_map R A := rfl
62+
63+
lemma algebra_map_clm_to_linear_map :
64+
(algebra_map_clm R A).to_linear_map = algebra.linear_map R A := rfl
65+
5366
end topological_algebra
5467

5568
section topological_algebra

0 commit comments

Comments
 (0)