diff --git a/src/analysis/inner_product_space/pi_L2.lean b/src/analysis/inner_product_space/pi_L2.lean index ab8caa6dc83db..d4ae133a5b30f 100644 --- a/src/analysis/inner_product_space/pi_L2.lean +++ b/src/analysis/inner_product_space/pi_L2.lean @@ -701,7 +701,7 @@ sum has an orthonormal basis indexed by `fin n` and subordinate to that direct s /-- An `n`-dimensional `inner_product_space` equipped with a decomposition as an internal direct sum has an orthonormal basis indexed by `fin n` and subordinate to that direct sum. This function provides the mapping by which it is subordinate. -/ -def direct_sum.is_internal.subordinate_orthonormal_basis_index +@[irreducible] def direct_sum.is_internal.subordinate_orthonormal_basis_index (a : fin n) (hV' : @orthogonal_family 𝕜 _ _ _ _ (λ i, V i) _ (λ i, (V i).subtypeₗᵢ)) : ι := ((hV.sigma_orthonormal_basis_index_equiv hn hV').symm a).1 @@ -712,12 +712,10 @@ lemma direct_sum.is_internal.subordinate_orthonormal_basis_subordinate (hV.subordinate_orthonormal_basis hn hV' a) ∈ V (hV.subordinate_orthonormal_basis_index hn a hV') := by simpa only [direct_sum.is_internal.subordinate_orthonormal_basis, - orthonormal_basis.coe_reindex] + orthonormal_basis.coe_reindex, direct_sum.is_internal.subordinate_orthonormal_basis_index] using hV.collected_orthonormal_basis_mem hV' (λ i, (std_orthonormal_basis 𝕜 (V i))) ((hV.sigma_orthonormal_basis_index_equiv hn hV').symm a) -attribute [irreducible] direct_sum.is_internal.subordinate_orthonormal_basis_index - end subordinate_orthonormal_basis end finite_dimensional diff --git a/src/analysis/inner_product_space/spectrum.lean b/src/analysis/inner_product_space/spectrum.lean index f6d03066c8430..07ce6120ccac8 100644 --- a/src/analysis/inner_product_space/spectrum.lean +++ b/src/analysis/inner_product_space/spectrum.lean @@ -179,7 +179,7 @@ finite-dimensional inner product space `E`. TODO Postcompose with a permutation so that these eigenvectors are listed in increasing order of eigenvalue. -/ -noncomputable def eigenvector_basis : orthonormal_basis (fin n) 𝕜 E := +@[irreducible] noncomputable def eigenvector_basis : orthonormal_basis (fin n) 𝕜 E := hT.direct_sum_is_internal.subordinate_orthonormal_basis hn hT.orthogonal_family_eigenspaces' @@ -187,7 +187,7 @@ hT.direct_sum_is_internal.subordinate_orthonormal_basis hn for a self-adjoint operator `T` on `E`. TODO Postcompose with a permutation so that these eigenvalues are listed in increasing order. -/ -noncomputable def eigenvalues (i : fin n) : ℝ := +@[irreducible] noncomputable def eigenvalues (i : fin n) : ℝ := @is_R_or_C.re 𝕜 _ $ hT.direct_sum_is_internal.subordinate_orthonormal_basis_index hn i hT.orthogonal_family_eigenspaces' @@ -198,11 +198,13 @@ begin let v : E := hT.eigenvector_basis hn i, let μ : 𝕜 := hT.direct_sum_is_internal.subordinate_orthonormal_basis_index hn i hT.orthogonal_family_eigenspaces', + simp_rw [eigenvalues], change has_eigenvector T (is_R_or_C.re μ) v, have key : has_eigenvector T μ v, { have H₁ : v ∈ eigenspace T μ, - { exact hT.direct_sum_is_internal.subordinate_orthonormal_basis_subordinate - hn i hT.orthogonal_family_eigenspaces' }, + { simp_rw [v, eigenvector_basis], + exact hT.direct_sum_is_internal.subordinate_orthonormal_basis_subordinate + hn i hT.orthogonal_family_eigenspaces' }, have H₂ : v ≠ 0 := by simpa using (hT.eigenvector_basis hn).to_basis.ne_zero i, exact ⟨H₁, H₂⟩ }, have re_μ : ↑(is_R_or_C.re μ) = μ, @@ -212,9 +214,7 @@ begin end lemma has_eigenvalue_eigenvalues (i : fin n) : has_eigenvalue T (hT.eigenvalues hn i) := - module.End.has_eigenvalue_of_has_eigenvector (hT.has_eigenvector_eigenvector_basis hn i) - -attribute [irreducible] eigenvector_basis eigenvalues +module.End.has_eigenvalue_of_has_eigenvector (hT.has_eigenvector_eigenvector_basis hn i) @[simp] lemma apply_eigenvector_basis (i : fin n) : T (hT.eigenvector_basis hn i) = (hT.eigenvalues hn i : 𝕜) • hT.eigenvector_basis hn i := diff --git a/src/data/complex/exponential.lean b/src/data/complex/exponential.lean index 21c7ea846b9e9..f41c5375fa3cc 100644 --- a/src/data/complex/exponential.lean +++ b/src/data/complex/exponential.lean @@ -337,7 +337,7 @@ the complex exponential function -/ ⟨λ n, ∑ m in range n, z ^ m / m!, is_cau_exp z⟩ /-- The complex exponential function, defined via its Taylor series -/ -@[pp_nodot] def exp (z : ℂ) : ℂ := lim (exp' z) +@[irreducible, pp_nodot] def exp (z : ℂ) : ℂ := lim (exp' z) /-- The complex sine function, defined via `exp` -/ @[pp_nodot] def sin (z : ℂ) : ℂ := ((exp (-z * I) - exp (z * I)) * I) / 2 @@ -392,8 +392,9 @@ namespace complex variables (x y : ℂ) @[simp] lemma exp_zero : exp 0 = 1 := -lim_eq_of_equiv_const $ - λ ε ε0, ⟨1, λ j hj, begin +begin + rw exp, + refine lim_eq_of_equiv_const (λ ε ε0, ⟨1, λ j hj, _⟩), convert ε0, cases j, { exact absurd hj (not_le_of_gt zero_lt_one) }, @@ -403,33 +404,29 @@ lim_eq_of_equiv_const $ { rw ← ih dec_trivial, simp only [sum_range_succ, pow_succ], simp } } -end⟩ +end lemma exp_add : exp (x + y) = exp x * exp y := -show lim (⟨_, is_cau_exp (x + y)⟩ : cau_seq ℂ abs) = - lim (show cau_seq ℂ abs, from ⟨_, is_cau_exp x⟩) - * lim (show cau_seq ℂ abs, from ⟨_, is_cau_exp y⟩), -from -have hj : ∀ j : ℕ, ∑ m in range j, (x + y) ^ m / m! = - ∑ i in range j, ∑ k in range (i + 1), x ^ k / k! * (y ^ (i - k) / (i - k)!), - from assume j, - finset.sum_congr rfl (λ m hm, begin - rw [add_pow, div_eq_mul_inv, sum_mul], - refine finset.sum_congr rfl (λ i hi, _), - have h₁ : (m.choose i : ℂ) ≠ 0 := nat.cast_ne_zero.2 - (pos_iff_ne_zero.1 (nat.choose_pos (nat.le_of_lt_succ (mem_range.1 hi)))), - have h₂ := nat.choose_mul_factorial_mul_factorial (nat.le_of_lt_succ $ finset.mem_range.1 hi), - rw [← h₂, nat.cast_mul, nat.cast_mul, mul_inv, mul_inv], - simp only [mul_left_comm (m.choose i : ℂ), mul_assoc, mul_left_comm (m.choose i : ℂ)⁻¹, - mul_comm (m.choose i : ℂ)], - rw inv_mul_cancel h₁, - simp [div_eq_mul_inv, mul_comm, mul_assoc, mul_left_comm] - end), -by rw lim_mul_lim; - exact eq.symm (lim_eq_lim_of_equiv (by dsimp; simp only [hj]; - exact cauchy_product (is_cau_abs_exp x) (is_cau_exp y))) - -attribute [irreducible] complex.exp +begin + have hj : ∀ j : ℕ, ∑ m in range j, (x + y) ^ m / m! = + ∑ i in range j, ∑ k in range (i + 1), x ^ k / k! * (y ^ (i - k) / (i - k)!), + { assume j, + refine finset.sum_congr rfl (λ m hm, _), + rw [add_pow, div_eq_mul_inv, sum_mul], + refine finset.sum_congr rfl (λ i hi, _), + have h₁ : (m.choose i : ℂ) ≠ 0 := nat.cast_ne_zero.2 + (pos_iff_ne_zero.1 (nat.choose_pos (nat.le_of_lt_succ (mem_range.1 hi)))), + have h₂ := nat.choose_mul_factorial_mul_factorial (nat.le_of_lt_succ $ finset.mem_range.1 hi), + rw [← h₂, nat.cast_mul, nat.cast_mul, mul_inv, mul_inv], + simp only [mul_left_comm (m.choose i : ℂ), mul_assoc, mul_left_comm (m.choose i : ℂ)⁻¹, + mul_comm (m.choose i : ℂ)], + rw inv_mul_cancel h₁, + simp [div_eq_mul_inv, mul_comm, mul_assoc, mul_left_comm] }, + simp_rw [exp, exp', lim_mul_lim], + apply (lim_eq_lim_of_equiv _).symm, + simp only [hj], + exact cauchy_product (is_cau_abs_exp x) (is_cau_exp y) +end lemma exp_list_sum (l : list ℂ) : exp l.sum = (l.map exp).prod := @monoid_hom.map_list_prod (multiplicative ℂ) ℂ _ _ ⟨exp, exp_zero, exp_add⟩ l diff --git a/src/ring_theory/fractional_ideal.lean b/src/ring_theory/fractional_ideal.lean index 458ac481a7a8f..7a28d8888942c 100644 --- a/src/ring_theory/fractional_ideal.lean +++ b/src/ring_theory/fractional_ideal.lean @@ -1094,14 +1094,15 @@ variables (S) def span_singleton (x : P) : fractional_ideal S P := ⟨span R {x}, is_fractional_span_singleton x⟩ -local attribute [semireducible] span_singleton +-- local attribute [semireducible] span_singleton @[simp] lemma coe_span_singleton (x : P) : - (span_singleton S x : submodule R P) = span R {x} := rfl + (span_singleton S x : submodule R P) = span R {x} := +by { rw span_singleton, refl } @[simp] lemma mem_span_singleton {x y : P} : x ∈ span_singleton S y ↔ ∃ (z : R), z • y = x := -submodule.mem_span_singleton +by { rw span_singleton, exact submodule.mem_span_singleton } lemma mem_span_singleton_self (x : P) : x ∈ span_singleton S x := @@ -1111,12 +1112,13 @@ variables {S} lemma span_singleton_eq_span_singleton [no_zero_smul_divisors R P] {x y : P} : span_singleton S x = span_singleton S y ↔ ∃ z : Rˣ, z • x = y := -by { rw [← submodule.span_singleton_eq_span_singleton], exact subtype.mk_eq_mk } +by { rw [← submodule.span_singleton_eq_span_singleton, span_singleton, span_singleton], + exact subtype.mk_eq_mk } lemma eq_span_singleton_of_principal (I : fractional_ideal S P) [is_principal (I : submodule R P)] : I = span_singleton S (generator (I : submodule R P)) := -coe_to_submodule_injective (span_singleton_generator ↑I).symm +by { rw span_singleton, exact coe_to_submodule_injective (span_singleton_generator ↑I).symm } lemma is_principal_iff (I : fractional_ideal S P) : is_principal (I : submodule R P) ↔ ∃ x, I = span_singleton S x := @@ -1301,7 +1303,7 @@ begin obtain ⟨a, aI, -, ha⟩ := exists_eq_span_singleton_mul I, use (algebra_map R K a)⁻¹ * algebra_map R K (generator aI), suffices : I = span_singleton R⁰ ((algebra_map R K a)⁻¹ * algebra_map R K (generator aI)), - { exact congr_arg subtype.val this }, + { rw span_singleton at this, exact congr_arg subtype.val this }, conv_lhs { rw [ha, ←span_singleton_generator aI] }, rw [ideal.submodule_span_eq, coe_ideal_span_singleton (generator aI), span_singleton_mul_span_singleton]