@@ -10,6 +10,7 @@ import topology.uniform_space.compact_convergence
1010import topology.algebra.star
1111import algebra.algebra.subalgebra.basic
1212import tactic.field_simp
13+ import algebra.star.star_alg_hom
1314
1415/-!
1516# Algebraic structures over continuous functions
@@ -835,4 +836,34 @@ instance [has_star R] [has_star β] [has_smul R β] [star_module R β]
835836
836837end star_structure
837838
839+ variables {X Y Z : Type *} [topological_space X] [topological_space Y] [topological_space Z]
840+ variables (𝕜 : Type *) [comm_semiring 𝕜]
841+ variables (A : Type *) [topological_space A] [semiring A] [topological_semiring A] [star_ring A]
842+ variables [has_continuous_star A] [algebra 𝕜 A]
843+
844+ /-- The functorial map taking `f : C(X, Y)` to `C(Y, A) →⋆ₐ[𝕜] C(X, A)` given by pre-composition
845+ with the continuous function `f`. See `continuous_map.comp_monoid_hom'` and
846+ `continuous_map.comp_add_monoid_hom'`, `continuous_map.comp_right_alg_hom` for bundlings of
847+ pre-composition into a `monoid_hom`, an `add_monoid_hom` and an `alg_hom`, respectively, under
848+ suitable assumptions on `A`. -/
849+ @[simps] def comp_star_alg_hom' (f : C(X, Y)) : C(Y, A) →⋆ₐ[𝕜] C(X, A) :=
850+ { to_fun := λ g, g.comp f,
851+ map_one' := one_comp _,
852+ map_mul' := λ _ _, rfl,
853+ map_zero' := zero_comp _,
854+ map_add' := λ _ _, rfl,
855+ commutes' := λ _, rfl,
856+ map_star' := λ _, rfl }
857+
858+ /-- `continuous_map.comp_star_alg_hom'` sends the identity continuous map to the identity
859+ `star_alg_hom` -/
860+ lemma comp_star_alg_hom'_id :
861+ comp_star_alg_hom' 𝕜 A (continuous_map.id X) = star_alg_hom.id 𝕜 C(X, A) :=
862+ star_alg_hom.ext $ λ _, continuous_map.ext $ λ _, rfl
863+
864+ /-- `continuous_map.comp_star_alg_hom` is functorial. -/
865+ lemma comp_star_alg_hom'_comp (g : C(Y, Z)) (f : C(X, Y)) :
866+ comp_star_alg_hom' 𝕜 A (g.comp f) = (comp_star_alg_hom' 𝕜 A f).comp (comp_star_alg_hom' 𝕜 A g) :=
867+ star_alg_hom.ext $ λ _, continuous_map.ext $ λ _, rfl
868+
838869end continuous_map
0 commit comments