diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 39166c444356f..94c9f84ddff57 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -53,12 +53,13 @@ theorem tendsto_norm_zero' {π•œ : Type*} [NormedAddCommGroup π•œ] : namespace NormedField -theorem tendsto_norm_inverse_nhdsWithin_0_atTop {π•œ : Type*} [NormedField π•œ] : +theorem tendsto_norm_inverse_nhdsWithin_0_atTop {π•œ : Type*} [NormedDivisionRing π•œ] : Tendsto (fun x : π•œ ↦ β€–x⁻¹‖) (𝓝[β‰ ] 0) atTop := (tendsto_inv_zero_atTop.comp tendsto_norm_zero').congr fun x ↦ (norm_inv x).symm #align normed_field.tendsto_norm_inverse_nhds_within_0_at_top NormedField.tendsto_norm_inverse_nhdsWithin_0_atTop -theorem tendsto_norm_zpow_nhdsWithin_0_atTop {π•œ : Type*} [NormedField π•œ] {m : β„€} (hm : m < 0) : +theorem tendsto_norm_zpow_nhdsWithin_0_atTop {π•œ : Type*} [NormedDivisionRing π•œ] {m : β„€} + (hm : m < 0) : Tendsto (fun x : π•œ ↦ β€–x ^ mβ€–) (𝓝[β‰ ] 0) atTop := by rcases neg_surjective m with ⟨m, rfl⟩ rw [neg_lt_zero] at hm; lift m to β„• using hm.le; rw [Int.natCast_pos] at hm @@ -67,8 +68,8 @@ theorem tendsto_norm_zpow_nhdsWithin_0_atTop {π•œ : Type*} [NormedField π•œ] { #align normed_field.tendsto_norm_zpow_nhds_within_0_at_top NormedField.tendsto_norm_zpow_nhdsWithin_0_atTop /-- The (scalar) product of a sequence that tends to zero with a bounded one also tends to zero. -/ -theorem tendsto_zero_smul_of_tendsto_zero_of_bounded {ΞΉ π•œ 𝔸 : Type*} [NormedField π•œ] - [NormedAddCommGroup 𝔸] [NormedSpace π•œ 𝔸] {l : Filter ΞΉ} {Ξ΅ : ΞΉ β†’ π•œ} {f : ΞΉ β†’ 𝔸} +theorem tendsto_zero_smul_of_tendsto_zero_of_bounded {ΞΉ π•œ 𝔸 : Type*} [NormedDivisionRing π•œ] + [NormedAddCommGroup 𝔸] [Module π•œ 𝔸] [BoundedSMul π•œ 𝔸] {l : Filter ΞΉ} {Ξ΅ : ΞΉ β†’ π•œ} {f : ΞΉ β†’ 𝔸} (hΞ΅ : Tendsto Ξ΅ l (𝓝 0)) (hf : Filter.IsBoundedUnder (Β· ≀ Β·) l (norm ∘ f)) : Tendsto (Ξ΅ β€’ f) l (𝓝 0) := by rw [← isLittleO_one_iff π•œ] at hΞ΅ ⊒ @@ -286,7 +287,7 @@ theorem tendsto_pow_atTop_nhds_zero_of_abs_lt_one {r : ℝ} (h : |r| < 1) : section Geometric -variable {K : Type*} [NormedField K] {ΞΎ : K} +variable {K : Type*} [NormedDivisionRing K] {ΞΎ : K} theorem hasSum_geometric_of_norm_lt_one (h : β€–ΞΎβ€– < 1) : HasSum (fun n : β„• ↦ ΞΎ ^ n) (1 - ΞΎ)⁻¹ := by have xi_ne_one : ΞΎ β‰  1 := by @@ -360,7 +361,7 @@ theorem summable_pow_mul_geometric_of_norm_lt_one {R : Type*} [NormedRing R] [Co summable_pow_mul_geometric_of_norm_lt_one /-- If `β€–rβ€– < 1`, then `βˆ‘' n : β„•, n * r ^ n = r / (1 - r) ^ 2`, `HasSum` version. -/ -theorem hasSum_coe_mul_geometric_of_norm_lt_one {π•œ : Type*} [NormedField π•œ] [CompleteSpace π•œ] +theorem hasSum_coe_mul_geometric_of_norm_lt_one {π•œ : Type*} [NormedDivisionRing π•œ] [CompleteSpace π•œ] {r : π•œ} (hr : β€–rβ€– < 1) : HasSum (fun n ↦ n * r ^ n : β„• β†’ π•œ) (r / (1 - r) ^ 2) := by have A : Summable (fun n ↦ (n : π•œ) * r ^ n : β„• β†’ π•œ) := by simpa only [pow_one] using summable_pow_mul_geometric_of_norm_lt_one 1 hr @@ -370,22 +371,28 @@ theorem hasSum_coe_mul_geometric_of_norm_lt_one {π•œ : Type*} [NormedField π•œ rintro rfl simp [lt_irrefl] at hr set s : π•œ := βˆ‘' n : β„•, n * r ^ n + have : Commute (1 - r) s := + .tsum_right _ fun _ => + .sub_left (.one_left _) (.mul_right (Nat.commute_cast _ _) (.pow_right (.refl _) _)) calc - s = (1 - r) * s / (1 - r) := (mul_div_cancel_leftβ‚€ _ (sub_ne_zero.2 hr'.symm)).symm + s = s * (1 - r) / (1 - r) := (mul_div_cancel_rightβ‚€ _ (sub_ne_zero.2 hr'.symm)).symm + _ = (1 - r) * s / (1 - r) := by rw [this.eq] _ = (s - r * s) / (1 - r) := by rw [_root_.sub_mul, one_mul] _ = (((0 : β„•) * r ^ 0 + βˆ‘' n : β„•, (n + 1 : β„•) * r ^ (n + 1)) - r * s) / (1 - r) := by rw [← tsum_eq_zero_add A] - _ = ((r * βˆ‘' n : β„•, (n + 1) * r ^ n) - r * s) / (1 - r) := by - simp [_root_.pow_succ', mul_left_comm _ r, _root_.tsum_mul_left] + _ = ((r * βˆ‘' n : β„•, ↑(n + 1) * r ^ n) - r * s) / (1 - r) := by + simp only [cast_zero, pow_zero, mul_one, _root_.pow_succ', (Nat.cast_commute _ r).left_comm, + _root_.tsum_mul_left, zero_add] _ = r / (1 - r) ^ 2 := by - simp [add_mul, tsum_add A B.summable, mul_add, B.tsum_eq, ← div_eq_mul_inv, sq, div_div] + simp [add_mul, tsum_add A B.summable, mul_add, B.tsum_eq, ← div_eq_mul_inv, sq, + div_mul_eq_div_div_swap] #align has_sum_coe_mul_geometric_of_norm_lt_1 hasSum_coe_mul_geometric_of_norm_lt_one @[deprecated] alias hasSum_coe_mul_geometric_of_norm_lt_1 := hasSum_coe_mul_geometric_of_norm_lt_one /-- If `β€–rβ€– < 1`, then `βˆ‘' n : β„•, n * r ^ n = r / (1 - r) ^ 2`. -/ -theorem tsum_coe_mul_geometric_of_norm_lt_one {π•œ : Type*} [NormedField π•œ] [CompleteSpace π•œ] {r : π•œ} - (hr : β€–rβ€– < 1) : (βˆ‘' n : β„•, n * r ^ n : π•œ) = r / (1 - r) ^ 2 := +theorem tsum_coe_mul_geometric_of_norm_lt_one {π•œ : Type*} [NormedDivisionRing π•œ] [CompleteSpace π•œ] + {r : π•œ} (hr : β€–rβ€– < 1) : (βˆ‘' n : β„•, n * r ^ n : π•œ) = r / (1 - r) ^ 2 := (hasSum_coe_mul_geometric_of_norm_lt_one hr).tsum_eq #align tsum_coe_mul_geometric_of_norm_lt_1 tsum_coe_mul_geometric_of_norm_lt_one @[deprecated] alias tsum_coe_mul_geometric_of_norm_lt_1 := tsum_coe_mul_geometric_of_norm_lt_one diff --git a/Mathlib/Topology/ContinuousFunction/Units.lean b/Mathlib/Topology/ContinuousFunction/Units.lean index daaff4ac1c0dd..359b581dd7395 100644 --- a/Mathlib/Topology/ContinuousFunction/Units.lean +++ b/Mathlib/Topology/ContinuousFunction/Units.lean @@ -103,17 +103,23 @@ end NormedRing section NormedField -variable [NormedField π•œ] [CompleteSpace π•œ] +variable [NormedField π•œ] [NormedDivisionRing R] [Algebra π•œ R] [CompleteSpace R] -theorem isUnit_iff_forall_ne_zero (f : C(X, π•œ)) : IsUnit f ↔ βˆ€ x, f x β‰  0 := by +theorem isUnit_iff_forall_ne_zero (f : C(X, R)) : IsUnit f ↔ βˆ€ x, f x β‰  0 := by simp_rw [f.isUnit_iff_forall_isUnit, isUnit_iff_ne_zero] #align continuous_map.is_unit_iff_forall_ne_zero ContinuousMap.isUnit_iff_forall_ne_zero -theorem spectrum_eq_range (f : C(X, π•œ)) : spectrum π•œ f = Set.range f := by +theorem spectrum_eq_preimage_range (f : C(X, R)) : + spectrum π•œ f = algebraMap _ _ ⁻¹' Set.range f := by ext x - simp only [spectrum.mem_iff, isUnit_iff_forall_ne_zero, not_forall, coe_sub, Pi.sub_apply, - algebraMap_apply, Algebra.id.smul_eq_mul, mul_one, Classical.not_not, Set.mem_range, - sub_eq_zero, @eq_comm _ x _] + simp only [spectrum.mem_iff, isUnit_iff_forall_ne_zero, not_forall, sub_apply, + algebraMap_apply, mul_one, Classical.not_not, Set.mem_range, + sub_eq_zero, @eq_comm _ (x β€’ 1 : R) _, Set.mem_preimage, Algebra.algebraMap_eq_smul_one, + smul_apply, one_apply] + +theorem spectrum_eq_range [CompleteSpace π•œ] (f : C(X, π•œ)) : spectrum π•œ f = Set.range f := by + rw [spectrum_eq_preimage_range, Algebra.id.map_eq_id] + exact Set.preimage_id #align continuous_map.spectrum_eq_range ContinuousMap.spectrum_eq_range end NormedField