diff --git a/src/analysis/normed_space/exponential.lean b/src/analysis/normed_space/exponential.lean index 1b5d2544146c7..870ab76e23e08 100644 --- a/src/analysis/normed_space/exponential.lean +++ b/src/analysis/normed_space/exponential.lean @@ -125,6 +125,10 @@ begin simp [h] end +lemma star_exp [t2_space 𝔸] [star_ring 𝔸] [has_continuous_star 𝔸] (x : 𝔸) : + star (exp 𝕂 x) = exp 𝕂 (star x) := +by simp_rw [exp_eq_tsum, ←star_pow, ←star_inv_nat_cast_smul, ←tsum_star] + variables (𝕂) lemma commute.exp_right [t2_space 𝔸] {x y : 𝔸} (h : commute x y) : commute x (exp 𝕂 y) := @@ -599,17 +603,3 @@ lemma exp_ℝ_β„‚_eq_exp_β„‚_β„‚ : (exp ℝ : β„‚ β†’ β„‚) = exp β„‚ := exp_eq_exp ℝ β„‚ β„‚ end scalar_tower - -lemma star_exp {π•œ A : Type*} [is_R_or_C π•œ] [normed_ring A] [normed_algebra π•œ A] - [star_ring A] [normed_star_group A] [complete_space A] - [star_module π•œ A] (a : A) : star (exp π•œ a) = exp π•œ (star a) := -begin - rw exp_eq_tsum, - have := continuous_linear_map.map_tsum - (starβ‚—α΅’ π•œ : A ≃ₗᡒ⋆[π•œ] A).to_linear_isometry.to_continuous_linear_map - (exp_series_summable' a), - dsimp at this, - convert this, - funext, - simp only [star_smul, star_pow, one_div, star_inv', star_nat_cast], -end diff --git a/src/analysis/normed_space/star/exponential.lean b/src/analysis/normed_space/star/exponential.lean index 8e56fb4d51985..f7a147543e781 100644 --- a/src/analysis/normed_space/star/exponential.lean +++ b/src/analysis/normed_space/star/exponential.lean @@ -23,7 +23,7 @@ subtypes `self_adjoint A` and `unitary A`. section star variables {A : Type*} -[normed_ring A] [normed_algebra β„‚ A] [star_ring A] [cstar_ring A] [complete_space A] +[normed_ring A] [normed_algebra β„‚ A] [star_ring A] [has_continuous_star A] [complete_space A] [star_module β„‚ A] open complex diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 04a1ebc14cfec..00071d7fc1c9a 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -246,6 +246,26 @@ lemma function.surjective.summable_iff_of_has_sum_iff {Ξ±' : Type*} [add_comm_mo summable f ↔ summable g := hes.exists.trans $ exists_congr $ @he +section has_continuous_star +variables [star_add_monoid Ξ±] [has_continuous_star Ξ±] + +lemma has_sum.star (h : has_sum f a) : has_sum (Ξ» b, star (f b)) (star a) := +by simpa only using h.map (star_add_equiv : Ξ± ≃+ Ξ±) continuous_star + +lemma summable.star (hf : summable f) : summable (Ξ» b, star (f b)) := +hf.has_sum.star.summable + +lemma summable.of_star (hf : summable (Ξ» b, star (f b))) : summable f := +by simpa only [star_star] using hf.star + +@[simp] lemma summable_star_iff : summable (Ξ» b, star (f b)) ↔ summable f := +⟨summable.of_star, summable.star⟩ + +@[simp] lemma summable_star_iff' : summable (star f) ↔ summable f := +summable_star_iff + +end has_continuous_star + variable [has_continuous_add Ξ±] lemma has_sum.add (hf : has_sum f a) (hg : has_sum g b) : has_sum (Ξ»b, f b + g b) (a + b) := @@ -474,6 +494,19 @@ end end has_continuous_add +section has_continuous_star +variables [star_add_monoid Ξ±] [has_continuous_star Ξ±] + +lemma tsum_star : star (βˆ‘' b, f b) = βˆ‘' b, star (f b) := +begin + by_cases hf : summable f, + { exact hf.has_sum.star.tsum_eq.symm, }, + { rw [tsum_eq_zero_of_not_summable hf, tsum_eq_zero_of_not_summable (mt summable.of_star hf), + star_zero] }, +end + +end has_continuous_star + section encodable open encodable variable [encodable Ξ³]