diff --git a/Mathlib/Analysis/NormedSpace/Spectrum.lean b/Mathlib/Analysis/NormedSpace/Spectrum.lean index 1567743ee20f1e..e05f0465c4534e 100644 --- a/Mathlib/Analysis/NormedSpace/Spectrum.lean +++ b/Mathlib/Analysis/NormedSpace/Spectrum.lean @@ -588,3 +588,65 @@ theorem equivAlgHom_symm_coe (f : A →ₐ[π•œ] π•œ) : ⇑(equivAlgHom.symm f) end CharacterSpace end WeakDual + +namespace SpectrumRestricts + +open NNReal ENNReal + +/-- If `π•œβ‚` is a normed field contained as subfield of a larger normed field `π•œβ‚‚`, and if `a : A` +is an element whose `π•œβ‚‚` spectrum restricts to `π•œβ‚`, then the spectral radii over each scalar +field coincide. -/ +lemma spectralRadius_eq {π•œβ‚ π•œβ‚‚ A : Type*} [NormedField π•œβ‚] [NormedField π•œβ‚‚] + [NormedRing A] [NormedAlgebra π•œβ‚ A] [NormedAlgebra π•œβ‚‚ A] [NormedAlgebra π•œβ‚ π•œβ‚‚] + [IsScalarTower π•œβ‚ π•œβ‚‚ A] {f : π•œβ‚‚ β†’ π•œβ‚} {a : A} (h : SpectrumRestricts a f) : + spectralRadius π•œβ‚ a = spectralRadius π•œβ‚‚ a := by + rw [spectralRadius, spectralRadius] + have := algebraMap_isometry π•œβ‚ π•œβ‚‚ |>.nnnorm_map_of_map_zero (map_zero _) + apply le_antisymm + all_goals apply iSupβ‚‚_le fun x hx ↦ ?_ + Β· refine congr_arg ((↑) : ℝβ‰₯0 β†’ ℝβ‰₯0∞) (this x) |>.symm.trans_le <| le_iSupβ‚‚ (Ξ± := ℝβ‰₯0∞) _ ?_ + exact (spectrum.algebraMap_mem_iff _).mpr hx + Β· have ⟨y, hy, hy'⟩ := h.algebraMap_image.symm β–Έ hx + subst hy' + exact this y β–Έ le_iSupβ‚‚ (Ξ± := ℝβ‰₯0∞) y hy + +variable {A : Type*} [Ring A] + +lemma nnreal_iff [Algebra ℝ A] {a : A} : + SpectrumRestricts a ContinuousMap.realToNNReal ↔ βˆ€ x ∈ spectrum ℝ a, 0 ≀ x := by + refine ⟨fun h x hx ↦ ?_, fun h ↦ ?_⟩ + Β· obtain ⟨x, -, rfl⟩ := h.algebraMap_image.symm β–Έ hx + exact coe_nonneg x + Β· exact .of_subset_range_algebraMap _ _ (fun _ ↦ Real.toNNReal_coe) + fun x hx ↦ ⟨⟨x, h x hx⟩, rfl⟩ + +lemma real_iff [Algebra β„‚ A] {a : A} : + SpectrumRestricts a Complex.reCLM ↔ βˆ€ x ∈ spectrum β„‚ a, x = x.re := by + refine ⟨fun h x hx ↦ ?_, fun h ↦ ?_⟩ + Β· obtain ⟨x, -, rfl⟩ := h.algebraMap_image.symm β–Έ hx + simp + Β· exact .of_subset_range_algebraMap _ _ Complex.ofReal_re fun x hx ↦ ⟨x.re, (h x hx).symm⟩ + +lemma nnreal_iff_spectralRadius_le [Algebra ℝ A] {a : A} {t : ℝβ‰₯0} (ht : spectralRadius ℝ a ≀ t) : + SpectrumRestricts a ContinuousMap.realToNNReal ↔ + spectralRadius ℝ (algebraMap ℝ A t - a) ≀ t := by + have : spectrum ℝ a βŠ† Set.Icc (-t) t := by + intro x hx + rw [Set.mem_Icc, ← abs_le, ← Real.norm_eq_abs, ← coe_nnnorm, NNReal.coe_le_coe, + ← ENNReal.coe_le_coe] + exact le_iSupβ‚‚ (Ξ± := ℝβ‰₯0∞) x hx |>.trans ht + rw [nnreal_iff] + refine ⟨fun h ↦ iSupβ‚‚_le fun x hx ↦ ?_, fun h ↦ ?_⟩ + Β· rw [← spectrum.singleton_sub_eq] at hx + obtain ⟨y, hy, rfl⟩ : βˆƒ y ∈ spectrum ℝ a, ↑t - y = x := by simpa using hx + obtain ⟨hty, hyt⟩ := Set.mem_Icc.mp <| this hy + lift y to ℝβ‰₯0 using h y hy + rw [← NNReal.coe_sub (by exact_mod_cast hyt)] + simp + Β· replace h : βˆ€ x ∈ spectrum ℝ a, β€–t - xβ€–β‚Š ≀ t := by + simpa [spectralRadius, iSupβ‚‚_le_iff, ← spectrum.singleton_sub_eq] using h + peel h with x hx h_le + rw [← NNReal.coe_le_coe, coe_nnnorm, Real.norm_eq_abs, abs_le] at h_le + linarith [h_le.2] + +end SpectrumRestricts