Skip to content

Commit

Permalink
chore: adjust style and docstring of ContinuousMap.compRightAlgHom (#…
Browse files Browse the repository at this point in the history
  • Loading branch information
j-loreaux committed Feb 20, 2024
1 parent c9eb34f commit bb76bc8
Showing 1 changed file with 7 additions and 17 deletions.
24 changes: 7 additions & 17 deletions Mathlib/Topology/ContinuousFunction/Algebra.lean
Expand Up @@ -749,27 +749,17 @@ protected def AlgHom.compLeftContinuous {α : Type*} [TopologicalSpace α] (g :

variable (A)

/-- Precomposition of functions into a normed ring by a continuous map is an algebra homomorphism.
-/
/-- Precomposition of functions into a topological semiring by a continuous map is an algebra
homomorphism. -/
@[simps]
def ContinuousMap.compRightAlgHom {α β : Type*} [TopologicalSpace α] [TopologicalSpace β]
(f : C(α, β)) : C(β, A) →ₐ[R] C(α, A) where
toFun g := g.comp f
map_zero' := by
ext
rfl
map_add' g₁ g₂ := by
ext
rfl
map_one' := by
ext
rfl
map_mul' g₁ g₂ := by
ext
rfl
commutes' r := by
ext
rfl
map_zero' := ext fun _ ↦ rfl
map_add' _ _ := ext fun _ ↦ rfl
map_one' := ext fun _ ↦ rfl
map_mul' _ _ := ext fun _ ↦ rfl
commutes' _ := ext fun _ ↦ rfl
#align continuous_map.comp_right_alg_hom ContinuousMap.compRightAlgHom

variable {A}
Expand Down

0 comments on commit bb76bc8

Please sign in to comment.