From 7a0fe7add9a7e3659523b92b9b704800091b2210 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 13:57:29 +0000 Subject: [PATCH 01/23] feat: port Data.Complex.Exponential From 21527810eafcfce443bd100e743b003207fa370c Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 13:57:29 +0000 Subject: [PATCH 02/23] Initial file copy from mathport --- Mathlib/Data/Complex/Exponential.lean | 2122 +++++++++++++++++++++++++ 1 file changed, 2122 insertions(+) create mode 100644 Mathlib/Data/Complex/Exponential.lean diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean new file mode 100644 index 0000000000000..c7cf3ba3c0921 --- /dev/null +++ b/Mathlib/Data/Complex/Exponential.lean @@ -0,0 +1,2122 @@ +/- +Copyright (c) 2018 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes, Abhimanyu Pallavi Sudhir + +! This file was ported from Lean 3 source module data.complex.exponential +! leanprover-community/mathlib commit 372edc36e5d2caafdd135769e0136b5a59186834 +! Please do not edit these lines, except to modify the commit id +! if you have ported upstream changes. +-/ +import Mathbin.Algebra.GeomSum +import Mathbin.Data.Complex.Basic +import Mathbin.Data.Nat.Choose.Sum + +/-! +# Exponential, trigonometric and hyperbolic trigonometric functions + +This file contains the definitions of the real and complex exponential, sine, cosine, tangent, +hyperbolic sine, hyperbolic cosine, and hyperbolic tangent functions. + +-/ + + +-- mathport name: exprabs' +local notation "abs'" => Abs.abs + +open IsAbsoluteValue + +open Classical BigOperators Nat ComplexConjugate + +section + +open Real IsAbsoluteValue Finset + +section + +variable {α : Type _} {β : Type _} [Ring β] [LinearOrderedField α] [Archimedean α] {abv : β → α} + [IsAbsoluteValue abv] + +theorem is_cau_of_decreasing_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) + (hnm : ∀ n ≥ m, f n.succ ≤ f n) : IsCauSeq abs f := fun ε ε0 => + by + let ⟨k, hk⟩ := Archimedean.arch a ε0 + have h : ∃ l, ∀ n ≥ m, a - l • ε < f n := + ⟨k + k + 1, fun n hnm => + lt_of_lt_of_le + (show a - (k + (k + 1)) • ε < -|f n| from + lt_neg.1 <| + lt_of_le_of_lt (ham n hnm) + (by + rw [neg_sub, lt_sub_iff_add_lt, add_nsmul, add_nsmul, one_nsmul] + exact add_lt_add_of_le_of_lt hk (lt_of_le_of_lt hk (lt_add_of_pos_right _ ε0)))) + (neg_le.2 <| abs_neg (f n) ▸ le_abs_self _)⟩ + let l := Nat.find h + have hl : ∀ n : ℕ, n ≥ m → f n > a - l • ε := Nat.find_spec h + have hl0 : l ≠ 0 := fun hl0 => + not_lt_of_ge (ham m le_rfl) + (lt_of_lt_of_le (by have := hl m (le_refl m) <;> simpa [hl0] using this) (le_abs_self (f m))) + cases' not_forall.1 (Nat.find_min h (Nat.pred_lt hl0)) with i hi + rw [not_imp, not_lt] at hi + exists i + intro j hj + have hfij : f j ≤ f i := (Nat.rel_of_forall_rel_succ_of_le_of_le (· ≥ ·) hnm hi.1 hj).le + rw [abs_of_nonpos (sub_nonpos.2 hfij), neg_sub, sub_lt_iff_lt_add'] + calc + f i ≤ a - Nat.pred l • ε := hi.2 + _ = a - l • ε + ε := by + conv => + rhs + rw [← Nat.succ_pred_eq_of_pos (Nat.pos_of_ne_zero hl0), succ_nsmul', sub_add, + add_sub_cancel] + _ < f j + ε := add_lt_add_right (hl j (le_trans hi.1 hj)) _ + +#align is_cau_of_decreasing_bounded is_cau_of_decreasing_bounded + +theorem is_cau_of_mono_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) + (hnm : ∀ n ≥ m, f n ≤ f n.succ) : IsCauSeq abs f := + by + refine' + @Eq.recOn (ℕ → α) _ (IsCauSeq abs) _ _ + (-⟨_, @is_cau_of_decreasing_bounded _ _ _ (fun n => -f n) a m (by simpa) (by simpa)⟩ : + CauSeq α abs).2 + ext + exact neg_neg _ +#align is_cau_of_mono_bounded is_cau_of_mono_bounded + +end + +section NoArchimedean + +variable {α : Type _} {β : Type _} [Ring β] [LinearOrderedField α] {abv : β → α} + [IsAbsoluteValue abv] + +theorem is_cau_series_of_abv_le_cau {f : ℕ → β} {g : ℕ → α} (n : ℕ) : + (∀ m, n ≤ m → abv (f m) ≤ g m) → + (IsCauSeq abs fun n => ∑ i in range n, g i) → IsCauSeq abv fun n => ∑ i in range n, f i := + by + intro hm hg ε ε0 + cases' hg (ε / 2) (div_pos ε0 (by norm_num)) with i hi + exists max n i + intro j ji + have hi₁ := hi j (le_trans (le_max_right n i) ji) + have hi₂ := hi (max n i) (le_max_right n i) + have sub_le := + abs_sub_le (∑ k in range j, g k) (∑ k in range i, g k) (∑ k in range (max n i), g k) + have := add_lt_add hi₁ hi₂ + rw [abs_sub_comm (∑ k in range (max n i), g k), add_halves ε] at this + refine' lt_of_le_of_lt (le_trans (le_trans _ (le_abs_self _)) sub_le) this + generalize hk : j - max n i = k + clear this hi₂ hi₁ hi ε0 ε hg sub_le + rw [tsub_eq_iff_eq_add_of_le ji] at hk + rw [hk] + clear hk ji j + induction' k with k' hi + · simp [abv_zero abv] + · simp only [Nat.succ_add, sum_range_succ_comm, sub_eq_add_neg, add_assoc] + refine' le_trans (abv_add _ _ _) _ + simp only [sub_eq_add_neg] at hi + exact add_le_add (hm _ (le_add_of_nonneg_of_le (Nat.zero_le _) (le_max_left _ _))) hi +#align is_cau_series_of_abv_le_cau is_cau_series_of_abv_le_cau + +theorem is_cau_series_of_abv_cau {f : ℕ → β} : + (IsCauSeq abs fun m => ∑ n in range m, abv (f n)) → IsCauSeq abv fun m => ∑ n in range m, f n := + is_cau_series_of_abv_le_cau 0 fun n h => le_rfl +#align is_cau_series_of_abv_cau is_cau_series_of_abv_cau + +end NoArchimedean + +section + +variable {α : Type _} [LinearOrderedField α] [Archimedean α] + +theorem is_cau_geo_series {β : Type _} [Ring β] [Nontrivial β] {abv : β → α} [IsAbsoluteValue abv] + (x : β) (hx1 : abv x < 1) : IsCauSeq abv fun n => ∑ m in range n, x ^ m := + have hx1' : abv x ≠ 1 := fun h => by simpa [h, lt_irrefl] using hx1 + is_cau_series_of_abv_cau + (by + simp only [abv_pow abv, geom_sum_eq hx1'] + conv in _ / _ => rw [← neg_div_neg_eq, neg_sub, neg_sub] + refine' @is_cau_of_mono_bounded _ _ _ _ ((1 : α) / (1 - abv x)) 0 _ _ + · intro n hn + rw [abs_of_nonneg] + refine' + div_le_div_of_le (le_of_lt <| sub_pos.2 hx1) + (sub_le_self _ (abv_pow abv x n ▸ abv_nonneg _ _)) + refine' div_nonneg (sub_nonneg.2 _) (sub_nonneg.2 <| le_of_lt hx1) + clear hn + induction' n with n ih + · simp + · rw [pow_succ, ← one_mul (1 : α)] + refine' mul_le_mul (le_of_lt hx1) ih (abv_pow abv x n ▸ abv_nonneg _ _) (by norm_num) + · intro n hn + refine' div_le_div_of_le (le_of_lt <| sub_pos.2 hx1) (sub_le_sub_left _ _) + rw [← one_mul (_ ^ n), pow_succ] + exact mul_le_mul_of_nonneg_right (le_of_lt hx1) (pow_nonneg (abv_nonneg _ _) _)) +#align is_cau_geo_series is_cau_geo_series + +theorem is_cau_geo_series_const (a : α) {x : α} (hx1 : |x| < 1) : + IsCauSeq abs fun m => ∑ n in range m, a * x ^ n := + by + have : IsCauSeq abs fun m => a * ∑ n in range m, x ^ n := + (CauSeq.const abs a * ⟨_, is_cau_geo_series x hx1⟩).2 + simpa only [mul_sum] +#align is_cau_geo_series_const is_cau_geo_series_const + +variable {β : Type _} [Ring β] {abv : β → α} [IsAbsoluteValue abv] + +theorem series_ratio_test {f : ℕ → β} (n : ℕ) (r : α) (hr0 : 0 ≤ r) (hr1 : r < 1) + (h : ∀ m, n ≤ m → abv (f m.succ) ≤ r * abv (f m)) : IsCauSeq abv fun m => ∑ n in range m, f n := + by + have har1 : |r| < 1 := by rwa [abs_of_nonneg hr0] + refine' + is_cau_series_of_abv_le_cau n.succ _ + (is_cau_geo_series_const (abv (f n.succ) * r⁻¹ ^ n.succ) har1) + intro m hmn + cases' Classical.em (r = 0) with r_zero r_ne_zero + · have m_pos := lt_of_lt_of_le (Nat.succ_pos n) hmn + have := h m.pred (Nat.le_of_succ_le_succ (by rwa [Nat.succ_pred_eq_of_pos m_pos])) + simpa [r_zero, Nat.succ_pred_eq_of_pos m_pos, pow_succ] + generalize hk : m - n.succ = k + have r_pos : 0 < r := lt_of_le_of_ne hr0 (Ne.symm r_ne_zero) + replace hk : m = k + n.succ := (tsub_eq_iff_eq_add_of_le hmn).1 hk + induction' k with k ih generalizing m n + · rw [hk, zero_add, mul_right_comm, inv_pow _ _, ← div_eq_mul_inv, mul_div_cancel] + exact (ne_of_lt (pow_pos r_pos _)).symm + · have kn : k + n.succ ≥ n.succ := by + rw [← zero_add n.succ] <;> exact add_le_add (zero_le _) (by simp) + rw [hk, Nat.succ_add, pow_succ' r, ← mul_assoc] + exact + le_trans (by rw [mul_comm] <;> exact h _ (Nat.le_of_succ_le kn)) + (mul_le_mul_of_nonneg_right (ih (k + n.succ) n h kn rfl) hr0) +#align series_ratio_test series_ratio_test + +theorem sum_range_diag_flip {α : Type _} [AddCommMonoid α] (n : ℕ) (f : ℕ → ℕ → α) : + (∑ m in range n, ∑ k in range (m + 1), f k (m - k)) = + ∑ m in range n, ∑ k in range (n - m), f m k := + by + rw [sum_sigma', sum_sigma'] <;> + exact + sum_bij (fun a _ => ⟨a.2, a.1 - a.2⟩) + (fun a ha => + have h₁ : a.1 < n := mem_range.1 (mem_sigma.1 ha).1 + have h₂ : a.2 < Nat.succ a.1 := mem_range.1 (mem_sigma.1 ha).2 + mem_sigma.2 + ⟨mem_range.2 (lt_of_lt_of_le h₂ h₁), + mem_range.2 ((tsub_lt_tsub_iff_right (Nat.le_of_lt_succ h₂)).2 h₁)⟩) + (fun _ _ => rfl) + (fun ⟨a₁, a₂⟩ ⟨b₁, b₂⟩ ha hb h => + have ha : a₁ < n ∧ a₂ ≤ a₁ := + ⟨mem_range.1 (mem_sigma.1 ha).1, Nat.le_of_lt_succ (mem_range.1 (mem_sigma.1 ha).2)⟩ + have hb : b₁ < n ∧ b₂ ≤ b₁ := + ⟨mem_range.1 (mem_sigma.1 hb).1, Nat.le_of_lt_succ (mem_range.1 (mem_sigma.1 hb).2)⟩ + have h : a₂ = b₂ ∧ _ := Sigma.mk.inj h + have h' : a₁ = b₁ - b₂ + a₂ := (tsub_eq_iff_eq_add_of_le ha.2).1 (eq_of_hEq h.2) + Sigma.mk.inj_iff.2 ⟨tsub_add_cancel_of_le hb.2 ▸ h'.symm ▸ h.1 ▸ rfl, hEq_of_eq h.1⟩) + fun ⟨a₁, a₂⟩ ha => + have ha : a₁ < n ∧ a₂ < n - a₁ := + ⟨mem_range.1 (mem_sigma.1 ha).1, mem_range.1 (mem_sigma.1 ha).2⟩ + ⟨⟨a₂ + a₁, a₁⟩, + ⟨mem_sigma.2 + ⟨mem_range.2 (lt_tsub_iff_right.1 ha.2), + mem_range.2 (Nat.lt_succ_of_le (Nat.le_add_left _ _))⟩, + Sigma.mk.inj_iff.2 ⟨rfl, hEq_of_eq (add_tsub_cancel_right _ _).symm⟩⟩⟩ +#align sum_range_diag_flip sum_range_diag_flip + +end + +section NoArchimedean + +variable {α : Type _} {β : Type _} [LinearOrderedField α] {abv : β → α} + +section + +variable [Semiring β] [IsAbsoluteValue abv] + +theorem abv_sum_le_sum_abv {γ : Type _} (f : γ → β) (s : Finset γ) : + abv (∑ k in s, f k) ≤ ∑ k in s, abv (f k) := + haveI := Classical.decEq γ + Finset.induction_on s (by simp [abv_zero abv]) fun a s has ih => by + rw [sum_insert has, sum_insert has] <;> exact le_trans (abv_add abv _ _) (add_le_add_left ih _) +#align abv_sum_le_sum_abv abv_sum_le_sum_abv + +end + +section + +variable [Ring β] [IsAbsoluteValue abv] + +theorem cauchy_product {a b : ℕ → β} (ha : IsCauSeq abs fun m => ∑ n in range m, abv (a n)) + (hb : IsCauSeq abv fun m => ∑ n in range m, b n) (ε : α) (ε0 : 0 < ε) : + ∃ i : ℕ, + ∀ j ≥ i, + abv + (((∑ k in range j, a k) * ∑ k in range j, b k) - + ∑ n in range j, ∑ m in range (n + 1), a m * b (n - m)) < + ε := + let ⟨Q, hQ⟩ := CauSeq.bounded ⟨_, hb⟩ + let ⟨P, hP⟩ := CauSeq.bounded ⟨_, ha⟩ + have hP0 : 0 < P := lt_of_le_of_lt (abs_nonneg _) (hP 0) + have hPε0 : 0 < ε / (2 * P) := div_pos ε0 (mul_pos (show (2 : α) > 0 by norm_num) hP0) + let ⟨N, hN⟩ := CauSeq.cauchy₂ ⟨_, hb⟩ hPε0 + have hQε0 : 0 < ε / (4 * Q) := + div_pos ε0 (mul_pos (show (0 : α) < 4 by norm_num) (lt_of_le_of_lt (abv_nonneg _ _) (hQ 0))) + let ⟨M, hM⟩ := CauSeq.cauchy₂ ⟨_, ha⟩ hQε0 + ⟨2 * (max N M + 1), fun K hK => + by + have h₁ : + (∑ m in range K, ∑ k in range (m + 1), a k * b (m - k)) = + ∑ m in range K, ∑ n in range (K - m), a m * b n := + by simpa using sum_range_diag_flip K fun m n => a m * b n + have h₂ : + (fun i => ∑ k in range (K - i), a i * b k) = fun i => a i * ∑ k in range (K - i), b k := by + simp [Finset.mul_sum] + have h₃ : + (∑ i in range K, a i * ∑ k in range (K - i), b k) = + (∑ i in range K, a i * ((∑ k in range (K - i), b k) - ∑ k in range K, b k)) + + ∑ i in range K, a i * ∑ k in range K, b k := + by rw [← sum_add_distrib] <;> simp [(mul_add _ _ _).symm] + have two_mul_two : (4 : α) = 2 * 2 := by norm_num + have hQ0 : Q ≠ 0 := fun h => by simpa [h, lt_irrefl] using hQε0 + have h2Q0 : 2 * Q ≠ 0 := mul_ne_zero two_ne_zero hQ0 + have hε : ε / (2 * P) * P + ε / (4 * Q) * (2 * Q) = ε := by + rw [← div_div, div_mul_cancel _ (Ne.symm (ne_of_lt hP0)), two_mul_two, mul_assoc, ← div_div, + div_mul_cancel _ h2Q0, add_halves] + have hNMK : max N M + 1 < K := + lt_of_lt_of_le (by rw [two_mul] <;> exact lt_add_of_pos_left _ (Nat.succ_pos _)) hK + have hKN : N < K := + calc + N ≤ max N M := le_max_left _ _ + _ < max N M + 1 := (Nat.lt_succ_self _) + _ < K := hNMK + + have hsumlesum : + (∑ i in range (max N M + 1), + abv (a i) * abv ((∑ k in range (K - i), b k) - ∑ k in range K, b k)) ≤ + ∑ i in range (max N M + 1), abv (a i) * (ε / (2 * P)) := + sum_le_sum fun m hmJ => + mul_le_mul_of_nonneg_left + (le_of_lt + (hN (K - m) + (le_tsub_of_add_le_left + (le_trans + (by + rw [two_mul] <;> + exact + add_le_add (le_of_lt (mem_range.1 hmJ)) + (le_trans (le_max_left _ _) (le_of_lt (lt_add_one _)))) + hK)) + K (le_of_lt hKN))) + (abv_nonneg abv _) + have hsumltP : (∑ n in range (max N M + 1), abv (a n)) < P := + calc + (∑ n in range (max N M + 1), abv (a n)) = |∑ n in range (max N M + 1), abv (a n)| := + Eq.symm (abs_of_nonneg (sum_nonneg fun x h => abv_nonneg abv (a x))) + _ < P := hP (max N M + 1) + + rw [h₁, h₂, h₃, sum_mul, ← sub_sub, sub_right_comm, sub_self, zero_sub, abv_neg abv] + refine' lt_of_le_of_lt (abv_sum_le_sum_abv _ _) _ + suffices + (∑ i in range (max N M + 1), + abv (a i) * abv ((∑ k in range (K - i), b k) - ∑ k in range K, b k)) + + ((∑ i in range K, abv (a i) * abv ((∑ k in range (K - i), b k) - ∑ k in range K, b k)) - + ∑ i in range (max N M + 1), + abv (a i) * abv ((∑ k in range (K - i), b k) - ∑ k in range K, b k)) < + ε / (2 * P) * P + ε / (4 * Q) * (2 * Q) + by + rw [hε] at this + simpa [abv_mul abv] + refine' + add_lt_add + (lt_of_le_of_lt hsumlesum + (by rw [← sum_mul, mul_comm] <;> exact (mul_lt_mul_left hPε0).mpr hsumltP)) + _ + rw [sum_range_sub_sum_range (le_of_lt hNMK)] + calc + (∑ i in (range K).filterₓ fun k => max N M + 1 ≤ k, + abv (a i) * abv ((∑ k in range (K - i), b k) - ∑ k in range K, b k)) ≤ + ∑ i in (range K).filterₓ fun k => max N M + 1 ≤ k, abv (a i) * (2 * Q) := + sum_le_sum fun n hn => + by + refine' mul_le_mul_of_nonneg_left _ (abv_nonneg _ _) + rw [sub_eq_add_neg] + refine' le_trans (abv_add _ _ _) _ + rw [two_mul, abv_neg abv] + exact add_le_add (le_of_lt (hQ _)) (le_of_lt (hQ _)) + _ < ε / (4 * Q) * (2 * Q) := by + rw [← sum_mul, ← sum_range_sub_sum_range (le_of_lt hNMK)] <;> + refine' + (mul_lt_mul_right <| by + rw [two_mul] <;> + exact + add_pos (lt_of_le_of_lt (abv_nonneg _ _) (hQ 0)) + (lt_of_le_of_lt (abv_nonneg _ _) (hQ 0))).2 + (lt_of_le_of_lt (le_abs_self _) + (hM _ (le_trans (Nat.le_succ_of_le (le_max_right _ _)) (le_of_lt hNMK)) _ + (Nat.le_succ_of_le (le_max_right _ _)))) + ⟩ +#align cauchy_product cauchy_product + +end + +end NoArchimedean + +end + +open Finset + +open CauSeq + +namespace Complex + +theorem is_cau_abs_exp (z : ℂ) : IsCauSeq Abs.abs fun n => ∑ m in range n, abs (z ^ m / m !) := + let ⟨n, hn⟩ := exists_nat_gt (abs z) + have hn0 : (0 : ℝ) < n := lt_of_le_of_lt (abs.NonNeg _) hn + series_ratio_test n (Complex.abs z / n) (div_nonneg (abs.NonNeg _) (le_of_lt hn0)) + (by rwa [div_lt_iff hn0, one_mul]) fun m hm => by + rw [abs_abs, abs_abs, Nat.factorial_succ, pow_succ, mul_comm m.succ, Nat.cast_mul, ← div_div, + mul_div_assoc, mul_div_right_comm, abs.map_mul, map_div₀, abs_cast_nat] <;> + exact + mul_le_mul_of_nonneg_right + (div_le_div_of_le_left (abs.nonneg _) hn0 (Nat.cast_le.2 (le_trans hm (Nat.le_succ _)))) + (abs.nonneg _) +#align complex.is_cau_abs_exp Complex.is_cau_abs_exp + +noncomputable section + +theorem is_cau_exp (z : ℂ) : IsCauSeq abs fun n => ∑ m in range n, z ^ m / m ! := + is_cau_series_of_abv_cau (is_cau_abs_exp z) +#align complex.is_cau_exp Complex.is_cau_exp + +/-- The Cauchy sequence consisting of partial sums of the Taylor series of +the complex exponential function -/ +@[pp_nodot] +def exp' (z : ℂ) : CauSeq ℂ Complex.abs := + ⟨fun n => ∑ m in range n, z ^ m / m !, is_cau_exp z⟩ +#align complex.exp' Complex.exp' + +/-- The complex exponential function, defined via its Taylor series -/ +@[pp_nodot] +irreducible_def exp (z : ℂ) : ℂ := + limUnder (exp' z) +#align complex.exp Complex.exp + +/-- The complex sine function, defined via `exp` -/ +@[pp_nodot] +def sin (z : ℂ) : ℂ := + (exp (-z * i) - exp (z * i)) * i / 2 +#align complex.sin Complex.sin + +/-- The complex cosine function, defined via `exp` -/ +@[pp_nodot] +def cos (z : ℂ) : ℂ := + (exp (z * i) + exp (-z * i)) / 2 +#align complex.cos Complex.cos + +/-- The complex tangent function, defined as `sin z / cos z` -/ +@[pp_nodot] +def tan (z : ℂ) : ℂ := + sin z / cos z +#align complex.tan Complex.tan + +/-- The complex hyperbolic sine function, defined via `exp` -/ +@[pp_nodot] +def sinh (z : ℂ) : ℂ := + (exp z - exp (-z)) / 2 +#align complex.sinh Complex.sinh + +/-- The complex hyperbolic cosine function, defined via `exp` -/ +@[pp_nodot] +def cosh (z : ℂ) : ℂ := + (exp z + exp (-z)) / 2 +#align complex.cosh Complex.cosh + +/-- The complex hyperbolic tangent function, defined as `sinh z / cosh z` -/ +@[pp_nodot] +def tanh (z : ℂ) : ℂ := + sinh z / cosh z +#align complex.tanh Complex.tanh + +end Complex + +namespace Real + +open Complex + +/-- The real exponential function, defined as the real part of the complex exponential -/ +@[pp_nodot] +def exp (x : ℝ) : ℝ := + (exp x).re +#align real.exp Real.exp + +/-- The real sine function, defined as the real part of the complex sine -/ +@[pp_nodot] +def sin (x : ℝ) : ℝ := + (sin x).re +#align real.sin Real.sin + +/-- The real cosine function, defined as the real part of the complex cosine -/ +@[pp_nodot] +def cos (x : ℝ) : ℝ := + (cos x).re +#align real.cos Real.cos + +/-- The real tangent function, defined as the real part of the complex tangent -/ +@[pp_nodot] +def tan (x : ℝ) : ℝ := + (tan x).re +#align real.tan Real.tan + +/-- The real hypebolic sine function, defined as the real part of the complex hyperbolic sine -/ +@[pp_nodot] +def sinh (x : ℝ) : ℝ := + (sinh x).re +#align real.sinh Real.sinh + +/-- The real hypebolic cosine function, defined as the real part of the complex hyperbolic cosine -/ +@[pp_nodot] +def cosh (x : ℝ) : ℝ := + (cosh x).re +#align real.cosh Real.cosh + +/-- The real hypebolic tangent function, defined as the real part of +the complex hyperbolic tangent -/ +@[pp_nodot] +def tanh (x : ℝ) : ℝ := + (tanh x).re +#align real.tanh Real.tanh + +end Real + +namespace Complex + +variable (x y : ℂ) + +@[simp] +theorem exp_zero : exp 0 = 1 := by + rw [exp] + refine' lim_eq_of_equiv_const fun ε ε0 => ⟨1, fun j hj => _⟩ + convert ε0 + cases j + · exact absurd hj (not_le_of_gt zero_lt_one) + · dsimp [exp'] + induction' j with j ih + · dsimp [exp'] <;> simp + · rw [← ih (by decide)] + simp only [sum_range_succ, pow_succ] + simp +#align complex.exp_zero Complex.exp_zero + +theorem exp_add : exp (x + y) = exp x * exp y := + by + 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)!) := + by + intro j + refine' Finset.sum_congr rfl fun m hm => _ + rw [add_pow, div_eq_mul_inv, sum_mul] + refine' Finset.sum_congr rfl fun 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) +#align complex.exp_add Complex.exp_add + +theorem exp_list_sum (l : List ℂ) : exp l.Sum = (l.map exp).Prod := + @MonoidHom.map_list_prod (Multiplicative ℂ) ℂ _ _ ⟨exp, exp_zero, exp_add⟩ l +#align complex.exp_list_sum Complex.exp_list_sum + +theorem exp_multiset_sum (s : Multiset ℂ) : exp s.Sum = (s.map exp).Prod := + @MonoidHom.map_multiset_prod (Multiplicative ℂ) ℂ _ _ ⟨exp, exp_zero, exp_add⟩ s +#align complex.exp_multiset_sum Complex.exp_multiset_sum + +theorem exp_sum {α : Type _} (s : Finset α) (f : α → ℂ) : + exp (∑ x in s, f x) = ∏ x in s, exp (f x) := + @MonoidHom.map_prod (Multiplicative ℂ) α ℂ _ _ ⟨exp, exp_zero, exp_add⟩ f s +#align complex.exp_sum Complex.exp_sum + +theorem exp_nat_mul (x : ℂ) : ∀ n : ℕ, exp (n * x) = exp x ^ n + | 0 => by rw [Nat.cast_zero, zero_mul, exp_zero, pow_zero] + | Nat.succ n => by rw [pow_succ', Nat.cast_add_one, add_mul, exp_add, ← exp_nat_mul, one_mul] +#align complex.exp_nat_mul Complex.exp_nat_mul + +theorem exp_ne_zero : exp x ≠ 0 := fun h => + zero_ne_one <| by rw [← exp_zero, ← add_neg_self x, exp_add, h] <;> simp +#align complex.exp_ne_zero Complex.exp_ne_zero + +theorem exp_neg : exp (-x) = (exp x)⁻¹ := by + rw [← mul_right_inj' (exp_ne_zero x), ← exp_add] <;> simp [mul_inv_cancel (exp_ne_zero x)] +#align complex.exp_neg Complex.exp_neg + +theorem exp_sub : exp (x - y) = exp x / exp y := by + simp [sub_eq_add_neg, exp_add, exp_neg, div_eq_mul_inv] +#align complex.exp_sub Complex.exp_sub + +theorem exp_int_mul (z : ℂ) (n : ℤ) : Complex.exp (n * z) = Complex.exp z ^ n := + by + cases n + · apply Complex.exp_nat_mul + · simpa [Complex.exp_neg, add_comm, ← neg_mul] using Complex.exp_nat_mul (-z) (1 + n) +#align complex.exp_int_mul Complex.exp_int_mul + +@[simp] +theorem exp_conj : exp (conj x) = conj (exp x) := + by + dsimp [exp] + rw [← lim_conj] + refine' congr_arg limUnder (CauSeq.ext fun _ => _) + dsimp [exp', Function.comp, cau_seq_conj] + rw [(starRingEnd _).map_sum] + refine' sum_congr rfl fun n hn => _ + rw [map_div₀, map_pow, ← of_real_nat_cast, conj_of_real] +#align complex.exp_conj Complex.exp_conj + +@[simp] +theorem of_real_exp_of_real_re (x : ℝ) : ((exp x).re : ℂ) = exp x := + eq_conj_iff_re.1 <| by rw [← exp_conj, conj_of_real] +#align complex.of_real_exp_of_real_re Complex.of_real_exp_of_real_re + +@[simp, norm_cast] +theorem of_real_exp (x : ℝ) : (Real.exp x : ℂ) = exp x := + of_real_exp_of_real_re _ +#align complex.of_real_exp Complex.of_real_exp + +@[simp] +theorem exp_of_real_im (x : ℝ) : (exp x).im = 0 := by rw [← of_real_exp_of_real_re, of_real_im] +#align complex.exp_of_real_im Complex.exp_of_real_im + +theorem exp_of_real_re (x : ℝ) : (exp x).re = Real.exp x := + rfl +#align complex.exp_of_real_re Complex.exp_of_real_re + +theorem two_sinh : 2 * sinh x = exp x - exp (-x) := + mul_div_cancel' _ two_ne_zero +#align complex.two_sinh Complex.two_sinh + +theorem two_cosh : 2 * cosh x = exp x + exp (-x) := + mul_div_cancel' _ two_ne_zero +#align complex.two_cosh Complex.two_cosh + +@[simp] +theorem sinh_zero : sinh 0 = 0 := by simp [sinh] +#align complex.sinh_zero Complex.sinh_zero + +@[simp] +theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] +#align complex.sinh_neg Complex.sinh_neg + +private theorem sinh_add_aux {a b c d : ℂ} : + (a - b) * (c + d) + (a + b) * (c - d) = 2 * (a * c - b * d) := by ring +#align complex.sinh_add_aux complex.sinh_add_aux + +theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := + by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← + mul_assoc, two_sinh, mul_left_comm, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, + mul_left_comm, two_cosh, ← mul_assoc, two_cosh] + exact sinh_add_aux +#align complex.sinh_add Complex.sinh_add + +@[simp] +theorem cosh_zero : cosh 0 = 1 := by simp [cosh] +#align complex.cosh_zero Complex.cosh_zero + +@[simp] +theorem cosh_neg : cosh (-x) = cosh x := by simp [add_comm, cosh, exp_neg] +#align complex.cosh_neg Complex.cosh_neg + +private theorem cosh_add_aux {a b c d : ℂ} : + (a + b) * (c + d) + (a - b) * (c - d) = 2 * (a * c + b * d) := by ring +#align complex.cosh_add_aux complex.cosh_add_aux + +theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := + by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← + mul_assoc, two_cosh, ← mul_assoc, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, + mul_left_comm, two_cosh, mul_left_comm, two_sinh] + exact cosh_add_aux +#align complex.cosh_add Complex.cosh_add + +theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by + simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] +#align complex.sinh_sub Complex.sinh_sub + +theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by + simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] +#align complex.cosh_sub Complex.cosh_sub + +theorem sinh_conj : sinh (conj x) = conj (sinh x) := by + rw [sinh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_sub, sinh, map_div₀, conj_bit0, + RingHom.map_one] +#align complex.sinh_conj Complex.sinh_conj + +@[simp] +theorem of_real_sinh_of_real_re (x : ℝ) : ((sinh x).re : ℂ) = sinh x := + eq_conj_iff_re.1 <| by rw [← sinh_conj, conj_of_real] +#align complex.of_real_sinh_of_real_re Complex.of_real_sinh_of_real_re + +@[simp, norm_cast] +theorem of_real_sinh (x : ℝ) : (Real.sinh x : ℂ) = sinh x := + of_real_sinh_of_real_re _ +#align complex.of_real_sinh Complex.of_real_sinh + +@[simp] +theorem sinh_of_real_im (x : ℝ) : (sinh x).im = 0 := by rw [← of_real_sinh_of_real_re, of_real_im] +#align complex.sinh_of_real_im Complex.sinh_of_real_im + +theorem sinh_of_real_re (x : ℝ) : (sinh x).re = Real.sinh x := + rfl +#align complex.sinh_of_real_re Complex.sinh_of_real_re + +theorem cosh_conj : cosh (conj x) = conj (cosh x) := by + rw [cosh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_add, cosh, map_div₀, conj_bit0, + RingHom.map_one] +#align complex.cosh_conj Complex.cosh_conj + +theorem of_real_cosh_of_real_re (x : ℝ) : ((cosh x).re : ℂ) = cosh x := + eq_conj_iff_re.1 <| by rw [← cosh_conj, conj_of_real] +#align complex.of_real_cosh_of_real_re Complex.of_real_cosh_of_real_re + +@[simp, norm_cast] +theorem of_real_cosh (x : ℝ) : (Real.cosh x : ℂ) = cosh x := + of_real_cosh_of_real_re _ +#align complex.of_real_cosh Complex.of_real_cosh + +@[simp] +theorem cosh_of_real_im (x : ℝ) : (cosh x).im = 0 := by rw [← of_real_cosh_of_real_re, of_real_im] +#align complex.cosh_of_real_im Complex.cosh_of_real_im + +@[simp] +theorem cosh_of_real_re (x : ℝ) : (cosh x).re = Real.cosh x := + rfl +#align complex.cosh_of_real_re Complex.cosh_of_real_re + +theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := + rfl +#align complex.tanh_eq_sinh_div_cosh Complex.tanh_eq_sinh_div_cosh + +@[simp] +theorem tanh_zero : tanh 0 = 0 := by simp [tanh] +#align complex.tanh_zero Complex.tanh_zero + +@[simp] +theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] +#align complex.tanh_neg Complex.tanh_neg + +theorem tanh_conj : tanh (conj x) = conj (tanh x) := by + rw [tanh, sinh_conj, cosh_conj, ← map_div₀, tanh] +#align complex.tanh_conj Complex.tanh_conj + +@[simp] +theorem of_real_tanh_of_real_re (x : ℝ) : ((tanh x).re : ℂ) = tanh x := + eq_conj_iff_re.1 <| by rw [← tanh_conj, conj_of_real] +#align complex.of_real_tanh_of_real_re Complex.of_real_tanh_of_real_re + +@[simp, norm_cast] +theorem of_real_tanh (x : ℝ) : (Real.tanh x : ℂ) = tanh x := + of_real_tanh_of_real_re _ +#align complex.of_real_tanh Complex.of_real_tanh + +@[simp] +theorem tanh_of_real_im (x : ℝ) : (tanh x).im = 0 := by rw [← of_real_tanh_of_real_re, of_real_im] +#align complex.tanh_of_real_im Complex.tanh_of_real_im + +theorem tanh_of_real_re (x : ℝ) : (tanh x).re = Real.tanh x := + rfl +#align complex.tanh_of_real_re Complex.tanh_of_real_re + +@[simp] +theorem cosh_add_sinh : cosh x + sinh x = exp x := by + rw [← mul_right_inj' (two_ne_zero' ℂ), mul_add, two_cosh, two_sinh, add_add_sub_cancel, two_mul] +#align complex.cosh_add_sinh Complex.cosh_add_sinh + +@[simp] +theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] +#align complex.sinh_add_cosh Complex.sinh_add_cosh + +@[simp] +theorem exp_sub_cosh : exp x - cosh x = sinh x := + sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm +#align complex.exp_sub_cosh Complex.exp_sub_cosh + +@[simp] +theorem exp_sub_sinh : exp x - sinh x = cosh x := + sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm +#align complex.exp_sub_sinh Complex.exp_sub_sinh + +@[simp] +theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by + rw [← mul_right_inj' (two_ne_zero' ℂ), mul_sub, two_cosh, two_sinh, add_sub_sub_cancel, two_mul] +#align complex.cosh_sub_sinh Complex.cosh_sub_sinh + +@[simp] +theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] +#align complex.sinh_sub_cosh Complex.sinh_sub_cosh + +@[simp] +theorem cosh_sq_sub_sinh_sq : cosh x ^ 2 - sinh x ^ 2 = 1 := by + rw [sq_sub_sq, cosh_add_sinh, cosh_sub_sinh, ← exp_add, add_neg_self, exp_zero] +#align complex.cosh_sq_sub_sinh_sq Complex.cosh_sq_sub_sinh_sq + +theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := + by + rw [← cosh_sq_sub_sinh_sq x] + ring +#align complex.cosh_sq Complex.cosh_sq + +theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := + by + rw [← cosh_sq_sub_sinh_sq x] + ring +#align complex.sinh_sq Complex.sinh_sq + +theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by rw [two_mul, cosh_add, sq, sq] +#align complex.cosh_two_mul Complex.cosh_two_mul + +theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := + by + rw [two_mul, sinh_add] + ring +#align complex.sinh_two_mul Complex.sinh_two_mul + +theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := + by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, cosh_add x (2 * x)] + simp only [cosh_two_mul, sinh_two_mul] + have h2 : sinh x * (2 * sinh x * cosh x) = 2 * cosh x * sinh x ^ 2 := by ring + rw [h2, sinh_sq] + ring +#align complex.cosh_three_mul Complex.cosh_three_mul + +theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := + by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, sinh_add x (2 * x)] + simp only [cosh_two_mul, sinh_two_mul] + have h2 : cosh x * (2 * sinh x * cosh x) = 2 * sinh x * cosh x ^ 2 := by ring + rw [h2, cosh_sq] + ring +#align complex.sinh_three_mul Complex.sinh_three_mul + +@[simp] +theorem sin_zero : sin 0 = 0 := by simp [sin] +#align complex.sin_zero Complex.sin_zero + +@[simp] +theorem sin_neg : sin (-x) = -sin x := by + simp [sin, sub_eq_add_neg, exp_neg, (neg_div _ _).symm, add_mul] +#align complex.sin_neg Complex.sin_neg + +theorem two_sin : 2 * sin x = (exp (-x * i) - exp (x * i)) * i := + mul_div_cancel' _ two_ne_zero +#align complex.two_sin Complex.two_sin + +theorem two_cos : 2 * cos x = exp (x * i) + exp (-x * i) := + mul_div_cancel' _ two_ne_zero +#align complex.two_cos Complex.two_cos + +theorem sinh_mul_i : sinh (x * i) = sin x * i := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, ← mul_assoc, two_sin, mul_assoc, I_mul_I, + mul_neg_one, neg_sub, neg_mul_eq_neg_mul] +#align complex.sinh_mul_I Complex.sinh_mul_i + +theorem cosh_mul_i : cosh (x * i) = cos x := by + rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, two_cos, neg_mul_eq_neg_mul] +#align complex.cosh_mul_I Complex.cosh_mul_i + +theorem tanh_mul_i : tanh (x * i) = tan x * i := by + rw [tanh_eq_sinh_div_cosh, cosh_mul_I, sinh_mul_I, mul_div_right_comm, tan] +#align complex.tanh_mul_I Complex.tanh_mul_i + +theorem cos_mul_i : cos (x * i) = cosh x := by rw [← cosh_mul_I] <;> ring_nf <;> simp +#align complex.cos_mul_I Complex.cos_mul_i + +theorem sin_mul_i : sin (x * i) = sinh x * i := + by + have h : i * sin (x * i) = -sinh x := + by + rw [mul_comm, ← sinh_mul_I] + ring_nf + simp + simpa only [neg_mul, div_I, neg_neg] using CancelFactors.cancel_factors_eq_div h I_ne_zero +#align complex.sin_mul_I Complex.sin_mul_i + +theorem tan_mul_i : tan (x * i) = tanh x * i := by + rw [tan, sin_mul_I, cos_mul_I, mul_div_right_comm, tanh_eq_sinh_div_cosh] +#align complex.tan_mul_I Complex.tan_mul_i + +theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := by + rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, add_mul, add_mul, mul_right_comm, ← sinh_mul_I, + mul_assoc, ← sinh_mul_I, ← cosh_mul_I, ← cosh_mul_I, sinh_add] +#align complex.sin_add Complex.sin_add + +@[simp] +theorem cos_zero : cos 0 = 1 := by simp [cos] +#align complex.cos_zero Complex.cos_zero + +@[simp] +theorem cos_neg : cos (-x) = cos x := by simp [cos, sub_eq_add_neg, exp_neg, add_comm] +#align complex.cos_neg Complex.cos_neg + +private theorem cos_add_aux {a b c d : ℂ} : + (a + b) * (c + d) - (b - a) * (d - c) * -1 = 2 * (a * c + b * d) := by ring +#align complex.cos_add_aux complex.cos_add_aux + +theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by + rw [← cosh_mul_I, add_mul, cosh_add, cosh_mul_I, cosh_mul_I, sinh_mul_I, sinh_mul_I, + mul_mul_mul_comm, I_mul_I, mul_neg_one, sub_eq_add_neg] +#align complex.cos_add Complex.cos_add + +theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by + simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] +#align complex.sin_sub Complex.sin_sub + +theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by + simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] +#align complex.cos_sub Complex.cos_sub + +theorem sin_add_mul_i (x y : ℂ) : sin (x + y * i) = sin x * cosh y + cos x * sinh y * i := by + rw [sin_add, cos_mul_I, sin_mul_I, mul_assoc] +#align complex.sin_add_mul_I Complex.sin_add_mul_i + +theorem sin_eq (z : ℂ) : sin z = sin z.re * cosh z.im + cos z.re * sinh z.im * i := by + convert sin_add_mul_I z.re z.im <;> exact (re_add_im z).symm +#align complex.sin_eq Complex.sin_eq + +theorem cos_add_mul_i (x y : ℂ) : cos (x + y * i) = cos x * cosh y - sin x * sinh y * i := by + rw [cos_add, cos_mul_I, sin_mul_I, mul_assoc] +#align complex.cos_add_mul_I Complex.cos_add_mul_i + +theorem cos_eq (z : ℂ) : cos z = cos z.re * cosh z.im - sin z.re * sinh z.im * i := by + convert cos_add_mul_I z.re z.im <;> exact (re_add_im z).symm +#align complex.cos_eq Complex.cos_eq + +theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := + by + have s1 := sin_add ((x + y) / 2) ((x - y) / 2) + have s2 := sin_sub ((x + y) / 2) ((x - y) / 2) + rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel, half_add_self] at s1 + rw [div_sub_div_same, ← sub_add, add_sub_cancel', half_add_self] at s2 + rw [s1, s2] + ring +#align complex.sin_sub_sin Complex.sin_sub_sin + +theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := + by + have s1 := cos_add ((x + y) / 2) ((x - y) / 2) + have s2 := cos_sub ((x + y) / 2) ((x - y) / 2) + rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel, half_add_self] at s1 + rw [div_sub_div_same, ← sub_add, add_sub_cancel', half_add_self] at s2 + rw [s1, s2] + ring +#align complex.cos_sub_cos Complex.cos_sub_cos + +theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := + by + have h2 : (2 : ℂ) ≠ 0 := by norm_num + calc + cos x + cos y = cos ((x + y) / 2 + (x - y) / 2) + cos ((x + y) / 2 - (x - y) / 2) := _ + _ = + cos ((x + y) / 2) * cos ((x - y) / 2) - sin ((x + y) / 2) * sin ((x - y) / 2) + + (cos ((x + y) / 2) * cos ((x - y) / 2) + sin ((x + y) / 2) * sin ((x - y) / 2)) := + _ + _ = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := _ + + · congr <;> field_simp [h2] <;> ring + · rw [cos_add, cos_sub] + ring +#align complex.cos_add_cos Complex.cos_add_cos + +theorem sin_conj : sin (conj x) = conj (sin x) := by + rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← RingHom.map_mul, + sinh_conj, mul_neg, sinh_neg, sinh_mul_I, mul_neg] +#align complex.sin_conj Complex.sin_conj + +@[simp] +theorem of_real_sin_of_real_re (x : ℝ) : ((sin x).re : ℂ) = sin x := + eq_conj_iff_re.1 <| by rw [← sin_conj, conj_of_real] +#align complex.of_real_sin_of_real_re Complex.of_real_sin_of_real_re + +@[simp, norm_cast] +theorem of_real_sin (x : ℝ) : (Real.sin x : ℂ) = sin x := + of_real_sin_of_real_re _ +#align complex.of_real_sin Complex.of_real_sin + +@[simp] +theorem sin_of_real_im (x : ℝ) : (sin x).im = 0 := by rw [← of_real_sin_of_real_re, of_real_im] +#align complex.sin_of_real_im Complex.sin_of_real_im + +theorem sin_of_real_re (x : ℝ) : (sin x).re = Real.sin x := + rfl +#align complex.sin_of_real_re Complex.sin_of_real_re + +theorem cos_conj : cos (conj x) = conj (cos x) := by + rw [← cosh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← cosh_mul_I, cosh_conj, mul_neg, cosh_neg] +#align complex.cos_conj Complex.cos_conj + +@[simp] +theorem of_real_cos_of_real_re (x : ℝ) : ((cos x).re : ℂ) = cos x := + eq_conj_iff_re.1 <| by rw [← cos_conj, conj_of_real] +#align complex.of_real_cos_of_real_re Complex.of_real_cos_of_real_re + +@[simp, norm_cast] +theorem of_real_cos (x : ℝ) : (Real.cos x : ℂ) = cos x := + of_real_cos_of_real_re _ +#align complex.of_real_cos Complex.of_real_cos + +@[simp] +theorem cos_of_real_im (x : ℝ) : (cos x).im = 0 := by rw [← of_real_cos_of_real_re, of_real_im] +#align complex.cos_of_real_im Complex.cos_of_real_im + +theorem cos_of_real_re (x : ℝ) : (cos x).re = Real.cos x := + rfl +#align complex.cos_of_real_re Complex.cos_of_real_re + +@[simp] +theorem tan_zero : tan 0 = 0 := by simp [tan] +#align complex.tan_zero Complex.tan_zero + +theorem tan_eq_sin_div_cos : tan x = sin x / cos x := + rfl +#align complex.tan_eq_sin_div_cos Complex.tan_eq_sin_div_cos + +theorem tan_mul_cos {x : ℂ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by + rw [tan_eq_sin_div_cos, div_mul_cancel _ hx] +#align complex.tan_mul_cos Complex.tan_mul_cos + +@[simp] +theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] +#align complex.tan_neg Complex.tan_neg + +theorem tan_conj : tan (conj x) = conj (tan x) := by rw [tan, sin_conj, cos_conj, ← map_div₀, tan] +#align complex.tan_conj Complex.tan_conj + +@[simp] +theorem of_real_tan_of_real_re (x : ℝ) : ((tan x).re : ℂ) = tan x := + eq_conj_iff_re.1 <| by rw [← tan_conj, conj_of_real] +#align complex.of_real_tan_of_real_re Complex.of_real_tan_of_real_re + +@[simp, norm_cast] +theorem of_real_tan (x : ℝ) : (Real.tan x : ℂ) = tan x := + of_real_tan_of_real_re _ +#align complex.of_real_tan Complex.of_real_tan + +@[simp] +theorem tan_of_real_im (x : ℝ) : (tan x).im = 0 := by rw [← of_real_tan_of_real_re, of_real_im] +#align complex.tan_of_real_im Complex.tan_of_real_im + +theorem tan_of_real_re (x : ℝ) : (tan x).re = Real.tan x := + rfl +#align complex.tan_of_real_re Complex.tan_of_real_re + +theorem cos_add_sin_i : cos x + sin x * i = exp (x * i) := by + rw [← cosh_add_sinh, sinh_mul_I, cosh_mul_I] +#align complex.cos_add_sin_I Complex.cos_add_sin_i + +theorem cos_sub_sin_i : cos x - sin x * i = exp (-x * i) := by + rw [neg_mul, ← cosh_sub_sinh, sinh_mul_I, cosh_mul_I] +#align complex.cos_sub_sin_I Complex.cos_sub_sin_i + +@[simp] +theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := + Eq.trans (by rw [cosh_mul_I, sinh_mul_I, mul_pow, I_sq, mul_neg_one, sub_neg_eq_add, add_comm]) + (cosh_sq_sub_sinh_sq (x * i)) +#align complex.sin_sq_add_cos_sq Complex.sin_sq_add_cos_sq + +@[simp] +theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] +#align complex.cos_sq_add_sin_sq Complex.cos_sq_add_sin_sq + +theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := by rw [two_mul, cos_add, ← sq, ← sq] +#align complex.cos_two_mul' Complex.cos_two_mul' + +theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := by + rw [cos_two_mul', eq_sub_iff_add_eq.2 (sin_sq_add_cos_sq x), ← sub_add, sub_add_eq_add_sub, + two_mul] +#align complex.cos_two_mul Complex.cos_two_mul + +theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := by + rw [two_mul, sin_add, two_mul, add_mul, mul_comm] +#align complex.sin_two_mul Complex.sin_two_mul + +theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := by + simp [cos_two_mul, div_add_div_same, mul_div_cancel_left, two_ne_zero, -one_div] +#align complex.cos_sq Complex.cos_sq + +theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel'] +#align complex.cos_sq' Complex.cos_sq' + +theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel] +#align complex.sin_sq Complex.sin_sq + +theorem inv_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := + by + have : cos x ^ 2 ≠ 0 := pow_ne_zero 2 hx + rw [tan_eq_sin_div_cos, div_pow] + field_simp [this] +#align complex.inv_one_add_tan_sq Complex.inv_one_add_tan_sq + +theorem tan_sq_div_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : + tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by + simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] +#align complex.tan_sq_div_one_add_tan_sq Complex.tan_sq_div_one_add_tan_sq + +theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := + by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, cos_add x (2 * x)] + simp only [cos_two_mul, sin_two_mul, mul_add, mul_sub, mul_one, sq] + have h2 : 4 * cos x ^ 3 = 2 * cos x * cos x * cos x + 2 * cos x * cos x ^ 2 := by ring + rw [h2, cos_sq'] + ring +#align complex.cos_three_mul Complex.cos_three_mul + +theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := + by + have h1 : x + 2 * x = 3 * x := by ring + rw [← h1, sin_add x (2 * x)] + simp only [cos_two_mul, sin_two_mul, cos_sq'] + have h2 : cos x * (2 * sin x * cos x) = 2 * sin x * cos x ^ 2 := by ring + rw [h2, cos_sq'] + ring +#align complex.sin_three_mul Complex.sin_three_mul + +theorem exp_mul_i : exp (x * i) = cos x + sin x * i := + (cos_add_sin_i _).symm +#align complex.exp_mul_I Complex.exp_mul_i + +theorem exp_add_mul_i : exp (x + y * i) = exp x * (cos y + sin y * i) := by rw [exp_add, exp_mul_I] +#align complex.exp_add_mul_I Complex.exp_add_mul_i + +theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im * i) := by + rw [← exp_add_mul_I, re_add_im] +#align complex.exp_eq_exp_re_mul_sin_add_cos Complex.exp_eq_exp_re_mul_sin_add_cos + +theorem exp_re : (exp x).re = Real.exp x.re * Real.cos x.im := + by + rw [exp_eq_exp_re_mul_sin_add_cos] + simp [exp_of_real_re, cos_of_real_re] +#align complex.exp_re Complex.exp_re + +theorem exp_im : (exp x).im = Real.exp x.re * Real.sin x.im := + by + rw [exp_eq_exp_re_mul_sin_add_cos] + simp [exp_of_real_re, sin_of_real_re] +#align complex.exp_im Complex.exp_im + +@[simp] +theorem exp_of_real_mul_i_re (x : ℝ) : (exp (x * i)).re = Real.cos x := by + simp [exp_mul_I, cos_of_real_re] +#align complex.exp_of_real_mul_I_re Complex.exp_of_real_mul_i_re + +@[simp] +theorem exp_of_real_mul_i_im (x : ℝ) : (exp (x * i)).im = Real.sin x := by + simp [exp_mul_I, sin_of_real_re] +#align complex.exp_of_real_mul_I_im Complex.exp_of_real_mul_i_im + +/-- **De Moivre's formula** -/ +theorem cos_add_sin_mul_i_pow (n : ℕ) (z : ℂ) : + (cos z + sin z * i) ^ n = cos (↑n * z) + sin (↑n * z) * i := + by + rw [← exp_mul_I, ← exp_mul_I] + induction' n with n ih + · rw [pow_zero, Nat.cast_zero, zero_mul, zero_mul, exp_zero] + · rw [pow_succ', ih, Nat.cast_succ, add_mul, add_mul, one_mul, exp_add] +#align complex.cos_add_sin_mul_I_pow Complex.cos_add_sin_mul_i_pow + +end Complex + +namespace Real + +open Complex + +variable (x y : ℝ) + +@[simp] +theorem exp_zero : exp 0 = 1 := by simp [Real.exp] +#align real.exp_zero Real.exp_zero + +theorem exp_add : exp (x + y) = exp x * exp y := by simp [exp_add, exp] +#align real.exp_add Real.exp_add + +theorem exp_list_sum (l : List ℝ) : exp l.Sum = (l.map exp).Prod := + @MonoidHom.map_list_prod (Multiplicative ℝ) ℝ _ _ ⟨exp, exp_zero, exp_add⟩ l +#align real.exp_list_sum Real.exp_list_sum + +theorem exp_multiset_sum (s : Multiset ℝ) : exp s.Sum = (s.map exp).Prod := + @MonoidHom.map_multiset_prod (Multiplicative ℝ) ℝ _ _ ⟨exp, exp_zero, exp_add⟩ s +#align real.exp_multiset_sum Real.exp_multiset_sum + +theorem exp_sum {α : Type _} (s : Finset α) (f : α → ℝ) : + exp (∑ x in s, f x) = ∏ x in s, exp (f x) := + @MonoidHom.map_prod (Multiplicative ℝ) α ℝ _ _ ⟨exp, exp_zero, exp_add⟩ f s +#align real.exp_sum Real.exp_sum + +theorem exp_nat_mul (x : ℝ) : ∀ n : ℕ, exp (n * x) = exp x ^ n + | 0 => by rw [Nat.cast_zero, zero_mul, exp_zero, pow_zero] + | Nat.succ n => by rw [pow_succ', Nat.cast_add_one, add_mul, exp_add, ← exp_nat_mul, one_mul] +#align real.exp_nat_mul Real.exp_nat_mul + +theorem exp_ne_zero : exp x ≠ 0 := fun h => + exp_ne_zero x <| by rw [exp, ← of_real_inj] at h <;> simp_all +#align real.exp_ne_zero Real.exp_ne_zero + +theorem exp_neg : exp (-x) = (exp x)⁻¹ := by + rw [← of_real_inj, exp, of_real_exp_of_real_re, of_real_neg, exp_neg, of_real_inv, of_real_exp] +#align real.exp_neg Real.exp_neg + +theorem exp_sub : exp (x - y) = exp x / exp y := by + simp [sub_eq_add_neg, exp_add, exp_neg, div_eq_mul_inv] +#align real.exp_sub Real.exp_sub + +@[simp] +theorem sin_zero : sin 0 = 0 := by simp [sin] +#align real.sin_zero Real.sin_zero + +@[simp] +theorem sin_neg : sin (-x) = -sin x := by simp [sin, exp_neg, (neg_div _ _).symm, add_mul] +#align real.sin_neg Real.sin_neg + +theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := by + rw [← of_real_inj] <;> simp [sin, sin_add] +#align real.sin_add Real.sin_add + +@[simp] +theorem cos_zero : cos 0 = 1 := by simp [cos] +#align real.cos_zero Real.cos_zero + +@[simp] +theorem cos_neg : cos (-x) = cos x := by simp [cos, exp_neg] +#align real.cos_neg Real.cos_neg + +@[simp] +theorem cos_abs : cos (|x|) = cos x := by + cases le_total x 0 <;> simp only [*, _root_.abs_of_nonneg, abs_of_nonpos, cos_neg] +#align real.cos_abs Real.cos_abs + +theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by + rw [← of_real_inj] <;> simp [cos, cos_add] +#align real.cos_add Real.cos_add + +theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by + simp [sub_eq_add_neg, sin_add, sin_neg, cos_neg] +#align real.sin_sub Real.sin_sub + +theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by + simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] +#align real.cos_sub Real.cos_sub + +theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := + by + rw [← of_real_inj] + simp only [sin, cos, of_real_sin_of_real_re, of_real_sub, of_real_add, of_real_div, of_real_mul, + of_real_one, of_real_bit0] + convert sin_sub_sin _ _ <;> norm_cast +#align real.sin_sub_sin Real.sin_sub_sin + +theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := + by + rw [← of_real_inj] + simp only [cos, neg_mul, of_real_sin, of_real_sub, of_real_add, of_real_cos_of_real_re, + of_real_div, of_real_mul, of_real_one, of_real_neg, of_real_bit0] + convert cos_sub_cos _ _ + ring +#align real.cos_sub_cos Real.cos_sub_cos + +theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := + by + rw [← of_real_inj] + simp only [cos, of_real_sub, of_real_add, of_real_cos_of_real_re, of_real_div, of_real_mul, + of_real_one, of_real_bit0] + convert cos_add_cos _ _ <;> norm_cast +#align real.cos_add_cos Real.cos_add_cos + +theorem tan_eq_sin_div_cos : tan x = sin x / cos x := by + rw [← of_real_inj, of_real_tan, tan_eq_sin_div_cos, of_real_div, of_real_sin, of_real_cos] +#align real.tan_eq_sin_div_cos Real.tan_eq_sin_div_cos + +theorem tan_mul_cos {x : ℝ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by + rw [tan_eq_sin_div_cos, div_mul_cancel _ hx] +#align real.tan_mul_cos Real.tan_mul_cos + +@[simp] +theorem tan_zero : tan 0 = 0 := by simp [tan] +#align real.tan_zero Real.tan_zero + +@[simp] +theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] +#align real.tan_neg Real.tan_neg + +@[simp] +theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := + of_real_inj.1 <| by simp +#align real.sin_sq_add_cos_sq Real.sin_sq_add_cos_sq + +@[simp] +theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq_add_cos_sq] +#align real.cos_sq_add_sin_sq Real.cos_sq_add_sin_sq + +theorem sin_sq_le_one : sin x ^ 2 ≤ 1 := by + rw [← sin_sq_add_cos_sq x] <;> exact le_add_of_nonneg_right (sq_nonneg _) +#align real.sin_sq_le_one Real.sin_sq_le_one + +theorem cos_sq_le_one : cos x ^ 2 ≤ 1 := by + rw [← sin_sq_add_cos_sq x] <;> exact le_add_of_nonneg_left (sq_nonneg _) +#align real.cos_sq_le_one Real.cos_sq_le_one + +theorem abs_sin_le_one : |sin x| ≤ 1 := + abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, sin_sq_le_one] +#align real.abs_sin_le_one Real.abs_sin_le_one + +theorem abs_cos_le_one : |cos x| ≤ 1 := + abs_le_one_iff_mul_self_le_one.2 <| by simp only [← sq, cos_sq_le_one] +#align real.abs_cos_le_one Real.abs_cos_le_one + +theorem sin_le_one : sin x ≤ 1 := + (abs_le.1 (abs_sin_le_one _)).2 +#align real.sin_le_one Real.sin_le_one + +theorem cos_le_one : cos x ≤ 1 := + (abs_le.1 (abs_cos_le_one _)).2 +#align real.cos_le_one Real.cos_le_one + +theorem neg_one_le_sin : -1 ≤ sin x := + (abs_le.1 (abs_sin_le_one _)).1 +#align real.neg_one_le_sin Real.neg_one_le_sin + +theorem neg_one_le_cos : -1 ≤ cos x := + (abs_le.1 (abs_cos_le_one _)).1 +#align real.neg_one_le_cos Real.neg_one_le_cos + +theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := by + rw [← of_real_inj] <;> simp [cos_two_mul] +#align real.cos_two_mul Real.cos_two_mul + +theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := by + rw [← of_real_inj] <;> simp [cos_two_mul'] +#align real.cos_two_mul' Real.cos_two_mul' + +theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := by + rw [← of_real_inj] <;> simp [sin_two_mul] +#align real.sin_two_mul Real.sin_two_mul + +theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := + of_real_inj.1 <| by simpa using cos_sq x +#align real.cos_sq Real.cos_sq + +theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel'] +#align real.cos_sq' Real.cos_sq' + +theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := + eq_sub_iff_add_eq.2 <| sin_sq_add_cos_sq _ +#align real.sin_sq Real.sin_sq + +theorem abs_sin_eq_sqrt_one_sub_cos_sq (x : ℝ) : |sin x| = sqrt (1 - cos x ^ 2) := by + rw [← sin_sq, sqrt_sq_eq_abs] +#align real.abs_sin_eq_sqrt_one_sub_cos_sq Real.abs_sin_eq_sqrt_one_sub_cos_sq + +theorem abs_cos_eq_sqrt_one_sub_sin_sq (x : ℝ) : |cos x| = sqrt (1 - sin x ^ 2) := by + rw [← cos_sq', sqrt_sq_eq_abs] +#align real.abs_cos_eq_sqrt_one_sub_sin_sq Real.abs_cos_eq_sqrt_one_sub_sin_sq + +theorem inv_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := + have : Complex.cos x ≠ 0 := mt (congr_arg re) hx + of_real_inj.1 <| by simpa using Complex.inv_one_add_tan_sq this +#align real.inv_one_add_tan_sq Real.inv_one_add_tan_sq + +theorem tan_sq_div_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : + tan x ^ 2 / (1 + tan x ^ 2) = sin x ^ 2 := by + simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] +#align real.tan_sq_div_one_add_tan_sq Real.tan_sq_div_one_add_tan_sq + +theorem inv_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : (sqrt (1 + tan x ^ 2))⁻¹ = cos x := by + rw [← sqrt_sq hx.le, ← sqrt_inv, inv_one_add_tan_sq hx.ne'] +#align real.inv_sqrt_one_add_tan_sq Real.inv_sqrt_one_add_tan_sq + +theorem tan_div_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : + tan x / sqrt (1 + tan x ^ 2) = sin x := by + rw [← tan_mul_cos hx.ne', ← inv_sqrt_one_add_tan_sq hx, div_eq_mul_inv] +#align real.tan_div_sqrt_one_add_tan_sq Real.tan_div_sqrt_one_add_tan_sq + +theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by + rw [← of_real_inj] <;> simp [cos_three_mul] +#align real.cos_three_mul Real.cos_three_mul + +theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by + rw [← of_real_inj] <;> simp [sin_three_mul] +#align real.sin_three_mul Real.sin_three_mul + +/-- The definition of `sinh` in terms of `exp`. -/ +theorem sinh_eq (x : ℝ) : sinh x = (exp x - exp (-x)) / 2 := + eq_div_of_mul_eq two_ne_zero <| by + rw [sinh, exp, exp, Complex.of_real_neg, Complex.sinh, mul_two, ← Complex.add_re, ← mul_two, + div_mul_cancel _ (two_ne_zero' ℂ), Complex.sub_re] +#align real.sinh_eq Real.sinh_eq + +@[simp] +theorem sinh_zero : sinh 0 = 0 := by simp [sinh] +#align real.sinh_zero Real.sinh_zero + +@[simp] +theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] +#align real.sinh_neg Real.sinh_neg + +theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by + rw [← of_real_inj] <;> simp [sinh_add] +#align real.sinh_add Real.sinh_add + +/-- The definition of `cosh` in terms of `exp`. -/ +theorem cosh_eq (x : ℝ) : cosh x = (exp x + exp (-x)) / 2 := + eq_div_of_mul_eq two_ne_zero <| by + rw [cosh, exp, exp, Complex.of_real_neg, Complex.cosh, mul_two, ← Complex.add_re, ← mul_two, + div_mul_cancel _ (two_ne_zero' ℂ), Complex.add_re] +#align real.cosh_eq Real.cosh_eq + +@[simp] +theorem cosh_zero : cosh 0 = 1 := by simp [cosh] +#align real.cosh_zero Real.cosh_zero + +@[simp] +theorem cosh_neg : cosh (-x) = cosh x := + of_real_inj.1 <| by simp +#align real.cosh_neg Real.cosh_neg + +@[simp] +theorem cosh_abs : cosh (|x|) = cosh x := by + cases le_total x 0 <;> simp [*, _root_.abs_of_nonneg, abs_of_nonpos] +#align real.cosh_abs Real.cosh_abs + +theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by + rw [← of_real_inj] <;> simp [cosh_add] +#align real.cosh_add Real.cosh_add + +theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by + simp [sub_eq_add_neg, sinh_add, sinh_neg, cosh_neg] +#align real.sinh_sub Real.sinh_sub + +theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by + simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] +#align real.cosh_sub Real.cosh_sub + +theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := + of_real_inj.1 <| by simp [tanh_eq_sinh_div_cosh] +#align real.tanh_eq_sinh_div_cosh Real.tanh_eq_sinh_div_cosh + +@[simp] +theorem tanh_zero : tanh 0 = 0 := by simp [tanh] +#align real.tanh_zero Real.tanh_zero + +@[simp] +theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] +#align real.tanh_neg Real.tanh_neg + +@[simp] +theorem cosh_add_sinh : cosh x + sinh x = exp x := by rw [← of_real_inj] <;> simp +#align real.cosh_add_sinh Real.cosh_add_sinh + +@[simp] +theorem sinh_add_cosh : sinh x + cosh x = exp x := by rw [add_comm, cosh_add_sinh] +#align real.sinh_add_cosh Real.sinh_add_cosh + +@[simp] +theorem exp_sub_cosh : exp x - cosh x = sinh x := + sub_eq_iff_eq_add.2 (sinh_add_cosh x).symm +#align real.exp_sub_cosh Real.exp_sub_cosh + +@[simp] +theorem exp_sub_sinh : exp x - sinh x = cosh x := + sub_eq_iff_eq_add.2 (cosh_add_sinh x).symm +#align real.exp_sub_sinh Real.exp_sub_sinh + +@[simp] +theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := + by + rw [← of_real_inj] + simp +#align real.cosh_sub_sinh Real.cosh_sub_sinh + +@[simp] +theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_sub_sinh] +#align real.sinh_sub_cosh Real.sinh_sub_cosh + +@[simp] +theorem cosh_sq_sub_sinh_sq (x : ℝ) : cosh x ^ 2 - sinh x ^ 2 = 1 := by rw [← of_real_inj] <;> simp +#align real.cosh_sq_sub_sinh_sq Real.cosh_sq_sub_sinh_sq + +theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by rw [← of_real_inj] <;> simp [cosh_sq] +#align real.cosh_sq Real.cosh_sq + +theorem cosh_sq' : cosh x ^ 2 = 1 + sinh x ^ 2 := + (cosh_sq x).trans (add_comm _ _) +#align real.cosh_sq' Real.cosh_sq' + +theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by rw [← of_real_inj] <;> simp [sinh_sq] +#align real.sinh_sq Real.sinh_sq + +theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by + rw [← of_real_inj] <;> simp [cosh_two_mul] +#align real.cosh_two_mul Real.cosh_two_mul + +theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by + rw [← of_real_inj] <;> simp [sinh_two_mul] +#align real.sinh_two_mul Real.sinh_two_mul + +theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by + rw [← of_real_inj] <;> simp [cosh_three_mul] +#align real.cosh_three_mul Real.cosh_three_mul + +theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by + rw [← of_real_inj] <;> simp [sinh_three_mul] +#align real.sinh_three_mul Real.sinh_three_mul + +open IsAbsoluteValue + +/-- This is an intermediate result that is later replaced by `real.add_one_le_exp`; use that lemma +instead. -/ +theorem add_one_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : x + 1 ≤ exp x := + calc + x + 1 ≤ limUnder (⟨fun n : ℕ => ((exp' x) n).re, isCauSeq_re (exp' x)⟩ : CauSeq ℝ Abs.abs) := + le_lim + (CauSeq.le_of_exists + ⟨2, fun j hj => + show x + (1 : ℝ) ≤ (∑ m in range j, (x ^ m / m ! : ℂ)).re + by + have h₁ : (((fun m : ℕ => (x ^ m / m ! : ℂ)) ∘ Nat.succ) 0).re = x := by simp + have h₂ : ((x : ℂ) ^ 0 / 0!).re = 1 := by simp + rw [← tsub_add_cancel_of_le hj, sum_range_succ', sum_range_succ', add_re, add_re, h₁, + h₂, add_assoc, ← coe_re_add_group_hom, re_add_group_hom.map_sum, + coe_re_add_group_hom] + refine' le_add_of_nonneg_of_le (sum_nonneg fun m hm => _) le_rfl + rw [← of_real_pow, ← of_real_nat_cast, ← of_real_div, of_real_re] + exact div_nonneg (pow_nonneg hx _) (Nat.cast_nonneg _)⟩) + _ = exp x := by rw [exp, Complex.exp, ← cau_seq_re, lim_re] + +#align real.add_one_le_exp_of_nonneg Real.add_one_le_exp_of_nonneg + +theorem one_le_exp {x : ℝ} (hx : 0 ≤ x) : 1 ≤ exp x := by linarith [add_one_le_exp_of_nonneg hx] +#align real.one_le_exp Real.one_le_exp + +theorem exp_pos (x : ℝ) : 0 < exp x := + (le_total 0 x).elim (lt_of_lt_of_le zero_lt_one ∘ one_le_exp) fun h => by + rw [← neg_neg x, Real.exp_neg] <;> + exact inv_pos.2 (lt_of_lt_of_le zero_lt_one (one_le_exp (neg_nonneg.2 h))) +#align real.exp_pos Real.exp_pos + +@[simp] +theorem abs_exp (x : ℝ) : |exp x| = exp x := + abs_of_pos (exp_pos _) +#align real.abs_exp Real.abs_exp + +@[mono] +theorem exp_strictMono : StrictMono exp := fun x y h => by + rw [← sub_add_cancel y x, Real.exp_add] <;> + exact + (lt_mul_iff_one_lt_left (exp_pos _)).2 + (lt_of_lt_of_le (by linarith) (add_one_le_exp_of_nonneg (by linarith))) +#align real.exp_strict_mono Real.exp_strictMono + +@[mono] +theorem exp_monotone : Monotone exp := + exp_strictMono.Monotone +#align real.exp_monotone Real.exp_monotone + +@[simp] +theorem exp_lt_exp {x y : ℝ} : exp x < exp y ↔ x < y := + exp_strictMono.lt_iff_lt +#align real.exp_lt_exp Real.exp_lt_exp + +@[simp] +theorem exp_le_exp {x y : ℝ} : exp x ≤ exp y ↔ x ≤ y := + exp_strictMono.le_iff_le +#align real.exp_le_exp Real.exp_le_exp + +theorem exp_injective : Function.Injective exp := + exp_strictMono.Injective +#align real.exp_injective Real.exp_injective + +@[simp] +theorem exp_eq_exp {x y : ℝ} : exp x = exp y ↔ x = y := + exp_injective.eq_iff +#align real.exp_eq_exp Real.exp_eq_exp + +@[simp] +theorem exp_eq_one_iff : exp x = 1 ↔ x = 0 := + exp_injective.eq_iff' exp_zero +#align real.exp_eq_one_iff Real.exp_eq_one_iff + +@[simp] +theorem one_lt_exp_iff {x : ℝ} : 1 < exp x ↔ 0 < x := by rw [← exp_zero, exp_lt_exp] +#align real.one_lt_exp_iff Real.one_lt_exp_iff + +@[simp] +theorem exp_lt_one_iff {x : ℝ} : exp x < 1 ↔ x < 0 := by rw [← exp_zero, exp_lt_exp] +#align real.exp_lt_one_iff Real.exp_lt_one_iff + +@[simp] +theorem exp_le_one_iff {x : ℝ} : exp x ≤ 1 ↔ x ≤ 0 := + exp_zero ▸ exp_le_exp +#align real.exp_le_one_iff Real.exp_le_one_iff + +@[simp] +theorem one_le_exp_iff {x : ℝ} : 1 ≤ exp x ↔ 0 ≤ x := + exp_zero ▸ exp_le_exp +#align real.one_le_exp_iff Real.one_le_exp_iff + +/-- `real.cosh` is always positive -/ +theorem cosh_pos (x : ℝ) : 0 < Real.cosh x := + (cosh_eq x).symm ▸ half_pos (add_pos (exp_pos x) (exp_pos (-x))) +#align real.cosh_pos Real.cosh_pos + +theorem sinh_lt_cosh : sinh x < cosh x := + lt_of_pow_lt_pow 2 (cosh_pos _).le <| (cosh_sq x).symm ▸ lt_add_one _ +#align real.sinh_lt_cosh Real.sinh_lt_cosh + +end Real + +namespace Complex + +theorem sum_div_factorial_le {α : Type _} [LinearOrderedField α] (n j : ℕ) (hn : 0 < n) : + (∑ m in Filter (fun k => n ≤ k) (range j), (1 / m ! : α)) ≤ n.succ / (n ! * n) := + calc + (∑ m in Filter (fun k => n ≤ k) (range j), (1 / m ! : α)) = + ∑ m in range (j - n), 1 / (m + n)! := + sum_bij (fun m _ => m - n) + (fun m hm => + mem_range.2 <| + (tsub_lt_tsub_iff_right (by simp at hm <;> tauto)).2 (by simp at hm <;> tauto)) + (fun m hm => by rw [tsub_add_cancel_of_le] <;> simp at * <;> tauto) + (fun a₁ a₂ ha₁ ha₂ h => by + rwa [tsub_eq_iff_eq_add_of_le, tsub_add_eq_add_tsub, eq_comm, tsub_eq_iff_eq_add_of_le, + add_left_inj, eq_comm] at h <;> + simp at * <;> + tauto) + fun b hb => + ⟨b + n, + mem_filter.2 ⟨mem_range.2 <| lt_tsub_iff_right.mp (mem_range.1 hb), Nat.le_add_left _ _⟩, + by rw [add_tsub_cancel_right]⟩ + _ ≤ ∑ m in range (j - n), (n ! * n.succ ^ m)⁻¹ := + by + refine' sum_le_sum fun m n => _ + rw [one_div, inv_le_inv] + · rw [← Nat.cast_pow, ← Nat.cast_mul, Nat.cast_le, add_comm] + exact Nat.factorial_mul_pow_le_factorial + · exact Nat.cast_pos.2 (Nat.factorial_pos _) + · + exact + mul_pos (Nat.cast_pos.2 (Nat.factorial_pos _)) + (pow_pos (Nat.cast_pos.2 (Nat.succ_pos _)) _) + _ = n !⁻¹ * ∑ m in range (j - n), n.succ⁻¹ ^ m := by + simp [mul_inv, mul_sum.symm, sum_mul.symm, -Nat.factorial_succ, mul_comm, inv_pow] + _ = (n.succ - n.succ * n.succ⁻¹ ^ (j - n)) / (n ! * n) := + by + have h₁ : (n.succ : α) ≠ 1 := + @Nat.cast_one α _ ▸ mt Nat.cast_inj.1 (mt Nat.succ.inj (pos_iff_ne_zero.1 hn)) + have h₂ : (n.succ : α) ≠ 0 := Nat.cast_ne_zero.2 (Nat.succ_ne_zero _) + have h₃ : (n ! * n : α) ≠ 0 := + mul_ne_zero (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 (Nat.factorial_pos _))) + (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 hn)) + have h₄ : (n.succ - 1 : α) = n := by simp + rw [geom_sum_inv h₁ h₂, eq_div_iff_mul_eq h₃, mul_comm _ (n ! * n : α), ← + mul_assoc (n !⁻¹ : α), ← mul_inv_rev, h₄, ← mul_assoc (n ! * n : α), mul_comm (n : α) n !, + mul_inv_cancel h₃] <;> + simp [mul_add, add_mul, mul_assoc, mul_comm] + _ ≤ n.succ / (n ! * n) := + by + refine' Iff.mpr (div_le_div_right (mul_pos _ _)) _ + exact Nat.cast_pos.2 (Nat.factorial_pos _) + exact Nat.cast_pos.2 hn + exact + sub_le_self _ + (mul_nonneg (Nat.cast_nonneg _) (pow_nonneg (inv_nonneg.2 (Nat.cast_nonneg _)) _)) + +#align complex.sum_div_factorial_le Complex.sum_div_factorial_le + +theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : + abs (exp x - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n * (n.succ * (n ! * n)⁻¹) := + by + rw [← lim_const (∑ m in range n, _), exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] + refine' lim_le (CauSeq.le_of_exists ⟨n, fun j hj => _⟩) + simp_rw [← sub_eq_add_neg] + show + abs ((∑ m in range j, x ^ m / m !) - ∑ m in range n, x ^ m / m !) ≤ + abs x ^ n * (n.succ * (n ! * n)⁻¹) + rw [sum_range_sub_sum_range hj] + calc + abs (∑ m in (range j).filterₓ fun k => n ≤ k, (x ^ m / m ! : ℂ)) = + abs (∑ m in (range j).filterₓ fun k => n ≤ k, (x ^ n * (x ^ (m - n) / m !) : ℂ)) := + by + refine' congr_arg abs (sum_congr rfl fun m hm => _) + rw [mem_filter, mem_range] at hm + rw [← mul_div_assoc, ← pow_add, add_tsub_cancel_of_le hm.2] + _ ≤ ∑ m in Filter (fun k => n ≤ k) (range j), abs (x ^ n * (_ / m !)) := + (abv_sum_le_sum_abv _ _) + _ ≤ ∑ m in Filter (fun k => n ≤ k) (range j), abs x ^ n * (1 / m !) := + by + refine' sum_le_sum fun m hm => _ + rw [map_mul, map_pow, map_div₀, abs_cast_nat] + refine' mul_le_mul_of_nonneg_left ((div_le_div_right _).2 _) _ + · exact Nat.cast_pos.2 (Nat.factorial_pos _) + · rw [abv_pow abs] + exact pow_le_one _ (abs.nonneg _) hx + · exact pow_nonneg (abs.nonneg _) _ + _ = abs x ^ n * ∑ m in (range j).filterₓ fun k => n ≤ k, (1 / m ! : ℝ) := by + simp [abs_mul, abv_pow abs, abs_div, mul_sum.symm] + _ ≤ abs x ^ n * (n.succ * (n ! * n)⁻¹) := + mul_le_mul_of_nonneg_left (sum_div_factorial_le _ _ hn) (pow_nonneg (abs.nonneg _) _) + +#align complex.exp_bound Complex.exp_bound + +theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : + abs (exp x - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n / n ! * 2 := + by + rw [← lim_const (∑ m in range n, _), exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] + refine' lim_le (CauSeq.le_of_exists ⟨n, fun j hj => _⟩) + simp_rw [← sub_eq_add_neg] + show abs ((∑ m in range j, x ^ m / m !) - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n / n ! * 2 + let k := j - n + have hj : j = n + k := (add_tsub_cancel_of_le hj).symm + rw [hj, sum_range_add_sub_sum_range] + calc + abs (∑ i : ℕ in range k, x ^ (n + i) / ((n + i)! : ℂ)) ≤ + ∑ i : ℕ in range k, abs (x ^ (n + i) / ((n + i)! : ℂ)) := + abv_sum_le_sum_abv _ _ + _ ≤ ∑ i : ℕ in range k, abs x ^ (n + i) / (n + i)! := by + simp only [Complex.abs_cast_nat, map_div₀, abv_pow abs] + _ ≤ ∑ i : ℕ in range k, abs x ^ (n + i) / (n ! * n.succ ^ i) := _ + _ = ∑ i : ℕ in range k, abs x ^ n / n ! * (abs x ^ i / n.succ ^ i) := _ + _ ≤ abs x ^ n / ↑n ! * 2 := _ + + · refine' sum_le_sum fun m hm => div_le_div (pow_nonneg (abs.nonneg x) (n + m)) le_rfl _ _ + · exact_mod_cast mul_pos n.factorial_pos (pow_pos n.succ_pos _) + · exact_mod_cast Nat.factorial_mul_pow_le_factorial + · refine' Finset.sum_congr rfl fun _ _ => _ + simp only [pow_add, div_eq_inv_mul, mul_inv, mul_left_comm, mul_assoc] + · rw [← mul_sum] + apply mul_le_mul_of_nonneg_left + · simp_rw [← div_pow] + rw [geom_sum_eq, div_le_iff_of_neg] + · trans (-1 : ℝ) + · linarith + · simp only [neg_le_sub_iff_le_add, div_pow, Nat.cast_succ, le_add_iff_nonneg_left] + exact + div_nonneg (pow_nonneg (abs.nonneg x) k) + (pow_nonneg (add_nonneg n.cast_nonneg zero_le_one) k) + · linarith + · linarith + · exact div_nonneg (pow_nonneg (abs.nonneg x) n) (Nat.cast_nonneg n !) +#align complex.exp_bound' Complex.exp_bound' + +theorem abs_exp_sub_one_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1) ≤ 2 * abs x := + calc + abs (exp x - 1) = abs (exp x - ∑ m in range 1, x ^ m / m !) := by simp [sum_range_succ] + _ ≤ abs x ^ 1 * (Nat.succ 1 * (1! * (1 : ℕ))⁻¹) := (exp_bound hx (by decide)) + _ = 2 * abs x := by simp [two_mul, mul_two, mul_add, mul_comm] + +#align complex.abs_exp_sub_one_le Complex.abs_exp_sub_one_le + +theorem abs_exp_sub_one_sub_id_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1 - x) ≤ abs x ^ 2 := + calc + abs (exp x - 1 - x) = abs (exp x - ∑ m in range 2, x ^ m / m !) := by + simp [sub_eq_add_neg, sum_range_succ_comm, add_assoc] + _ ≤ abs x ^ 2 * (Nat.succ 2 * (2! * (2 : ℕ))⁻¹) := (exp_bound hx (by decide)) + _ ≤ abs x ^ 2 * 1 := (mul_le_mul_of_nonneg_left (by norm_num) (sq_nonneg (abs x))) + _ = abs x ^ 2 := by rw [mul_one] + +#align complex.abs_exp_sub_one_sub_id_le Complex.abs_exp_sub_one_sub_id_le + +end Complex + +namespace Real + +open Complex Finset + +theorem exp_bound {x : ℝ} (hx : |x| ≤ 1) {n : ℕ} (hn : 0 < n) : + |exp x - ∑ m in range n, x ^ m / m !| ≤ |x| ^ n * (n.succ / (n ! * n)) := + by + have hxc : Complex.abs x ≤ 1 := by exact_mod_cast hx + convert exp_bound hxc hn <;> norm_cast +#align real.exp_bound Real.exp_bound + +theorem exp_bound' {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) {n : ℕ} (hn : 0 < n) : + Real.exp x ≤ (∑ m in Finset.range n, x ^ m / m !) + x ^ n * (n + 1) / (n ! * n) := + by + have h3 : |x| = x := by simpa + have h4 : |x| ≤ 1 := by rwa [h3] + have h' := Real.exp_bound h4 hn + rw [h3] at h' + have h'' := (abs_sub_le_iff.1 h').1 + have t := sub_le_iff_le_add'.1 h'' + simpa [mul_div_assoc] using t +#align real.exp_bound' Real.exp_bound' + +theorem abs_exp_sub_one_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1| ≤ 2 * |x| := + by + have : Complex.abs x ≤ 1 := by exact_mod_cast hx + exact_mod_cast Complex.abs_exp_sub_one_le this +#align real.abs_exp_sub_one_le Real.abs_exp_sub_one_le + +theorem abs_exp_sub_one_sub_id_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1 - x| ≤ x ^ 2 := + by + rw [← _root_.sq_abs] + have : Complex.abs x ≤ 1 := by exact_mod_cast hx + exact_mod_cast Complex.abs_exp_sub_one_sub_id_le this +#align real.abs_exp_sub_one_sub_id_le Real.abs_exp_sub_one_sub_id_le + +/-- A finite initial segment of the exponential series, followed by an arbitrary tail. +For fixed `n` this is just a linear map wrt `r`, and each map is a simple linear function +of the previous (see `exp_near_succ`), with `exp_near n x r ⟶ exp x` as `n ⟶ ∞`, +for any `r`. -/ +def expNear (n : ℕ) (x r : ℝ) : ℝ := + (∑ m in range n, x ^ m / m !) + x ^ n / n ! * r +#align real.exp_near Real.expNear + +@[simp] +theorem expNear_zero (x r) : expNear 0 x r = r := by simp [exp_near] +#align real.exp_near_zero Real.expNear_zero + +@[simp] +theorem expNear_succ (n x r) : expNear (n + 1) x r = expNear n x (1 + x / (n + 1) * r) := by + simp [exp_near, range_succ, mul_add, add_left_comm, add_assoc, pow_succ, div_eq_mul_inv, + mul_inv] <;> + ac_rfl +#align real.exp_near_succ Real.expNear_succ + +theorem expNear_sub (n x r₁ r₂) : expNear n x r₁ - expNear n x r₂ = x ^ n / n ! * (r₁ - r₂) := by + simp [exp_near, mul_sub] +#align real.exp_near_sub Real.expNear_sub + +theorem exp_approx_end (n m : ℕ) (x : ℝ) (e₁ : n + 1 = m) (h : |x| ≤ 1) : + |exp x - expNear m x 0| ≤ |x| ^ m / m ! * ((m + 1) / m) := + by + simp [exp_near] + convert exp_bound h _ using 1 + field_simp [mul_comm] + linarith +#align real.exp_approx_end Real.exp_approx_end + +theorem exp_approx_succ {n} {x a₁ b₁ : ℝ} (m : ℕ) (e₁ : n + 1 = m) (a₂ b₂ : ℝ) + (e : |1 + x / m * a₂ - a₁| ≤ b₁ - |x| / m * b₂) + (h : |exp x - expNear m x a₂| ≤ |x| ^ m / m ! * b₂) : + |exp x - expNear n x a₁| ≤ |x| ^ n / n ! * b₁ := + by + refine' (_root_.abs_sub_le _ _ _).trans ((add_le_add_right h _).trans _) + subst e₁; rw [exp_near_succ, exp_near_sub, _root_.abs_mul] + convert mul_le_mul_of_nonneg_left (le_sub_iff_add_le'.1 e) _ + · simp [mul_add, pow_succ', div_eq_mul_inv, _root_.abs_mul, _root_.abs_inv, ← pow_abs, mul_inv] + ac_rfl + · simp [_root_.div_nonneg, _root_.abs_nonneg] +#align real.exp_approx_succ Real.exp_approx_succ + +theorem exp_approx_end' {n} {x a b : ℝ} (m : ℕ) (e₁ : n + 1 = m) (rm : ℝ) (er : ↑m = rm) + (h : |x| ≤ 1) (e : |1 - a| ≤ b - |x| / rm * ((rm + 1) / rm)) : + |exp x - expNear n x a| ≤ |x| ^ n / n ! * b := by + subst er <;> exact exp_approx_succ _ e₁ _ _ (by simpa using e) (exp_approx_end _ _ _ e₁ h) +#align real.exp_approx_end' Real.exp_approx_end' + +theorem exp_1_approx_succ_eq {n} {a₁ b₁ : ℝ} {m : ℕ} (en : n + 1 = m) {rm : ℝ} (er : ↑m = rm) + (h : |exp 1 - expNear m 1 ((a₁ - 1) * rm)| ≤ |1| ^ m / m ! * (b₁ * rm)) : + |exp 1 - expNear n 1 a₁| ≤ |1| ^ n / n ! * b₁ := + by + subst er + refine' exp_approx_succ _ en _ _ _ h + field_simp [show (m : ℝ) ≠ 0 by norm_cast <;> linarith] +#align real.exp_1_approx_succ_eq Real.exp_1_approx_succ_eq + +theorem exp_approx_start (x a b : ℝ) (h : |exp x - expNear 0 x a| ≤ |x| ^ 0 / 0! * b) : + |exp x - a| ≤ b := by simpa using h +#align real.exp_approx_start Real.exp_approx_start + +theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| ^ 4 * (5 / 96) := + calc + |cos x - (1 - x ^ 2 / 2)| = abs (Complex.cos x - (1 - x ^ 2 / 2)) := by + rw [← abs_of_real] <;> simp [of_real_bit0, of_real_one, of_real_inv] + _ = abs ((Complex.exp (x * i) + Complex.exp (-x * i) - (2 - x ^ 2)) / 2) := by + simp [Complex.cos, sub_div, add_div, neg_div, div_self (two_ne_zero' ℂ)] + _ = + abs + (((Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) + + (Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !)) / + 2) := + (congr_arg abs + (congr_arg (fun x : ℂ => x / 2) + (by + simp only [sum_range_succ] + simp [pow_succ] + apply Complex.ext <;> simp [div_eq_mul_inv, norm_sq] <;> ring))) + _ ≤ + abs ((Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) / 2) + + abs ((Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) / 2) := + by rw [add_div] <;> exact complex.abs.add_le _ _ + _ = + abs (Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) / 2 + + abs (Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) / 2 := + by simp [map_div₀] + _ ≤ + Complex.abs (x * i) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + + Complex.abs (-x * i) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 := + (add_le_add ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide))) + ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide)))) + _ ≤ |x| ^ 4 * (5 / 96) := by + norm_num <;> simp [mul_assoc, mul_comm, mul_left_comm, mul_div_assoc] + +#align real.cos_bound Real.cos_bound + +theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| ^ 4 * (5 / 96) := + calc + |sin x - (x - x ^ 3 / 6)| = abs (Complex.sin x - (x - x ^ 3 / 6)) := by + rw [← abs_of_real] <;> simp [of_real_bit0, of_real_one, of_real_inv] + _ = abs (((Complex.exp (-x * i) - Complex.exp (x * i)) * i - (2 * x - x ^ 3 / 3)) / 2) := by + simp [Complex.sin, sub_div, add_div, neg_div, mul_div_cancel_left _ (two_ne_zero' ℂ), div_div, + show (3 : ℂ) * 2 = 6 by norm_num] + _ = + abs + (((Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) - + (Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !)) * + i / + 2) := + (congr_arg abs + (congr_arg (fun x : ℂ => x / 2) + (by + simp only [sum_range_succ] + simp [pow_succ] + apply Complex.ext <;> simp [div_eq_mul_inv, norm_sq] <;> ring))) + _ ≤ + abs ((Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) * i / 2) + + abs (-((Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) * i) / 2) := + by rw [sub_mul, sub_eq_add_neg, add_div] <;> exact complex.abs.add_le _ _ + _ = + abs (Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) / 2 + + abs (Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) / 2 := + by simp [add_comm, map_div₀] + _ ≤ + Complex.abs (x * i) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + + Complex.abs (-x * i) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 := + (add_le_add ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide))) + ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide)))) + _ ≤ |x| ^ 4 * (5 / 96) := by + norm_num <;> simp [mul_assoc, mul_comm, mul_left_comm, mul_div_assoc] + +#align real.sin_bound Real.sin_bound + +theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := + calc + 0 < 1 - x ^ 2 / 2 - |x| ^ 4 * (5 / 96) := + sub_pos.2 <| + lt_sub_iff_add_lt.2 + (calc + |x| ^ 4 * (5 / 96) + x ^ 2 / 2 ≤ 1 * (5 / 96) + 1 / 2 := + add_le_add (mul_le_mul_of_nonneg_right (pow_le_one _ (abs_nonneg _) hx) (by norm_num)) + ((div_le_div_right (by norm_num)).2 + (by + rw [sq, ← abs_mul_self, _root_.abs_mul] <;> + exact mul_le_one hx (abs_nonneg _) hx)) + _ < 1 := by norm_num + ) + _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 + +#align real.cos_pos_of_le_one Real.cos_pos_of_le_one + +theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < sin x := + calc + 0 < x - x ^ 3 / 6 - |x| ^ 4 * (5 / 96) := + sub_pos.2 <| + lt_sub_iff_add_lt.2 + (calc + |x| ^ 4 * (5 / 96) + x ^ 3 / 6 ≤ x * (5 / 96) + x / 6 := + add_le_add + (mul_le_mul_of_nonneg_right + (calc + |x| ^ 4 ≤ |x| ^ 1 := + pow_le_pow_of_le_one (abs_nonneg _) + (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]) (by decide) + _ = x := by simp [_root_.abs_of_nonneg (le_of_lt hx0)] + ) + (by norm_num)) + ((div_le_div_right (by norm_num)).2 + (calc + x ^ 3 ≤ x ^ 1 := pow_le_pow_of_le_one (le_of_lt hx0) hx (by decide) + _ = x := pow_one _ + )) + _ < x := by linarith + ) + _ ≤ sin x := + sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]))).2 + +#align real.sin_pos_of_pos_of_le_one Real.sin_pos_of_pos_of_le_one + +theorem sin_pos_of_pos_of_le_two {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 2) : 0 < sin x := + have : x / 2 ≤ 1 := (div_le_iff (by norm_num)).mpr (by simpa) + calc + 0 < 2 * sin (x / 2) * cos (x / 2) := + mul_pos (mul_pos (by norm_num) (sin_pos_of_pos_of_le_one (half_pos hx0) this)) + (cos_pos_of_le_one (by rwa [_root_.abs_of_nonneg (le_of_lt (half_pos hx0))])) + _ = sin x := by rw [← sin_two_mul, two_mul, add_halves] + +#align real.sin_pos_of_pos_of_le_two Real.sin_pos_of_pos_of_le_two + +theorem cos_one_le : cos 1 ≤ 2 / 3 := + calc + cos 1 ≤ |(1 : ℝ)| ^ 4 * (5 / 96) + (1 - 1 ^ 2 / 2) := + sub_le_iff_le_add.1 (abs_sub_le_iff.1 (cos_bound (by simp))).1 + _ ≤ 2 / 3 := by norm_num + +#align real.cos_one_le Real.cos_one_le + +theorem cos_one_pos : 0 < cos 1 := + cos_pos_of_le_one (le_of_eq abs_one) +#align real.cos_one_pos Real.cos_one_pos + +theorem cos_two_neg : cos 2 < 0 := + calc + cos 2 = cos (2 * 1) := congr_arg cos (mul_one _).symm + _ = _ := (Real.cos_two_mul 1) + _ ≤ 2 * (2 / 3) ^ 2 - 1 := + (sub_le_sub_right + (mul_le_mul_of_nonneg_left + (by + rw [sq, sq] + exact mul_self_le_mul_self (le_of_lt cos_one_pos) cos_one_le) + zero_le_two) + _) + _ < 0 := by norm_num + +#align real.cos_two_neg Real.cos_two_neg + +theorem exp_bound_div_one_sub_of_interval_approx {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) : + (∑ j : ℕ in Finset.range 3, x ^ j / j.factorial) + + x ^ 3 * ((3 : ℕ) + 1) / ((3 : ℕ).factorial * (3 : ℕ)) ≤ + ∑ j in Finset.range 3, x ^ j := + by + norm_num [Finset.sum] + rw [add_assoc, add_comm (x + 1) (x ^ 3 * 4 / 18), ← add_assoc, add_le_add_iff_right, ← + add_le_add_iff_left (-(x ^ 2 / 2)), ← add_assoc, CommRing.add_left_neg (x ^ 2 / 2), zero_add, + neg_add_eq_sub, sub_half, sq, pow_succ, sq] + have i1 : x * 4 / 18 ≤ 1 / 2 := by linarith + have i2 : 0 ≤ x * 4 / 18 := by linarith + have i3 := mul_le_mul h1 h1 le_rfl h1 + rw [zero_mul] at i3 + have t := mul_le_mul le_rfl i1 i2 i3 + rw [← mul_assoc] + rwa [mul_one_div, ← mul_div_assoc, ← mul_assoc] at t +#align real.exp_bound_div_one_sub_of_interval_approx Real.exp_bound_div_one_sub_of_interval_approx + +theorem exp_bound_div_one_sub_of_interval {x : ℝ} (h1 : 0 ≤ x) (h2 : x < 1) : + Real.exp x ≤ 1 / (1 - x) := + haveI h : (∑ j in Finset.range 3, x ^ j) ≤ 1 / (1 - x) := + by + norm_num [Finset.sum] + have h1x : 0 < 1 - x := by simpa + rw [le_div_iff h1x] + norm_num [← add_assoc, mul_sub_left_distrib, mul_one, add_mul, sub_add_eq_sub_sub, + pow_succ' x 2] + have hx3 : 0 ≤ x ^ 3 := by + norm_num + exact h1 + linarith + (exp_bound' h1 h2.le <| by linarith).trans + ((exp_bound_div_one_sub_of_interval_approx h1 h2.le).trans h) +#align real.exp_bound_div_one_sub_of_interval Real.exp_bound_div_one_sub_of_interval + +theorem one_sub_le_exp_minus_of_pos {y : ℝ} (h : 0 ≤ y) : 1 - y ≤ Real.exp (-y) := + by + rw [Real.exp_neg] + have r1 : (1 - y) * Real.exp y ≤ 1 := + by + cases le_or_lt (1 - y) 0 + · have h'' : (1 - y) * y.exp ≤ 0 := by + rw [mul_nonpos_iff] + right + exact ⟨h_1, y.exp_pos.le⟩ + linarith + have hy1 : y < 1 := by linarith + rw [← le_div_iff' h_1] + exact exp_bound_div_one_sub_of_interval h hy1 + rw [inv_eq_one_div] + rw [le_div_iff' y.exp_pos] + rwa [mul_comm] at r1 +#align real.one_sub_le_exp_minus_of_pos Real.one_sub_le_exp_minus_of_pos + +theorem add_one_le_exp_of_nonpos {x : ℝ} (h : x ≤ 0) : x + 1 ≤ Real.exp x := + by + rw [add_comm] + have h1 : 0 ≤ -x := by linarith + simpa using one_sub_le_exp_minus_of_pos h1 +#align real.add_one_le_exp_of_nonpos Real.add_one_le_exp_of_nonpos + +theorem add_one_le_exp (x : ℝ) : x + 1 ≤ Real.exp x := + by + cases le_or_lt 0 x + · exact Real.add_one_le_exp_of_nonneg h + exact add_one_le_exp_of_nonpos h.le +#align real.add_one_le_exp Real.add_one_le_exp + +theorem one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t / n) ^ n ≤ exp (-t) := + by + rcases eq_or_ne n 0 with (rfl | hn) + · simp + rwa [Nat.cast_zero] at ht' + convert pow_le_pow_of_le_left _ (add_one_le_exp (-(t / n))) n + · abel + · rw [← Real.exp_nat_mul] + congr 1 + field_simp [nat.cast_ne_zero.mpr hn] + ring + · rwa [add_comm, ← sub_eq_add_neg, sub_nonneg, div_le_one] + positivity +#align real.one_sub_div_pow_le_exp_neg Real.one_sub_div_pow_le_exp_neg + +end Real + +namespace Tactic + +open Positivity Real + +/-- Extension for the `positivity` tactic: `real.exp` is always positive. -/ +@[positivity] +unsafe def positivity_exp : expr → tactic strictness + | q(Real.exp $(a)) => positive <$> mk_app `real.exp_pos [a] + | e => pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `real.exp r`" +#align tactic.positivity_exp tactic.positivity_exp + +end Tactic + +namespace Complex + +@[simp] +theorem abs_cos_add_sin_mul_i (x : ℝ) : abs (cos x + sin x * i) = 1 := + by + have := Real.sin_sq_add_cos_sq x + simp_all [add_comm, abs, norm_sq, sq, sin_of_real_re, cos_of_real_re, mul_re] +#align complex.abs_cos_add_sin_mul_I Complex.abs_cos_add_sin_mul_i + +@[simp] +theorem abs_exp_of_real (x : ℝ) : abs (exp x) = Real.exp x := by + rw [← of_real_exp] <;> exact abs_of_nonneg (le_of_lt (Real.exp_pos _)) +#align complex.abs_exp_of_real Complex.abs_exp_of_real + +@[simp] +theorem abs_exp_of_real_mul_i (x : ℝ) : abs (exp (x * i)) = 1 := by + rw [exp_mul_I, abs_cos_add_sin_mul_I] +#align complex.abs_exp_of_real_mul_I Complex.abs_exp_of_real_mul_i + +theorem abs_exp (z : ℂ) : abs (exp z) = Real.exp z.re := by + rw [exp_eq_exp_re_mul_sin_add_cos, map_mul, abs_exp_of_real, abs_cos_add_sin_mul_I, mul_one] +#align complex.abs_exp Complex.abs_exp + +theorem abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = y.re := by + rw [abs_exp, abs_exp, Real.exp_eq_exp] +#align complex.abs_exp_eq_iff_re_eq Complex.abs_exp_eq_iff_re_eq + +end Complex + From ae8ba385d4849e3bf5a6e7753d12bc9db0022946 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 13:57:29 +0000 Subject: [PATCH 03/23] automated fixes Mathbin -> Mathlib fix certain import statements move "by" to end of line add import to Mathlib.lean --- Mathlib.lean | 1 + Mathlib/Data/Complex/Exponential.lean | 138 +++++++++----------------- 2 files changed, 48 insertions(+), 91 deletions(-) diff --git a/Mathlib.lean b/Mathlib.lean index 3b2d1399f8c43..9afb3c2cfd0a3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -471,6 +471,7 @@ import Mathlib.Data.Bundle import Mathlib.Data.ByteArray import Mathlib.Data.Char import Mathlib.Data.Complex.Basic +import Mathlib.Data.Complex.Exponential import Mathlib.Data.Countable.Basic import Mathlib.Data.Countable.Defs import Mathlib.Data.Countable.Small diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index c7cf3ba3c0921..729887fea36d5 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -8,9 +8,9 @@ Authors: Chris Hughes, Abhimanyu Pallavi Sudhir ! Please do not edit these lines, except to modify the commit id ! if you have ported upstream changes. -/ -import Mathbin.Algebra.GeomSum -import Mathbin.Data.Complex.Basic -import Mathbin.Data.Nat.Choose.Sum +import Mathlib.Algebra.GeomSum +import Mathlib.Data.Complex.Basic +import Mathlib.Data.Nat.Choose.Sum /-! # Exponential, trigonometric and hyperbolic trigonometric functions @@ -38,8 +38,7 @@ variable {α : Type _} {β : Type _} [Ring β] [LinearOrderedField α] [Archimed [IsAbsoluteValue abv] theorem is_cau_of_decreasing_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) - (hnm : ∀ n ≥ m, f n.succ ≤ f n) : IsCauSeq abs f := fun ε ε0 => - by + (hnm : ∀ n ≥ m, f n.succ ≤ f n) : IsCauSeq abs f := fun ε ε0 => by let ⟨k, hk⟩ := Archimedean.arch a ε0 have h : ∃ l, ∀ n ≥ m, a - l • ε < f n := ⟨k + k + 1, fun n hnm => @@ -74,8 +73,7 @@ theorem is_cau_of_decreasing_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : #align is_cau_of_decreasing_bounded is_cau_of_decreasing_bounded theorem is_cau_of_mono_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) - (hnm : ∀ n ≥ m, f n ≤ f n.succ) : IsCauSeq abs f := - by + (hnm : ∀ n ≥ m, f n ≤ f n.succ) : IsCauSeq abs f := by refine' @Eq.recOn (ℕ → α) _ (IsCauSeq abs) _ _ (-⟨_, @is_cau_of_decreasing_bounded _ _ _ (fun n => -f n) a m (by simpa) (by simpa)⟩ : @@ -93,8 +91,7 @@ variable {α : Type _} {β : Type _} [Ring β] [LinearOrderedField α] {abv : β theorem is_cau_series_of_abv_le_cau {f : ℕ → β} {g : ℕ → α} (n : ℕ) : (∀ m, n ≤ m → abv (f m) ≤ g m) → - (IsCauSeq abs fun n => ∑ i in range n, g i) → IsCauSeq abv fun n => ∑ i in range n, f i := - by + (IsCauSeq abs fun n => ∑ i in range n, g i) → IsCauSeq abv fun n => ∑ i in range n, f i := by intro hm hg ε ε0 cases' hg (ε / 2) (div_pos ε0 (by norm_num)) with i hi exists max n i @@ -156,8 +153,7 @@ theorem is_cau_geo_series {β : Type _} [Ring β] [Nontrivial β] {abv : β → #align is_cau_geo_series is_cau_geo_series theorem is_cau_geo_series_const (a : α) {x : α} (hx1 : |x| < 1) : - IsCauSeq abs fun m => ∑ n in range m, a * x ^ n := - by + IsCauSeq abs fun m => ∑ n in range m, a * x ^ n := by have : IsCauSeq abs fun m => a * ∑ n in range m, x ^ n := (CauSeq.const abs a * ⟨_, is_cau_geo_series x hx1⟩).2 simpa only [mul_sum] @@ -193,8 +189,7 @@ theorem series_ratio_test {f : ℕ → β} (n : ℕ) (r : α) (hr0 : 0 ≤ r) (h theorem sum_range_diag_flip {α : Type _} [AddCommMonoid α] (n : ℕ) (f : ℕ → ℕ → α) : (∑ m in range n, ∑ k in range (m + 1), f k (m - k)) = - ∑ m in range n, ∑ k in range (n - m), f m k := - by + ∑ m in range n, ∑ k in range (n - m), f m k := by rw [sum_sigma', sum_sigma'] <;> exact sum_bij (fun a _ => ⟨a.2, a.1 - a.2⟩) @@ -507,8 +502,7 @@ theorem exp_zero : exp 0 = 1 := by simp #align complex.exp_zero Complex.exp_zero -theorem exp_add : exp (x + y) = exp x * exp y := - by +theorem exp_add : exp (x + y) = exp x * exp y := by have hj : ∀ j : ℕ, (∑ m in range j, (x + y) ^ m / m !) = @@ -562,16 +556,14 @@ theorem exp_sub : exp (x - y) = exp x / exp y := by simp [sub_eq_add_neg, exp_add, exp_neg, div_eq_mul_inv] #align complex.exp_sub Complex.exp_sub -theorem exp_int_mul (z : ℂ) (n : ℤ) : Complex.exp (n * z) = Complex.exp z ^ n := - by +theorem exp_int_mul (z : ℂ) (n : ℤ) : Complex.exp (n * z) = Complex.exp z ^ n := by cases n · apply Complex.exp_nat_mul · simpa [Complex.exp_neg, add_comm, ← neg_mul] using Complex.exp_nat_mul (-z) (1 + n) #align complex.exp_int_mul Complex.exp_int_mul @[simp] -theorem exp_conj : exp (conj x) = conj (exp x) := - by +theorem exp_conj : exp (conj x) = conj (exp x) := by dsimp [exp] rw [← lim_conj] refine' congr_arg limUnder (CauSeq.ext fun _ => _) @@ -619,8 +611,7 @@ private theorem sinh_add_aux {a b c d : ℂ} : (a - b) * (c + d) + (a + b) * (c - d) = 2 * (a * c - b * d) := by ring #align complex.sinh_add_aux complex.sinh_add_aux -theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := - by +theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← mul_assoc, two_sinh, mul_left_comm, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, mul_left_comm, two_cosh, ← mul_assoc, two_cosh] @@ -639,8 +630,7 @@ private theorem cosh_add_aux {a b c d : ℂ} : (a + b) * (c + d) + (a - b) * (c - d) = 2 * (a * c + b * d) := by ring #align complex.cosh_add_aux complex.cosh_add_aux -theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := - by +theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← mul_assoc, two_cosh, ← mul_assoc, two_sinh, ← mul_right_inj' (two_ne_zero' ℂ), mul_add, mul_left_comm, two_cosh, mul_left_comm, two_sinh] @@ -768,14 +758,12 @@ theorem cosh_sq_sub_sinh_sq : cosh x ^ 2 - sinh x ^ 2 = 1 := by rw [sq_sub_sq, cosh_add_sinh, cosh_sub_sinh, ← exp_add, add_neg_self, exp_zero] #align complex.cosh_sq_sub_sinh_sq Complex.cosh_sq_sub_sinh_sq -theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := - by +theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by rw [← cosh_sq_sub_sinh_sq x] ring #align complex.cosh_sq Complex.cosh_sq -theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := - by +theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by rw [← cosh_sq_sub_sinh_sq x] ring #align complex.sinh_sq Complex.sinh_sq @@ -783,14 +771,12 @@ theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by rw [two_mul, cosh_add, sq, sq] #align complex.cosh_two_mul Complex.cosh_two_mul -theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := - by +theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by rw [two_mul, sinh_add] ring #align complex.sinh_two_mul Complex.sinh_two_mul -theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := - by +theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by have h1 : x + 2 * x = 3 * x := by ring rw [← h1, cosh_add x (2 * x)] simp only [cosh_two_mul, sinh_two_mul] @@ -799,8 +785,7 @@ theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := ring #align complex.cosh_three_mul Complex.cosh_three_mul -theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := - by +theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by have h1 : x + 2 * x = 3 * x := by ring rw [← h1, sinh_add x (2 * x)] simp only [cosh_two_mul, sinh_two_mul] @@ -842,8 +827,7 @@ theorem tanh_mul_i : tanh (x * i) = tan x * i := by theorem cos_mul_i : cos (x * i) = cosh x := by rw [← cosh_mul_I] <;> ring_nf <;> simp #align complex.cos_mul_I Complex.cos_mul_i -theorem sin_mul_i : sin (x * i) = sinh x * i := - by +theorem sin_mul_i : sin (x * i) = sinh x * i := by have h : i * sin (x * i) = -sinh x := by rw [mul_comm, ← sinh_mul_I] @@ -902,8 +886,7 @@ theorem cos_eq (z : ℂ) : cos z = cos z.re * cosh z.im - sin z.re * sinh z.im * convert cos_add_mul_I z.re z.im <;> exact (re_add_im z).symm #align complex.cos_eq Complex.cos_eq -theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := - by +theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := by have s1 := sin_add ((x + y) / 2) ((x - y) / 2) have s2 := sin_sub ((x + y) / 2) ((x - y) / 2) rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel, half_add_self] at s1 @@ -912,8 +895,7 @@ theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) ring #align complex.sin_sub_sin Complex.sin_sub_sin -theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := - by +theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := by have s1 := cos_add ((x + y) / 2) ((x - y) / 2) have s2 := cos_sub ((x + y) / 2) ((x - y) / 2) rw [div_add_div_same, add_sub, add_right_comm, add_sub_cancel, half_add_self] at s1 @@ -922,8 +904,7 @@ theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) ring #align complex.cos_sub_cos Complex.cos_sub_cos -theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := - by +theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := by have h2 : (2 : ℂ) ≠ 0 := by norm_num calc cos x + cos y = cos ((x + y) / 2 + (x - y) / 2) + cos ((x + y) / 2 - (x - y) / 2) := _ @@ -1060,8 +1041,7 @@ theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, a theorem sin_sq : sin x ^ 2 = 1 - cos x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel] #align complex.sin_sq Complex.sin_sq -theorem inv_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := - by +theorem inv_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := by have : cos x ^ 2 ≠ 0 := pow_ne_zero 2 hx rw [tan_eq_sin_div_cos, div_pow] field_simp [this] @@ -1072,8 +1052,7 @@ theorem tan_sq_div_one_add_tan_sq {x : ℂ} (hx : cos x ≠ 0) : simp only [← tan_mul_cos hx, mul_pow, ← inv_one_add_tan_sq hx, div_eq_mul_inv, one_mul] #align complex.tan_sq_div_one_add_tan_sq Complex.tan_sq_div_one_add_tan_sq -theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := - by +theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by have h1 : x + 2 * x = 3 * x := by ring rw [← h1, cos_add x (2 * x)] simp only [cos_two_mul, sin_two_mul, mul_add, mul_sub, mul_one, sq] @@ -1082,8 +1061,7 @@ theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := ring #align complex.cos_three_mul Complex.cos_three_mul -theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := - by +theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by have h1 : x + 2 * x = 3 * x := by ring rw [← h1, sin_add x (2 * x)] simp only [cos_two_mul, sin_two_mul, cos_sq'] @@ -1103,14 +1081,12 @@ theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im rw [← exp_add_mul_I, re_add_im] #align complex.exp_eq_exp_re_mul_sin_add_cos Complex.exp_eq_exp_re_mul_sin_add_cos -theorem exp_re : (exp x).re = Real.exp x.re * Real.cos x.im := - by +theorem exp_re : (exp x).re = Real.exp x.re * Real.cos x.im := by rw [exp_eq_exp_re_mul_sin_add_cos] simp [exp_of_real_re, cos_of_real_re] #align complex.exp_re Complex.exp_re -theorem exp_im : (exp x).im = Real.exp x.re * Real.sin x.im := - by +theorem exp_im : (exp x).im = Real.exp x.re * Real.sin x.im := by rw [exp_eq_exp_re_mul_sin_add_cos] simp [exp_of_real_re, sin_of_real_re] #align complex.exp_im Complex.exp_im @@ -1127,8 +1103,7 @@ theorem exp_of_real_mul_i_im (x : ℝ) : (exp (x * i)).im = Real.sin x := by /-- **De Moivre's formula** -/ theorem cos_add_sin_mul_i_pow (n : ℕ) (z : ℂ) : - (cos z + sin z * i) ^ n = cos (↑n * z) + sin (↑n * z) * i := - by + (cos z + sin z * i) ^ n = cos (↑n * z) + sin (↑n * z) * i := by rw [← exp_mul_I, ← exp_mul_I] induction' n with n ih · rw [pow_zero, Nat.cast_zero, zero_mul, zero_mul, exp_zero] @@ -1217,16 +1192,14 @@ theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] #align real.cos_sub Real.cos_sub -theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := - by +theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := by rw [← of_real_inj] simp only [sin, cos, of_real_sin_of_real_re, of_real_sub, of_real_add, of_real_div, of_real_mul, of_real_one, of_real_bit0] convert sin_sub_sin _ _ <;> norm_cast #align real.sin_sub_sin Real.sin_sub_sin -theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := - by +theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := by rw [← of_real_inj] simp only [cos, neg_mul, of_real_sin, of_real_sub, of_real_add, of_real_cos_of_real_re, of_real_div, of_real_mul, of_real_one, of_real_neg, of_real_bit0] @@ -1234,8 +1207,7 @@ theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) ring #align real.cos_sub_cos Real.cos_sub_cos -theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := - by +theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := by rw [← of_real_inj] simp only [cos, of_real_sub, of_real_add, of_real_cos_of_real_re, of_real_div, of_real_mul, of_real_one, of_real_bit0] @@ -1440,8 +1412,7 @@ theorem exp_sub_sinh : exp x - sinh x = cosh x := #align real.exp_sub_sinh Real.exp_sub_sinh @[simp] -theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := - by +theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by rw [← of_real_inj] simp #align real.cosh_sub_sinh Real.cosh_sub_sinh @@ -1643,8 +1614,7 @@ theorem sum_div_factorial_le {α : Type _} [LinearOrderedField α] (n j : ℕ) ( #align complex.sum_div_factorial_le Complex.sum_div_factorial_le theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : - abs (exp x - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n * (n.succ * (n ! * n)⁻¹) := - by + abs (exp x - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n * (n.succ * (n ! * n)⁻¹) := by rw [← lim_const (∑ m in range n, _), exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] refine' lim_le (CauSeq.le_of_exists ⟨n, fun j hj => _⟩) simp_rw [← sub_eq_add_neg] @@ -1678,8 +1648,7 @@ theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : #align complex.exp_bound Complex.exp_bound theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : - abs (exp x - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n / n ! * 2 := - by + abs (exp x - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n / n ! * 2 := by rw [← lim_const (∑ m in range n, _), exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] refine' lim_le (CauSeq.le_of_exists ⟨n, fun j hj => _⟩) simp_rw [← sub_eq_add_neg] @@ -1742,15 +1711,13 @@ namespace Real open Complex Finset theorem exp_bound {x : ℝ} (hx : |x| ≤ 1) {n : ℕ} (hn : 0 < n) : - |exp x - ∑ m in range n, x ^ m / m !| ≤ |x| ^ n * (n.succ / (n ! * n)) := - by + |exp x - ∑ m in range n, x ^ m / m !| ≤ |x| ^ n * (n.succ / (n ! * n)) := by have hxc : Complex.abs x ≤ 1 := by exact_mod_cast hx convert exp_bound hxc hn <;> norm_cast #align real.exp_bound Real.exp_bound theorem exp_bound' {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) {n : ℕ} (hn : 0 < n) : - Real.exp x ≤ (∑ m in Finset.range n, x ^ m / m !) + x ^ n * (n + 1) / (n ! * n) := - by + Real.exp x ≤ (∑ m in Finset.range n, x ^ m / m !) + x ^ n * (n + 1) / (n ! * n) := by have h3 : |x| = x := by simpa have h4 : |x| ≤ 1 := by rwa [h3] have h' := Real.exp_bound h4 hn @@ -1760,14 +1727,12 @@ theorem exp_bound' {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) {n : ℕ} (hn : 0 < n simpa [mul_div_assoc] using t #align real.exp_bound' Real.exp_bound' -theorem abs_exp_sub_one_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1| ≤ 2 * |x| := - by +theorem abs_exp_sub_one_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1| ≤ 2 * |x| := by have : Complex.abs x ≤ 1 := by exact_mod_cast hx exact_mod_cast Complex.abs_exp_sub_one_le this #align real.abs_exp_sub_one_le Real.abs_exp_sub_one_le -theorem abs_exp_sub_one_sub_id_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1 - x| ≤ x ^ 2 := - by +theorem abs_exp_sub_one_sub_id_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1 - x| ≤ x ^ 2 := by rw [← _root_.sq_abs] have : Complex.abs x ≤ 1 := by exact_mod_cast hx exact_mod_cast Complex.abs_exp_sub_one_sub_id_le this @@ -1797,8 +1762,7 @@ theorem expNear_sub (n x r₁ r₂) : expNear n x r₁ - expNear n x r₂ = x ^ #align real.exp_near_sub Real.expNear_sub theorem exp_approx_end (n m : ℕ) (x : ℝ) (e₁ : n + 1 = m) (h : |x| ≤ 1) : - |exp x - expNear m x 0| ≤ |x| ^ m / m ! * ((m + 1) / m) := - by + |exp x - expNear m x 0| ≤ |x| ^ m / m ! * ((m + 1) / m) := by simp [exp_near] convert exp_bound h _ using 1 field_simp [mul_comm] @@ -1808,8 +1772,7 @@ theorem exp_approx_end (n m : ℕ) (x : ℝ) (e₁ : n + 1 = m) (h : |x| ≤ 1) theorem exp_approx_succ {n} {x a₁ b₁ : ℝ} (m : ℕ) (e₁ : n + 1 = m) (a₂ b₂ : ℝ) (e : |1 + x / m * a₂ - a₁| ≤ b₁ - |x| / m * b₂) (h : |exp x - expNear m x a₂| ≤ |x| ^ m / m ! * b₂) : - |exp x - expNear n x a₁| ≤ |x| ^ n / n ! * b₁ := - by + |exp x - expNear n x a₁| ≤ |x| ^ n / n ! * b₁ := by refine' (_root_.abs_sub_le _ _ _).trans ((add_le_add_right h _).trans _) subst e₁; rw [exp_near_succ, exp_near_sub, _root_.abs_mul] convert mul_le_mul_of_nonneg_left (le_sub_iff_add_le'.1 e) _ @@ -1826,8 +1789,7 @@ theorem exp_approx_end' {n} {x a b : ℝ} (m : ℕ) (e₁ : n + 1 = m) (rm : ℝ theorem exp_1_approx_succ_eq {n} {a₁ b₁ : ℝ} {m : ℕ} (en : n + 1 = m) {rm : ℝ} (er : ↑m = rm) (h : |exp 1 - expNear m 1 ((a₁ - 1) * rm)| ≤ |1| ^ m / m ! * (b₁ * rm)) : - |exp 1 - expNear n 1 a₁| ≤ |1| ^ n / n ! * b₁ := - by + |exp 1 - expNear n 1 a₁| ≤ |1| ^ n / n ! * b₁ := by subst er refine' exp_approx_succ _ en _ _ _ h field_simp [show (m : ℝ) ≠ 0 by norm_cast <;> linarith] @@ -1996,8 +1958,7 @@ theorem cos_two_neg : cos 2 < 0 := theorem exp_bound_div_one_sub_of_interval_approx {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) : (∑ j : ℕ in Finset.range 3, x ^ j / j.factorial) + x ^ 3 * ((3 : ℕ) + 1) / ((3 : ℕ).factorial * (3 : ℕ)) ≤ - ∑ j in Finset.range 3, x ^ j := - by + ∑ j in Finset.range 3, x ^ j := by norm_num [Finset.sum] rw [add_assoc, add_comm (x + 1) (x ^ 3 * 4 / 18), ← add_assoc, add_le_add_iff_right, ← add_le_add_iff_left (-(x ^ 2 / 2)), ← add_assoc, CommRing.add_left_neg (x ^ 2 / 2), zero_add, @@ -2028,8 +1989,7 @@ theorem exp_bound_div_one_sub_of_interval {x : ℝ} (h1 : 0 ≤ x) (h2 : x < 1) ((exp_bound_div_one_sub_of_interval_approx h1 h2.le).trans h) #align real.exp_bound_div_one_sub_of_interval Real.exp_bound_div_one_sub_of_interval -theorem one_sub_le_exp_minus_of_pos {y : ℝ} (h : 0 ≤ y) : 1 - y ≤ Real.exp (-y) := - by +theorem one_sub_le_exp_minus_of_pos {y : ℝ} (h : 0 ≤ y) : 1 - y ≤ Real.exp (-y) := by rw [Real.exp_neg] have r1 : (1 - y) * Real.exp y ≤ 1 := by @@ -2047,22 +2007,19 @@ theorem one_sub_le_exp_minus_of_pos {y : ℝ} (h : 0 ≤ y) : 1 - y ≤ Real.exp rwa [mul_comm] at r1 #align real.one_sub_le_exp_minus_of_pos Real.one_sub_le_exp_minus_of_pos -theorem add_one_le_exp_of_nonpos {x : ℝ} (h : x ≤ 0) : x + 1 ≤ Real.exp x := - by +theorem add_one_le_exp_of_nonpos {x : ℝ} (h : x ≤ 0) : x + 1 ≤ Real.exp x := by rw [add_comm] have h1 : 0 ≤ -x := by linarith simpa using one_sub_le_exp_minus_of_pos h1 #align real.add_one_le_exp_of_nonpos Real.add_one_le_exp_of_nonpos -theorem add_one_le_exp (x : ℝ) : x + 1 ≤ Real.exp x := - by +theorem add_one_le_exp (x : ℝ) : x + 1 ≤ Real.exp x := by cases le_or_lt 0 x · exact Real.add_one_le_exp_of_nonneg h exact add_one_le_exp_of_nonpos h.le #align real.add_one_le_exp Real.add_one_le_exp -theorem one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t / n) ^ n ≤ exp (-t) := - by +theorem one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t / n) ^ n ≤ exp (-t) := by rcases eq_or_ne n 0 with (rfl | hn) · simp rwa [Nat.cast_zero] at ht' @@ -2094,8 +2051,7 @@ end Tactic namespace Complex @[simp] -theorem abs_cos_add_sin_mul_i (x : ℝ) : abs (cos x + sin x * i) = 1 := - by +theorem abs_cos_add_sin_mul_i (x : ℝ) : abs (cos x + sin x * i) = 1 := by have := Real.sin_sq_add_cos_sq x simp_all [add_comm, abs, norm_sq, sq, sin_of_real_re, cos_of_real_re, mul_re] #align complex.abs_cos_add_sin_mul_I Complex.abs_cos_add_sin_mul_i From 3aef49b00425e99c1bed581da43c4ba52ef25cd5 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 14:35:54 +0000 Subject: [PATCH 04/23] minor fix to Complex.Basic --- Mathlib/Data/Complex/Basic.lean | 9 ++------- 1 file changed, 2 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 0b4d16acaaee4..78e51ac7801b3 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -935,21 +935,16 @@ private theorem abs_add (z w : ℂ) : (abs z + w) ≤ (abs z) + abs w := -- #align complex.abs_theory.abs_add complex.abs_theory.abs_add /-- The complex absolute value function, defined as the square root of the norm squared. -/ -noncomputable def Complex.abs : AbsoluteValue ℂ ℝ - where +noncomputable def _root_.Complex.abs : AbsoluteValue ℂ ℝ where toFun x := abs x map_mul' := abs_mul nonneg' := abs_nonneg' eq_zero' _ := (Real.sqrt_eq_zero <| normSq_nonneg _).trans normSq_eq_zero add_le' := abs_add -#align complex.abs Complex.AbsTheory.Complex.abs +#align complex.abs Complex.abs end AbsTheory --- Porting note: Added this to make the following work. --- Hope this is as intended. Also the name has been expanded. -open Complex.AbsTheory - theorem abs_def : (Complex.abs : ℂ → ℝ) = fun z => (normSq z).sqrt := rfl #align complex.abs_def Complex.abs_def From e619681bc1ed9bb46a639f81a7d58f16f9e69d77 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 14:39:31 +0000 Subject: [PATCH 05/23] fix build --- Mathlib/Data/Complex/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 78e51ac7801b3..c82c3cef694c9 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -1099,13 +1099,13 @@ theorem abs_le_sqrt_two_mul_max (z : ℂ) : Complex.abs z ≤ Real.sqrt 2 * max theorem abs_re_div_abs_le_one (z : ℂ) : |z.re / Complex.abs z| ≤ 1 := if hz : z = 0 then by simp [hz, zero_le_one] else by simp_rw [_root_.abs_div, abs_abs, - div_le_iff (AbsoluteValue.pos Complex.AbsTheory.Complex.abs hz), one_mul, abs_re_le_abs] + div_le_iff (AbsoluteValue.pos Complex.abs hz), one_mul, abs_re_le_abs] #align complex.abs_re_div_abs_le_one Complex.abs_re_div_abs_le_one theorem abs_im_div_abs_le_one (z : ℂ) : |z.im / Complex.abs z| ≤ 1 := if hz : z = 0 then by simp [hz, zero_le_one] else by simp_rw [_root_.abs_div, abs_abs, - div_le_iff (AbsoluteValue.pos Complex.AbsTheory.Complex.abs hz), one_mul, abs_im_le_abs] + div_le_iff (AbsoluteValue.pos Complex.abs hz), one_mul, abs_im_le_abs] #align complex.abs_im_div_abs_le_one Complex.abs_im_div_abs_le_one -- Porting note: removed `norm_cast` attribute because the RHS can't start with `↑` From 7515deac1685945ef5f2949c6e52a541013a5409 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 14:49:09 +0000 Subject: [PATCH 06/23] start work --- Mathlib/Data/Complex/Exponential.lean | 400 +++++++++++++------------- 1 file changed, 204 insertions(+), 196 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 729887fea36d5..cf78db73a8725 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -26,7 +26,7 @@ local notation "abs'" => Abs.abs open IsAbsoluteValue -open Classical BigOperators Nat ComplexConjugate +open Classical BigOperators ComplexConjugate section @@ -37,7 +37,7 @@ section variable {α : Type _} {β : Type _} [Ring β] [LinearOrderedField α] [Archimedean α] {abv : β → α} [IsAbsoluteValue abv] -theorem is_cau_of_decreasing_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) +theorem isCauSeq_of_decreasing_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) (hnm : ∀ n ≥ m, f n.succ ≤ f n) : IsCauSeq abs f := fun ε ε0 => by let ⟨k, hk⟩ := Archimedean.arch a ε0 have h : ∃ l, ∀ n ≥ m, a - l • ε < f n := @@ -54,7 +54,7 @@ theorem is_cau_of_decreasing_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : have hl : ∀ n : ℕ, n ≥ m → f n > a - l • ε := Nat.find_spec h have hl0 : l ≠ 0 := fun hl0 => not_lt_of_ge (ham m le_rfl) - (lt_of_lt_of_le (by have := hl m (le_refl m) <;> simpa [hl0] using this) (le_abs_self (f m))) + (lt_of_lt_of_le (by have := hl m (le_refl m); simpa [hl0] using this) (le_abs_self (f m))) cases' not_forall.1 (Nat.find_min h (Nat.pred_lt hl0)) with i hi rw [not_imp, not_lt] at hi exists i @@ -69,18 +69,18 @@ theorem is_cau_of_decreasing_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : rw [← Nat.succ_pred_eq_of_pos (Nat.pos_of_ne_zero hl0), succ_nsmul', sub_add, add_sub_cancel] _ < f j + ε := add_lt_add_right (hl j (le_trans hi.1 hj)) _ - -#align is_cau_of_decreasing_bounded is_cau_of_decreasing_bounded -theorem is_cau_of_mono_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) +#align is_cau_of_decreasing_bounded isCauSeq_of_decreasing_bounded + +theorem isCauSeq_of_mono_bounded (f : ℕ → α) {a : α} {m : ℕ} (ham : ∀ n ≥ m, |f n| ≤ a) (hnm : ∀ n ≥ m, f n ≤ f n.succ) : IsCauSeq abs f := by refine' - @Eq.recOn (ℕ → α) _ (IsCauSeq abs) _ _ - (-⟨_, @is_cau_of_decreasing_bounded _ _ _ (fun n => -f n) a m (by simpa) (by simpa)⟩ : + @Eq.ndrecOn (ℕ → α) _ (IsCauSeq abs) _ _ + (-⟨_, @isCauSeq_of_decreasing_bounded _ _ _ (fun n => -f n) a m (by simpa) (by simpa)⟩ : CauSeq α abs).2 ext exact neg_neg _ -#align is_cau_of_mono_bounded is_cau_of_mono_bounded +#align is_cau_of_mono_bounded isCauSeq_of_mono_bounded end @@ -89,7 +89,7 @@ section NoArchimedean variable {α : Type _} {β : Type _} [Ring β] [LinearOrderedField α] {abv : β → α} [IsAbsoluteValue abv] -theorem is_cau_series_of_abv_le_cau {f : ℕ → β} {g : ℕ → α} (n : ℕ) : +theorem isCauSeq_series_of_abv_le_of_isCauSeq {f : ℕ → β} {g : ℕ → α} (n : ℕ) : (∀ m, n ≤ m → abv (f m) ≤ g m) → (IsCauSeq abs fun n => ∑ i in range n, g i) → IsCauSeq abv fun n => ∑ i in range n, f i := by intro hm hg ε ε0 @@ -107,19 +107,21 @@ theorem is_cau_series_of_abv_le_cau {f : ℕ → β} {g : ℕ → α} (n : ℕ) clear this hi₂ hi₁ hi ε0 ε hg sub_le rw [tsub_eq_iff_eq_add_of_le ji] at hk rw [hk] + dsimp only clear hk ji j induction' k with k' hi · simp [abv_zero abv] - · simp only [Nat.succ_add, sum_range_succ_comm, sub_eq_add_neg, add_assoc] - refine' le_trans (abv_add _ _ _) _ + · simp only [Nat.succ_add, Nat.succ_eq_add_one, Finset.sum_range_succ_comm] + simp only [add_assoc, sub_eq_add_neg] + refine le_trans (abv_add _ _ _) ?_ simp only [sub_eq_add_neg] at hi - exact add_le_add (hm _ (le_add_of_nonneg_of_le (Nat.zero_le _) (le_max_left _ _))) hi -#align is_cau_series_of_abv_le_cau is_cau_series_of_abv_le_cau + refine add_le_add (hm _ (le_add_of_nonneg_of_le (Nat.zero_le _) (le_max_left _ _))) hi +#align is_cau_series_of_abv_le_cau isCauSeq_series_of_abv_le_of_isCauSeq -theorem is_cau_series_of_abv_cau {f : ℕ → β} : +theorem isCauSeq_series_of_abv_isCauSeq {f : ℕ → β} : (IsCauSeq abs fun m => ∑ n in range m, abv (f n)) → IsCauSeq abv fun m => ∑ n in range m, f n := - is_cau_series_of_abv_le_cau 0 fun n h => le_rfl -#align is_cau_series_of_abv_cau is_cau_series_of_abv_cau + isCauSeq_series_of_abv_le_of_isCauSeq 0 fun _ _ => le_rfl +#align is_cau_series_of_abv_cau isCauSeq_series_of_abv_isCauSeq end NoArchimedean @@ -127,14 +129,14 @@ section variable {α : Type _} [LinearOrderedField α] [Archimedean α] -theorem is_cau_geo_series {β : Type _} [Ring β] [Nontrivial β] {abv : β → α} [IsAbsoluteValue abv] +theorem isCauSeq_geo_series {β : Type _} [Ring β] [Nontrivial β] {abv : β → α} [IsAbsoluteValue abv] (x : β) (hx1 : abv x < 1) : IsCauSeq abv fun n => ∑ m in range n, x ^ m := - have hx1' : abv x ≠ 1 := fun h => by simpa [h, lt_irrefl] using hx1 - is_cau_series_of_abv_cau + have hx1' : abv x ≠ 1 := fun h => by simp [h, lt_irrefl] at hx1 + isCauSeq_series_of_abv_isCauSeq (by simp only [abv_pow abv, geom_sum_eq hx1'] conv in _ / _ => rw [← neg_div_neg_eq, neg_sub, neg_sub] - refine' @is_cau_of_mono_bounded _ _ _ _ ((1 : α) / (1 - abv x)) 0 _ _ + refine' @isCauSeq_of_mono_bounded _ _ _ _ ((1 : α) / (1 - abv x)) 0 _ _ · intro n hn rw [abs_of_nonneg] refine' @@ -146,18 +148,19 @@ theorem is_cau_geo_series {β : Type _} [Ring β] [Nontrivial β] {abv : β → · simp · rw [pow_succ, ← one_mul (1 : α)] refine' mul_le_mul (le_of_lt hx1) ih (abv_pow abv x n ▸ abv_nonneg _ _) (by norm_num) - · intro n hn + · intro n _ refine' div_le_div_of_le (le_of_lt <| sub_pos.2 hx1) (sub_le_sub_left _ _) - rw [← one_mul (_ ^ n), pow_succ] + rw [←one_mul (abv x ^ n), pow_succ] exact mul_le_mul_of_nonneg_right (le_of_lt hx1) (pow_nonneg (abv_nonneg _ _) _)) -#align is_cau_geo_series is_cau_geo_series +#align is_cau_geo_series isCauSeq_geo_series -theorem is_cau_geo_series_const (a : α) {x : α} (hx1 : |x| < 1) : - IsCauSeq abs fun m => ∑ n in range m, a * x ^ n := by - have : IsCauSeq abs fun m => a * ∑ n in range m, x ^ n := - (CauSeq.const abs a * ⟨_, is_cau_geo_series x hx1⟩).2 - simpa only [mul_sum] -#align is_cau_geo_series_const is_cau_geo_series_const +theorem isCauSeq_geo_series_const (a : α) {x : α} (hx1 : |x| < 1) : + IsCauSeq abs fun m => ∑ n in range m, (a * x ^ n) := by + have : IsCauSeq abs fun m => a * ∑ n in range m, (x ^ n) := + (CauSeq.const abs a * + (show CauSeq α abs from ⟨fun n => ∑ m in range n, x ^ m, isCauSeq_geo_series x hx1⟩)).2 + simpa [mul_sum] using this +#align is_cau_geo_series_const isCauSeq_geo_series_const variable {β : Type _} [Ring β] {abv : β → α} [IsAbsoluteValue abv] @@ -166,56 +169,59 @@ theorem series_ratio_test {f : ℕ → β} (n : ℕ) (r : α) (hr0 : 0 ≤ r) (h by have har1 : |r| < 1 := by rwa [abs_of_nonneg hr0] refine' - is_cau_series_of_abv_le_cau n.succ _ - (is_cau_geo_series_const (abv (f n.succ) * r⁻¹ ^ n.succ) har1) + isCauSeq_series_of_abv_le_of_isCauSeq n.succ _ + (isCauSeq_geo_series_const (abv (f n.succ) * r⁻¹ ^ n.succ) har1) intro m hmn cases' Classical.em (r = 0) with r_zero r_ne_zero · have m_pos := lt_of_lt_of_le (Nat.succ_pos n) hmn have := h m.pred (Nat.le_of_succ_le_succ (by rwa [Nat.succ_pred_eq_of_pos m_pos])) - simpa [r_zero, Nat.succ_pred_eq_of_pos m_pos, pow_succ] + simpa [r_zero, Nat.succ_pred_eq_of_pos m_pos, pow_succ] using this generalize hk : m - n.succ = k have r_pos : 0 < r := lt_of_le_of_ne hr0 (Ne.symm r_ne_zero) replace hk : m = k + n.succ := (tsub_eq_iff_eq_add_of_le hmn).1 hk induction' k with k ih generalizing m n - · rw [hk, zero_add, mul_right_comm, inv_pow _ _, ← div_eq_mul_inv, mul_div_cancel] + · rw [hk, Nat.zero_add, mul_right_comm, inv_pow _ _, ← div_eq_mul_inv, mul_div_cancel] + exact le_refl _ exact (ne_of_lt (pow_pos r_pos _)).symm · have kn : k + n.succ ≥ n.succ := by - rw [← zero_add n.succ] <;> exact add_le_add (zero_le _) (by simp) - rw [hk, Nat.succ_add, pow_succ' r, ← mul_assoc] - exact + rw [← zero_add n.succ]; exact add_le_add (Nat.zero_le _) (by simp) + erw [hk, Nat.succ_add, pow_succ' r, ← mul_assoc] + refine le_trans (by rw [mul_comm] <;> exact h _ (Nat.le_of_succ_le kn)) - (mul_le_mul_of_nonneg_right (ih (k + n.succ) n h kn rfl) hr0) + (mul_le_mul_of_nonneg_right ?_ hr0) + exact ih _ h _ (by simp) rfl + #align series_ratio_test series_ratio_test theorem sum_range_diag_flip {α : Type _} [AddCommMonoid α] (n : ℕ) (f : ℕ → ℕ → α) : (∑ m in range n, ∑ k in range (m + 1), f k (m - k)) = ∑ m in range n, ∑ k in range (n - m), f m k := by - rw [sum_sigma', sum_sigma'] <;> - exact - sum_bij (fun a _ => ⟨a.2, a.1 - a.2⟩) - (fun a ha => - have h₁ : a.1 < n := mem_range.1 (mem_sigma.1 ha).1 - have h₂ : a.2 < Nat.succ a.1 := mem_range.1 (mem_sigma.1 ha).2 - mem_sigma.2 - ⟨mem_range.2 (lt_of_lt_of_le h₂ h₁), - mem_range.2 ((tsub_lt_tsub_iff_right (Nat.le_of_lt_succ h₂)).2 h₁)⟩) - (fun _ _ => rfl) - (fun ⟨a₁, a₂⟩ ⟨b₁, b₂⟩ ha hb h => - have ha : a₁ < n ∧ a₂ ≤ a₁ := - ⟨mem_range.1 (mem_sigma.1 ha).1, Nat.le_of_lt_succ (mem_range.1 (mem_sigma.1 ha).2)⟩ - have hb : b₁ < n ∧ b₂ ≤ b₁ := - ⟨mem_range.1 (mem_sigma.1 hb).1, Nat.le_of_lt_succ (mem_range.1 (mem_sigma.1 hb).2)⟩ - have h : a₂ = b₂ ∧ _ := Sigma.mk.inj h - have h' : a₁ = b₁ - b₂ + a₂ := (tsub_eq_iff_eq_add_of_le ha.2).1 (eq_of_hEq h.2) - Sigma.mk.inj_iff.2 ⟨tsub_add_cancel_of_le hb.2 ▸ h'.symm ▸ h.1 ▸ rfl, hEq_of_eq h.1⟩) - fun ⟨a₁, a₂⟩ ha => - have ha : a₁ < n ∧ a₂ < n - a₁ := - ⟨mem_range.1 (mem_sigma.1 ha).1, mem_range.1 (mem_sigma.1 ha).2⟩ - ⟨⟨a₂ + a₁, a₁⟩, - ⟨mem_sigma.2 - ⟨mem_range.2 (lt_tsub_iff_right.1 ha.2), - mem_range.2 (Nat.lt_succ_of_le (Nat.le_add_left _ _))⟩, - Sigma.mk.inj_iff.2 ⟨rfl, hEq_of_eq (add_tsub_cancel_right _ _).symm⟩⟩⟩ + rw [sum_sigma', sum_sigma'] + exact + sum_bij (fun a _ => ⟨a.2, a.1 - a.2⟩) + (fun a ha => + have h₁ : a.1 < n := mem_range.1 (mem_sigma.1 ha).1 + have h₂ : a.2 < Nat.succ a.1 := mem_range.1 (mem_sigma.1 ha).2 + mem_sigma.2 + ⟨mem_range.2 (lt_of_lt_of_le h₂ h₁), + mem_range.2 ((tsub_lt_tsub_iff_right (Nat.le_of_lt_succ h₂)).2 h₁)⟩) + (fun _ _ => rfl) + (fun ⟨a₁, a₂⟩ ⟨b₁, b₂⟩ ha hb h => + have ha : a₁ < n ∧ a₂ ≤ a₁ := + ⟨mem_range.1 (mem_sigma.1 ha).1, Nat.le_of_lt_succ (mem_range.1 (mem_sigma.1 ha).2)⟩ + have hb : b₁ < n ∧ b₂ ≤ b₁ := + ⟨mem_range.1 (mem_sigma.1 hb).1, Nat.le_of_lt_succ (mem_range.1 (mem_sigma.1 hb).2)⟩ + have h : a₂ = b₂ ∧ _ := by simpa using h + have h' : a₁ = b₁ - b₂ + a₂ := (tsub_eq_iff_eq_add_of_le ha.2).1 (eq_of_heq h.2) + Sigma.mk.inj_iff.2 ⟨tsub_add_cancel_of_le hb.2 ▸ h'.symm ▸ h.1 ▸ rfl, heq_of_eq h.1⟩) + fun ⟨a₁, a₂⟩ ha => + have ha : a₁ < n ∧ a₂ < n - a₁ := + ⟨mem_range.1 (mem_sigma.1 ha).1, mem_range.1 (mem_sigma.1 ha).2⟩ + ⟨⟨a₂ + a₁, a₁⟩, + ⟨mem_sigma.2 + ⟨mem_range.2 (lt_tsub_iff_right.1 ha.2), + mem_range.2 (Nat.lt_succ_of_le (Nat.le_add_left _ _))⟩, + Sigma.mk.inj_iff.2 ⟨rfl, heq_of_eq (add_tsub_cancel_right _ _).symm⟩⟩⟩ #align sum_range_diag_flip sum_range_diag_flip end @@ -232,7 +238,7 @@ theorem abv_sum_le_sum_abv {γ : Type _} (f : γ → β) (s : Finset γ) : abv (∑ k in s, f k) ≤ ∑ k in s, abv (f k) := haveI := Classical.decEq γ Finset.induction_on s (by simp [abv_zero abv]) fun a s has ih => by - rw [sum_insert has, sum_insert has] <;> exact le_trans (abv_add abv _ _) (add_le_add_left ih _) + rw [sum_insert has, sum_insert has]; exact le_trans (abv_add abv _ _) (add_le_add_left ih _) #align abv_sum_le_sum_abv abv_sum_le_sum_abv end @@ -270,21 +276,21 @@ theorem cauchy_product {a b : ℕ → β} (ha : IsCauSeq abs fun m => ∑ n in r (∑ i in range K, a i * ∑ k in range (K - i), b k) = (∑ i in range K, a i * ((∑ k in range (K - i), b k) - ∑ k in range K, b k)) + ∑ i in range K, a i * ∑ k in range K, b k := - by rw [← sum_add_distrib] <;> simp [(mul_add _ _ _).symm] + by rw [← sum_add_distrib]; simp [(mul_add _ _ _).symm] have two_mul_two : (4 : α) = 2 * 2 := by norm_num - have hQ0 : Q ≠ 0 := fun h => by simpa [h, lt_irrefl] using hQε0 + have hQ0 : Q ≠ 0 := fun h => by simp [h, lt_irrefl] at hQε0 have h2Q0 : 2 * Q ≠ 0 := mul_ne_zero two_ne_zero hQ0 have hε : ε / (2 * P) * P + ε / (4 * Q) * (2 * Q) = ε := by rw [← div_div, div_mul_cancel _ (Ne.symm (ne_of_lt hP0)), two_mul_two, mul_assoc, ← div_div, div_mul_cancel _ h2Q0, add_halves] have hNMK : max N M + 1 < K := - lt_of_lt_of_le (by rw [two_mul] <;> exact lt_add_of_pos_left _ (Nat.succ_pos _)) hK + lt_of_lt_of_le (by rw [two_mul]; exact lt_add_of_pos_left _ (Nat.succ_pos _)) hK have hKN : N < K := calc N ≤ max N M := le_max_left _ _ _ < max N M + 1 := (Nat.lt_succ_self _) _ < K := hNMK - + have hsumlesum : (∑ i in range (max N M + 1), abv (a i) * abv ((∑ k in range (K - i), b k) - ∑ k in range K, b k)) ≤ @@ -296,19 +302,19 @@ theorem cauchy_product {a b : ℕ → β} (ha : IsCauSeq abs fun m => ∑ n in r (le_tsub_of_add_le_left (le_trans (by - rw [two_mul] <;> - exact - add_le_add (le_of_lt (mem_range.1 hmJ)) - (le_trans (le_max_left _ _) (le_of_lt (lt_add_one _)))) + rw [two_mul] + exact + add_le_add (le_of_lt (mem_range.1 hmJ)) + (le_trans (le_max_left _ _) (le_of_lt (lt_add_one _)))) hK)) K (le_of_lt hKN))) (abv_nonneg abv _) have hsumltP : (∑ n in range (max N M + 1), abv (a n)) < P := calc (∑ n in range (max N M + 1), abv (a n)) = |∑ n in range (max N M + 1), abv (a n)| := - Eq.symm (abs_of_nonneg (sum_nonneg fun x h => abv_nonneg abv (a x))) + Eq.symm (abs_of_nonneg (sum_nonneg fun x _ => abv_nonneg abv (a x))) _ < P := hP (max N M + 1) - + rw [h₁, h₂, h₃, sum_mul, ← sub_sub, sub_right_comm, sub_self, zero_sub, abv_neg abv] refine' lt_of_le_of_lt (abv_sum_le_sum_abv _ _) _ suffices @@ -320,18 +326,18 @@ theorem cauchy_product {a b : ℕ → β} (ha : IsCauSeq abs fun m => ∑ n in r ε / (2 * P) * P + ε / (4 * Q) * (2 * Q) by rw [hε] at this - simpa [abv_mul abv] + simpa [abv_mul abv] using this refine' add_lt_add (lt_of_le_of_lt hsumlesum - (by rw [← sum_mul, mul_comm] <;> exact (mul_lt_mul_left hPε0).mpr hsumltP)) + (by rw [← sum_mul, mul_comm]; exact (mul_lt_mul_left hPε0).mpr hsumltP)) _ rw [sum_range_sub_sum_range (le_of_lt hNMK)] calc - (∑ i in (range K).filterₓ fun k => max N M + 1 ≤ k, + (∑ i in (range K).filter fun k => max N M + 1 ≤ k, abv (a i) * abv ((∑ k in range (K - i), b k) - ∑ k in range K, b k)) ≤ - ∑ i in (range K).filterₓ fun k => max N M + 1 ≤ k, abv (a i) * (2 * Q) := - sum_le_sum fun n hn => + ∑ i in (range K).filter fun k => max N M + 1 ≤ k, abv (a i) * (2 * Q) := + sum_le_sum fun n _ => by refine' mul_le_mul_of_nonneg_left _ (abv_nonneg _ _) rw [sub_eq_add_neg] @@ -339,13 +345,13 @@ theorem cauchy_product {a b : ℕ → β} (ha : IsCauSeq abs fun m => ∑ n in r rw [two_mul, abv_neg abv] exact add_le_add (le_of_lt (hQ _)) (le_of_lt (hQ _)) _ < ε / (4 * Q) * (2 * Q) := by - rw [← sum_mul, ← sum_range_sub_sum_range (le_of_lt hNMK)] <;> + rw [← sum_mul, ← sum_range_sub_sum_range (le_of_lt hNMK)]; refine' (mul_lt_mul_right <| by - rw [two_mul] <;> - exact - add_pos (lt_of_le_of_lt (abv_nonneg _ _) (hQ 0)) - (lt_of_le_of_lt (abv_nonneg _ _) (hQ 0))).2 + rw [two_mul] + exact + add_pos (lt_of_le_of_lt (abv_nonneg _ _) (hQ 0)) + (lt_of_le_of_lt (abv_nonneg _ _) (hQ 0))).2 (lt_of_le_of_lt (le_abs_self _) (hM _ (le_trans (Nat.le_succ_of_le (le_max_right _ _)) (le_of_lt hNMK)) _ (Nat.le_succ_of_le (le_max_right _ _)))) @@ -364,119 +370,122 @@ open CauSeq namespace Complex -theorem is_cau_abs_exp (z : ℂ) : IsCauSeq Abs.abs fun n => ∑ m in range n, abs (z ^ m / m !) := +theorem isCauSeq_abs_exp (z : ℂ) : + IsCauSeq Abs.abs fun n => ∑ m in range n, abs (z ^ m / m.factorial) := let ⟨n, hn⟩ := exists_nat_gt (abs z) - have hn0 : (0 : ℝ) < n := lt_of_le_of_lt (abs.NonNeg _) hn - series_ratio_test n (Complex.abs z / n) (div_nonneg (abs.NonNeg _) (le_of_lt hn0)) + have hn0 : (0 : ℝ) < n := lt_of_le_of_lt (abs.nonneg _) hn + series_ratio_test n (abs z / n) (div_nonneg (abs.nonneg _) (le_of_lt hn0)) (by rwa [div_lt_iff hn0, one_mul]) fun m hm => by - rw [abs_abs, abs_abs, Nat.factorial_succ, pow_succ, mul_comm m.succ, Nat.cast_mul, ← div_div, - mul_div_assoc, mul_div_right_comm, abs.map_mul, map_div₀, abs_cast_nat] <;> + rw [abs_abs, abs_abs, Nat.factorial_succ, pow_succ, mul_comm m.succ, Nat.cast_mul, ← div_div, + mul_div_assoc, mul_div_right_comm, map_mul, map_div₀, abs_cast_nat] exact mul_le_mul_of_nonneg_right (div_le_div_of_le_left (abs.nonneg _) hn0 (Nat.cast_le.2 (le_trans hm (Nat.le_succ _)))) (abs.nonneg _) -#align complex.is_cau_abs_exp Complex.is_cau_abs_exp +#align complex.is_cau_abs_exp Complex.isCauSeq_abs_exp noncomputable section -theorem is_cau_exp (z : ℂ) : IsCauSeq abs fun n => ∑ m in range n, z ^ m / m ! := - is_cau_series_of_abv_cau (is_cau_abs_exp z) -#align complex.is_cau_exp Complex.is_cau_exp +theorem isCauSeq_exp (z : ℂ) : IsCauSeq abs fun n => ∑ m in range n, z ^ m / m.factorial := + isCauSeq_series_of_abv_isCauSeq (isCauSeq_abs_exp z) +#align complex.is_cau_exp Complex.isCauSeq_exp /-- The Cauchy sequence consisting of partial sums of the Taylor series of the complex exponential function -/ -@[pp_nodot] +----@[pp_nodot] Porting note: removed Porting note: removed def exp' (z : ℂ) : CauSeq ℂ Complex.abs := - ⟨fun n => ∑ m in range n, z ^ m / m !, is_cau_exp z⟩ + ⟨fun n => ∑ m in range n, z ^ m / m.factorial, isCauSeq_exp z⟩ #align complex.exp' Complex.exp' /-- The complex exponential function, defined via its Taylor series -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed irreducible_def exp (z : ℂ) : ℂ := - limUnder (exp' z) + CauSeq.lim (exp' z) #align complex.exp Complex.exp /-- The complex sine function, defined via `exp` -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def sin (z : ℂ) : ℂ := - (exp (-z * i) - exp (z * i)) * i / 2 + (exp (-z * I) - exp (z * I)) * I / 2 #align complex.sin Complex.sin /-- The complex cosine function, defined via `exp` -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def cos (z : ℂ) : ℂ := - (exp (z * i) + exp (-z * i)) / 2 + (exp (z * I) + exp (-z * I)) / 2 #align complex.cos Complex.cos /-- The complex tangent function, defined as `sin z / cos z` -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def tan (z : ℂ) : ℂ := sin z / cos z #align complex.tan Complex.tan /-- The complex hyperbolic sine function, defined via `exp` -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def sinh (z : ℂ) : ℂ := (exp z - exp (-z)) / 2 #align complex.sinh Complex.sinh /-- The complex hyperbolic cosine function, defined via `exp` -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def cosh (z : ℂ) : ℂ := (exp z + exp (-z)) / 2 #align complex.cosh Complex.cosh /-- The complex hyperbolic tangent function, defined as `sinh z / cosh z` -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def tanh (z : ℂ) : ℂ := sinh z / cosh z #align complex.tanh Complex.tanh -end Complex +end +end Complex +#print Complex.exp namespace Real -open Complex +--open Complex /-- The real exponential function, defined as the real part of the complex exponential -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def exp (x : ℝ) : ℝ := - (exp x).re + (Complex.exp x).re #align real.exp Real.exp /-- The real sine function, defined as the real part of the complex sine -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def sin (x : ℝ) : ℝ := (sin x).re #align real.sin Real.sin /-- The real cosine function, defined as the real part of the complex cosine -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def cos (x : ℝ) : ℝ := (cos x).re #align real.cos Real.cos /-- The real tangent function, defined as the real part of the complex tangent -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def tan (x : ℝ) : ℝ := (tan x).re #align real.tan Real.tan /-- The real hypebolic sine function, defined as the real part of the complex hyperbolic sine -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def sinh (x : ℝ) : ℝ := (sinh x).re #align real.sinh Real.sinh /-- The real hypebolic cosine function, defined as the real part of the complex hyperbolic cosine -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def cosh (x : ℝ) : ℝ := (cosh x).re #align real.cosh Real.cosh /-- The real hypebolic tangent function, defined as the real part of the complex hyperbolic tangent -/ -@[pp_nodot] +--@[pp_nodot] Porting note: removed def tanh (x : ℝ) : ℝ := (tanh x).re #align real.tanh Real.tanh @@ -506,18 +515,18 @@ theorem exp_add : exp (x + y) = exp x * exp y := by 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)!) := + ∑ I in range j, ∑ k in range (i + 1), x ^ k / k ! * (y ^ (i - k) / (i - k)!) := by intro j refine' Finset.sum_congr rfl fun m hm => _ rw [add_pow, div_eq_mul_inv, sum_mul] - refine' Finset.sum_congr rfl fun i hi => _ - have h₁ : (m.choose i : ℂ) ≠ 0 := + refine' Finset.sum_congr rfl fun 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 : ℂ)] + 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] @@ -803,32 +812,32 @@ theorem sin_neg : sin (-x) = -sin x := by simp [sin, sub_eq_add_neg, exp_neg, (neg_div _ _).symm, add_mul] #align complex.sin_neg Complex.sin_neg -theorem two_sin : 2 * sin x = (exp (-x * i) - exp (x * i)) * i := +theorem two_sin : 2 * sin x = (exp (-x * I) - exp (x * I)) * I := mul_div_cancel' _ two_ne_zero #align complex.two_sin Complex.two_sin -theorem two_cos : 2 * cos x = exp (x * i) + exp (-x * i) := +theorem two_cos : 2 * cos x = exp (x * I) + exp (-x * I) := mul_div_cancel' _ two_ne_zero #align complex.two_cos Complex.two_cos -theorem sinh_mul_i : sinh (x * i) = sin x * i := by +theorem sinh_mul_i : sinh (x * I) = sin x * I := by rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, ← mul_assoc, two_sin, mul_assoc, I_mul_I, mul_neg_one, neg_sub, neg_mul_eq_neg_mul] #align complex.sinh_mul_I Complex.sinh_mul_i -theorem cosh_mul_i : cosh (x * i) = cos x := by +theorem cosh_mul_i : cosh (x * I) = cos x := by rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, two_cos, neg_mul_eq_neg_mul] #align complex.cosh_mul_I Complex.cosh_mul_i -theorem tanh_mul_i : tanh (x * i) = tan x * i := by +theorem tanh_mul_i : tanh (x * I) = tan x * I := by rw [tanh_eq_sinh_div_cosh, cosh_mul_I, sinh_mul_I, mul_div_right_comm, tan] #align complex.tanh_mul_I Complex.tanh_mul_i -theorem cos_mul_i : cos (x * i) = cosh x := by rw [← cosh_mul_I] <;> ring_nf <;> simp +theorem cos_mul_i : cos (x * I) = cosh x := by rw [← cosh_mul_I] <;> ring_nf <;> simp #align complex.cos_mul_I Complex.cos_mul_i -theorem sin_mul_i : sin (x * i) = sinh x * i := by - have h : i * sin (x * i) = -sinh x := +theorem sin_mul_i : sin (x * I) = sinh x * I := by + have h : I * sin (x * I) = -sinh x := by rw [mul_comm, ← sinh_mul_I] ring_nf @@ -836,7 +845,7 @@ theorem sin_mul_i : sin (x * i) = sinh x * i := by simpa only [neg_mul, div_I, neg_neg] using CancelFactors.cancel_factors_eq_div h I_ne_zero #align complex.sin_mul_I Complex.sin_mul_i -theorem tan_mul_i : tan (x * i) = tanh x * i := by +theorem tan_mul_i : tan (x * I) = tanh x * I := by rw [tan, sin_mul_I, cos_mul_I, mul_div_right_comm, tanh_eq_sinh_div_cosh] #align complex.tan_mul_I Complex.tan_mul_i @@ -870,19 +879,19 @@ theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] #align complex.cos_sub Complex.cos_sub -theorem sin_add_mul_i (x y : ℂ) : sin (x + y * i) = sin x * cosh y + cos x * sinh y * i := by +theorem sin_add_mul_i (x y : ℂ) : sin (x + y * I) = sin x * cosh y + cos x * sinh y * I := by rw [sin_add, cos_mul_I, sin_mul_I, mul_assoc] #align complex.sin_add_mul_I Complex.sin_add_mul_i -theorem sin_eq (z : ℂ) : sin z = sin z.re * cosh z.im + cos z.re * sinh z.im * i := by +theorem sin_eq (z : ℂ) : sin z = sin z.re * cosh z.im + cos z.re * sinh z.im * I := by convert sin_add_mul_I z.re z.im <;> exact (re_add_im z).symm #align complex.sin_eq Complex.sin_eq -theorem cos_add_mul_i (x y : ℂ) : cos (x + y * i) = cos x * cosh y - sin x * sinh y * i := by +theorem cos_add_mul_i (x y : ℂ) : cos (x + y * I) = cos x * cosh y - sin x * sinh y * I := by rw [cos_add, cos_mul_I, sin_mul_I, mul_assoc] #align complex.cos_add_mul_I Complex.cos_add_mul_i -theorem cos_eq (z : ℂ) : cos z = cos z.re * cosh z.im - sin z.re * sinh z.im * i := by +theorem cos_eq (z : ℂ) : cos z = cos z.re * cosh z.im - sin z.re * sinh z.im * I := by convert cos_add_mul_I z.re z.im <;> exact (re_add_im z).symm #align complex.cos_eq Complex.cos_eq @@ -913,7 +922,7 @@ theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) (cos ((x + y) / 2) * cos ((x - y) / 2) + sin ((x + y) / 2) * sin ((x - y) / 2)) := _ _ = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := _ - + · congr <;> field_simp [h2] <;> ring · rw [cos_add, cos_sub] ring @@ -1001,18 +1010,18 @@ theorem tan_of_real_re (x : ℝ) : (tan x).re = Real.tan x := rfl #align complex.tan_of_real_re Complex.tan_of_real_re -theorem cos_add_sin_i : cos x + sin x * i = exp (x * i) := by +theorem cos_add_sin_i : cos x + sin x * I = exp (x * I) := by rw [← cosh_add_sinh, sinh_mul_I, cosh_mul_I] #align complex.cos_add_sin_I Complex.cos_add_sin_i -theorem cos_sub_sin_i : cos x - sin x * i = exp (-x * i) := by +theorem cos_sub_sin_i : cos x - sin x * I = exp (-x * I) := by rw [neg_mul, ← cosh_sub_sinh, sinh_mul_I, cosh_mul_I] #align complex.cos_sub_sin_I Complex.cos_sub_sin_i @[simp] theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := Eq.trans (by rw [cosh_mul_I, sinh_mul_I, mul_pow, I_sq, mul_neg_one, sub_neg_eq_add, add_comm]) - (cosh_sq_sub_sinh_sq (x * i)) + (cosh_sq_sub_sinh_sq (x * I)) #align complex.sin_sq_add_cos_sq Complex.sin_sq_add_cos_sq @[simp] @@ -1070,14 +1079,14 @@ theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by ring #align complex.sin_three_mul Complex.sin_three_mul -theorem exp_mul_i : exp (x * i) = cos x + sin x * i := +theorem exp_mul_i : exp (x * I) = cos x + sin x * I := (cos_add_sin_i _).symm #align complex.exp_mul_I Complex.exp_mul_i -theorem exp_add_mul_i : exp (x + y * i) = exp x * (cos y + sin y * i) := by rw [exp_add, exp_mul_I] +theorem exp_add_mul_i : exp (x + y * I) = exp x * (cos y + sin y * I) := by rw [exp_add, exp_mul_I] #align complex.exp_add_mul_I Complex.exp_add_mul_i -theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im * i) := by +theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im * I) := by rw [← exp_add_mul_I, re_add_im] #align complex.exp_eq_exp_re_mul_sin_add_cos Complex.exp_eq_exp_re_mul_sin_add_cos @@ -1092,18 +1101,18 @@ theorem exp_im : (exp x).im = Real.exp x.re * Real.sin x.im := by #align complex.exp_im Complex.exp_im @[simp] -theorem exp_of_real_mul_i_re (x : ℝ) : (exp (x * i)).re = Real.cos x := by +theorem exp_of_real_mul_i_re (x : ℝ) : (exp (x * I)).re = Real.cos x := by simp [exp_mul_I, cos_of_real_re] #align complex.exp_of_real_mul_I_re Complex.exp_of_real_mul_i_re @[simp] -theorem exp_of_real_mul_i_im (x : ℝ) : (exp (x * i)).im = Real.sin x := by +theorem exp_of_real_mul_i_im (x : ℝ) : (exp (x * I)).im = Real.sin x := by simp [exp_mul_I, sin_of_real_re] #align complex.exp_of_real_mul_I_im Complex.exp_of_real_mul_i_im /-- **De Moivre's formula** -/ theorem cos_add_sin_mul_i_pow (n : ℕ) (z : ℂ) : - (cos z + sin z * i) ^ n = cos (↑n * z) + sin (↑n * z) * i := by + (cos z + sin z * I) ^ n = cos (↑n * z) + sin (↑n * z) * I := by rw [← exp_mul_I, ← exp_mul_I] induction' n with n ih · rw [pow_zero, Nat.cast_zero, zero_mul, zero_mul, exp_zero] @@ -1177,7 +1186,7 @@ theorem cos_neg : cos (-x) = cos x := by simp [cos, exp_neg] @[simp] theorem cos_abs : cos (|x|) = cos x := by - cases le_total x 0 <;> simp only [*, _root_.abs_of_nonneg, abs_of_nonpos, cos_neg] + cases le_total x 0 <;> simp only [*, abs_of_nonneg, abs_of_nonpos, cos_neg] #align real.cos_abs Real.cos_abs theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by @@ -1366,7 +1375,7 @@ theorem cosh_neg : cosh (-x) = cosh x := @[simp] theorem cosh_abs : cosh (|x|) = cosh x := by - cases le_total x 0 <;> simp [*, _root_.abs_of_nonneg, abs_of_nonpos] + cases le_total x 0 <;> simp [*, abs_of_nonneg, abs_of_nonpos] #align real.cosh_abs Real.cosh_abs theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by @@ -1472,7 +1481,7 @@ theorem add_one_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : x + 1 ≤ exp x := rw [← of_real_pow, ← of_real_nat_cast, ← of_real_div, of_real_re] exact div_nonneg (pow_nonneg hx _) (Nat.cast_nonneg _)⟩) _ = exp x := by rw [exp, Complex.exp, ← cau_seq_re, lim_re] - + #align real.add_one_le_exp_of_nonneg Real.add_one_le_exp_of_nonneg theorem one_le_exp {x : ℝ} (hx : 0 ≤ x) : 1 ≤ exp x := by linarith [add_one_le_exp_of_nonneg hx] @@ -1610,7 +1619,7 @@ theorem sum_div_factorial_le {α : Type _} [LinearOrderedField α] (n j : ℕ) ( exact sub_le_self _ (mul_nonneg (Nat.cast_nonneg _) (pow_nonneg (inv_nonneg.2 (Nat.cast_nonneg _)) _)) - + #align complex.sum_div_factorial_le Complex.sum_div_factorial_le theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : @@ -1644,7 +1653,7 @@ theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : simp [abs_mul, abv_pow abs, abs_div, mul_sum.symm] _ ≤ abs x ^ n * (n.succ * (n ! * n)⁻¹) := mul_le_mul_of_nonneg_left (sum_div_factorial_le _ _ hn) (pow_nonneg (abs.nonneg _) _) - + #align complex.exp_bound Complex.exp_bound theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : @@ -1657,15 +1666,15 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : have hj : j = n + k := (add_tsub_cancel_of_le hj).symm rw [hj, sum_range_add_sub_sum_range] calc - abs (∑ i : ℕ in range k, x ^ (n + i) / ((n + i)! : ℂ)) ≤ - ∑ i : ℕ in range k, abs (x ^ (n + i) / ((n + i)! : ℂ)) := + abs (∑ I : ℕ in range k, x ^ (n + I) / ((n + I)! : ℂ)) ≤ + ∑ I : ℕ in range k, abs (x ^ (n + I) / ((n + I)! : ℂ)) := abv_sum_le_sum_abv _ _ - _ ≤ ∑ i : ℕ in range k, abs x ^ (n + i) / (n + i)! := by + _ ≤ ∑ I : ℕ in range k, abs x ^ (n + I) / (n + I)! := by simp only [Complex.abs_cast_nat, map_div₀, abv_pow abs] - _ ≤ ∑ i : ℕ in range k, abs x ^ (n + i) / (n ! * n.succ ^ i) := _ - _ = ∑ i : ℕ in range k, abs x ^ n / n ! * (abs x ^ i / n.succ ^ i) := _ + _ ≤ ∑ I : ℕ in range k, abs x ^ (n + I) / (n ! * n.succ ^ I) := _ + _ = ∑ I : ℕ in range k, abs x ^ n / n ! * (abs x ^ I / n.succ ^ I) := _ _ ≤ abs x ^ n / ↑n ! * 2 := _ - + · refine' sum_le_sum fun m hm => div_le_div (pow_nonneg (abs.nonneg x) (n + m)) le_rfl _ _ · exact_mod_cast mul_pos n.factorial_pos (pow_pos n.succ_pos _) · exact_mod_cast Nat.factorial_mul_pow_le_factorial @@ -1691,7 +1700,7 @@ theorem abs_exp_sub_one_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1) ≤ 2 abs (exp x - 1) = abs (exp x - ∑ m in range 1, x ^ m / m !) := by simp [sum_range_succ] _ ≤ abs x ^ 1 * (Nat.succ 1 * (1! * (1 : ℕ))⁻¹) := (exp_bound hx (by decide)) _ = 2 * abs x := by simp [two_mul, mul_two, mul_add, mul_comm] - + #align complex.abs_exp_sub_one_le Complex.abs_exp_sub_one_le theorem abs_exp_sub_one_sub_id_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1 - x) ≤ abs x ^ 2 := @@ -1701,7 +1710,7 @@ theorem abs_exp_sub_one_sub_id_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1 _ ≤ abs x ^ 2 * (Nat.succ 2 * (2! * (2 : ℕ))⁻¹) := (exp_bound hx (by decide)) _ ≤ abs x ^ 2 * 1 := (mul_le_mul_of_nonneg_left (by norm_num) (sq_nonneg (abs x))) _ = abs x ^ 2 := by rw [mul_one] - + #align complex.abs_exp_sub_one_sub_id_le Complex.abs_exp_sub_one_sub_id_le end Complex @@ -1733,7 +1742,7 @@ theorem abs_exp_sub_one_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1| ≤ 2 * |x| #align real.abs_exp_sub_one_le Real.abs_exp_sub_one_le theorem abs_exp_sub_one_sub_id_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1 - x| ≤ x ^ 2 := by - rw [← _root_.sq_abs] + rw [← sq_abs] have : Complex.abs x ≤ 1 := by exact_mod_cast hx exact_mod_cast Complex.abs_exp_sub_one_sub_id_le this #align real.abs_exp_sub_one_sub_id_le Real.abs_exp_sub_one_sub_id_le @@ -1773,12 +1782,12 @@ theorem exp_approx_succ {n} {x a₁ b₁ : ℝ} (m : ℕ) (e₁ : n + 1 = m) (a (e : |1 + x / m * a₂ - a₁| ≤ b₁ - |x| / m * b₂) (h : |exp x - expNear m x a₂| ≤ |x| ^ m / m ! * b₂) : |exp x - expNear n x a₁| ≤ |x| ^ n / n ! * b₁ := by - refine' (_root_.abs_sub_le _ _ _).trans ((add_le_add_right h _).trans _) - subst e₁; rw [exp_near_succ, exp_near_sub, _root_.abs_mul] + refine' (abs_sub_le _ _ _).trans ((add_le_add_right h _).trans _) + subst e₁; rw [exp_near_succ, exp_near_sub, abs_mul] convert mul_le_mul_of_nonneg_left (le_sub_iff_add_le'.1 e) _ - · simp [mul_add, pow_succ', div_eq_mul_inv, _root_.abs_mul, _root_.abs_inv, ← pow_abs, mul_inv] + · simp [mul_add, pow_succ', div_eq_mul_inv, abs_mul, abs_inv, ← pow_abs, mul_inv] ac_rfl - · simp [_root_.div_nonneg, _root_.abs_nonneg] + · simp [div_nonneg, abs_nonneg] #align real.exp_approx_succ Real.exp_approx_succ theorem exp_approx_end' {n} {x a b : ℝ} (m : ℕ) (e₁ : n + 1 = m) (rm : ℝ) (er : ↑m = rm) @@ -1803,12 +1812,12 @@ theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| calc |cos x - (1 - x ^ 2 / 2)| = abs (Complex.cos x - (1 - x ^ 2 / 2)) := by rw [← abs_of_real] <;> simp [of_real_bit0, of_real_one, of_real_inv] - _ = abs ((Complex.exp (x * i) + Complex.exp (-x * i) - (2 - x ^ 2)) / 2) := by + _ = abs ((Complex.exp (x * I) + Complex.exp (-x * I) - (2 - x ^ 2)) / 2) := by simp [Complex.cos, sub_div, add_div, neg_div, div_self (two_ne_zero' ℂ)] _ = abs - (((Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) + - (Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !)) / + (((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) + + (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !)) / 2) := (congr_arg abs (congr_arg (fun x : ℂ => x / 2) @@ -1817,35 +1826,35 @@ theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| simp [pow_succ] apply Complex.ext <;> simp [div_eq_mul_inv, norm_sq] <;> ring))) _ ≤ - abs ((Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) / 2) + - abs ((Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) / 2) := + abs ((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) / 2) + + abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) / 2) := by rw [add_div] <;> exact complex.abs.add_le _ _ _ = - abs (Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) / 2 + - abs (Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) / 2 := + abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) / 2 + + abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) / 2 := by simp [map_div₀] _ ≤ - Complex.abs (x * i) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + - Complex.abs (-x * i) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 := + Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + + Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 := (add_le_add ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide))) ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide)))) _ ≤ |x| ^ 4 * (5 / 96) := by norm_num <;> simp [mul_assoc, mul_comm, mul_left_comm, mul_div_assoc] - + #align real.cos_bound Real.cos_bound theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| ^ 4 * (5 / 96) := calc |sin x - (x - x ^ 3 / 6)| = abs (Complex.sin x - (x - x ^ 3 / 6)) := by rw [← abs_of_real] <;> simp [of_real_bit0, of_real_one, of_real_inv] - _ = abs (((Complex.exp (-x * i) - Complex.exp (x * i)) * i - (2 * x - x ^ 3 / 3)) / 2) := by + _ = abs (((Complex.exp (-x * I) - Complex.exp (x * I)) * I - (2 * x - x ^ 3 / 3)) / 2) := by simp [Complex.sin, sub_div, add_div, neg_div, mul_div_cancel_left _ (two_ne_zero' ℂ), div_div, show (3 : ℂ) * 2 = 6 by norm_num] _ = abs - (((Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) - - (Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !)) * - i / + (((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) - + (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !)) * + I / 2) := (congr_arg abs (congr_arg (fun x : ℂ => x / 2) @@ -1854,21 +1863,21 @@ theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| simp [pow_succ] apply Complex.ext <;> simp [div_eq_mul_inv, norm_sq] <;> ring))) _ ≤ - abs ((Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) * i / 2) + - abs (-((Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) * i) / 2) := + abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) * I / 2) + + abs (-((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) * I) / 2) := by rw [sub_mul, sub_eq_add_neg, add_div] <;> exact complex.abs.add_le _ _ _ = - abs (Complex.exp (x * i) - ∑ m in range 4, (x * i) ^ m / m !) / 2 + - abs (Complex.exp (-x * i) - ∑ m in range 4, (-x * i) ^ m / m !) / 2 := + abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) / 2 + + abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) / 2 := by simp [add_comm, map_div₀] _ ≤ - Complex.abs (x * i) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + - Complex.abs (-x * i) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 := + Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + + Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 := (add_le_add ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide))) ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide)))) _ ≤ |x| ^ 4 * (5 / 96) := by norm_num <;> simp [mul_assoc, mul_comm, mul_left_comm, mul_div_assoc] - + #align real.sin_bound Real.sin_bound theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := @@ -1881,12 +1890,12 @@ theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := add_le_add (mul_le_mul_of_nonneg_right (pow_le_one _ (abs_nonneg _) hx) (by norm_num)) ((div_le_div_right (by norm_num)).2 (by - rw [sq, ← abs_mul_self, _root_.abs_mul] <;> + rw [sq, ← abs_mul_self, abs_mul] <;> exact mul_le_one hx (abs_nonneg _) hx)) _ < 1 := by norm_num ) _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 - + #align real.cos_pos_of_le_one Real.cos_pos_of_le_one theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < sin x := @@ -1901,8 +1910,8 @@ theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < si (calc |x| ^ 4 ≤ |x| ^ 1 := pow_le_pow_of_le_one (abs_nonneg _) - (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]) (by decide) - _ = x := by simp [_root_.abs_of_nonneg (le_of_lt hx0)] + (by rwa [abs_of_nonneg (le_of_lt hx0)]) (by decide) + _ = x := by simp [abs_of_nonneg (le_of_lt hx0)] ) (by norm_num)) ((div_le_div_right (by norm_num)).2 @@ -1913,8 +1922,8 @@ theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < si _ < x := by linarith ) _ ≤ sin x := - sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]))).2 - + sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [abs_of_nonneg (le_of_lt hx0)]))).2 + #align real.sin_pos_of_pos_of_le_one Real.sin_pos_of_pos_of_le_one theorem sin_pos_of_pos_of_le_two {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 2) : 0 < sin x := @@ -1922,9 +1931,9 @@ theorem sin_pos_of_pos_of_le_two {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 2) : 0 < si calc 0 < 2 * sin (x / 2) * cos (x / 2) := mul_pos (mul_pos (by norm_num) (sin_pos_of_pos_of_le_one (half_pos hx0) this)) - (cos_pos_of_le_one (by rwa [_root_.abs_of_nonneg (le_of_lt (half_pos hx0))])) + (cos_pos_of_le_one (by rwa [abs_of_nonneg (le_of_lt (half_pos hx0))])) _ = sin x := by rw [← sin_two_mul, two_mul, add_halves] - + #align real.sin_pos_of_pos_of_le_two Real.sin_pos_of_pos_of_le_two theorem cos_one_le : cos 1 ≤ 2 / 3 := @@ -1932,7 +1941,7 @@ theorem cos_one_le : cos 1 ≤ 2 / 3 := cos 1 ≤ |(1 : ℝ)| ^ 4 * (5 / 96) + (1 - 1 ^ 2 / 2) := sub_le_iff_le_add.1 (abs_sub_le_iff.1 (cos_bound (by simp))).1 _ ≤ 2 / 3 := by norm_num - + #align real.cos_one_le Real.cos_one_le theorem cos_one_pos : 0 < cos 1 := @@ -1952,7 +1961,7 @@ theorem cos_two_neg : cos 2 < 0 := zero_le_two) _) _ < 0 := by norm_num - + #align real.cos_two_neg Real.cos_two_neg theorem exp_bound_div_one_sub_of_interval_approx {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) : @@ -2051,7 +2060,7 @@ end Tactic namespace Complex @[simp] -theorem abs_cos_add_sin_mul_i (x : ℝ) : abs (cos x + sin x * i) = 1 := by +theorem abs_cos_add_sin_mul_i (x : ℝ) : abs (cos x + sin x * I) = 1 := by have := Real.sin_sq_add_cos_sq x simp_all [add_comm, abs, norm_sq, sq, sin_of_real_re, cos_of_real_re, mul_re] #align complex.abs_cos_add_sin_mul_I Complex.abs_cos_add_sin_mul_i @@ -2062,7 +2071,7 @@ theorem abs_exp_of_real (x : ℝ) : abs (exp x) = Real.exp x := by #align complex.abs_exp_of_real Complex.abs_exp_of_real @[simp] -theorem abs_exp_of_real_mul_i (x : ℝ) : abs (exp (x * i)) = 1 := by +theorem abs_exp_of_real_mul_i (x : ℝ) : abs (exp (x * I)) = 1 := by rw [exp_mul_I, abs_cos_add_sin_mul_I] #align complex.abs_exp_of_real_mul_I Complex.abs_exp_of_real_mul_i @@ -2075,4 +2084,3 @@ theorem abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = #align complex.abs_exp_eq_iff_re_eq Complex.abs_exp_eq_iff_re_eq end Complex - From 15350fd409fce88d3ee687552c708057df66e66b Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 15:17:02 +0000 Subject: [PATCH 07/23] fix some names --- Mathlib/Data/Complex/Basic.lean | 52 ++++++++++++++++----------------- 1 file changed, 26 insertions(+), 26 deletions(-) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index c82c3cef694c9..405bdfac1b500 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -305,32 +305,32 @@ set_option linter.uppercaseLean3 false in #align complex.I Complex.I @[simp] -theorem i_re : I.re = 0 := +theorem I_re : I.re = 0 := rfl set_option linter.uppercaseLean3 false in -#align complex.I_re Complex.i_re +#align complex.I_re Complex.I_re @[simp] -theorem i_im : I.im = 1 := +theorem I_im : I.im = 1 := rfl set_option linter.uppercaseLean3 false in -#align complex.I_im Complex.i_im +#align complex.I_im Complex.I_im @[simp] -theorem i_mul_I : I * I = -1 := +theorem I_mul_I : I * I = -1 := ext_iff.2 <| by simp set_option linter.uppercaseLean3 false in -#align complex.I_mul_I Complex.i_mul_I +#align complex.I_mul_I Complex.I_mul_I -theorem i_mul (z : ℂ) : I * z = ⟨-z.im, z.re⟩ := +theorem I_mul (z : ℂ) : I * z = ⟨-z.im, z.re⟩ := ext_iff.2 <| by simp set_option linter.uppercaseLean3 false in -#align complex.I_mul Complex.i_mul +#align complex.I_mul Complex.I_mul -theorem i_ne_zero : (I : ℂ) ≠ 0 := +theorem I_ne_zero : (I : ℂ) ≠ 0 := mt (congr_arg im) zero_ne_one.symm set_option linter.uppercaseLean3 false in -#align complex.I_ne_zero Complex.i_ne_zero +#align complex.I_ne_zero Complex.I_ne_zero theorem mk_eq_add_mul_I (a b : ℝ) : Complex.mk a b = a + b * I := ext_iff.2 <| by simp [ofReal'] @@ -350,13 +350,13 @@ theorem mul_I_im (z : ℂ) : (z * I).im = z.re := by simp set_option linter.uppercaseLean3 false in #align complex.mul_I_im Complex.mul_I_im -theorem i_mul_re (z : ℂ) : (I * z).re = -z.im := by simp +theorem I_mul_re (z : ℂ) : (I * z).re = -z.im := by simp set_option linter.uppercaseLean3 false in -#align complex.I_mul_re Complex.i_mul_re +#align complex.I_mul_re Complex.I_mul_re -theorem i_mul_im (z : ℂ) : (I * z).im = z.re := by simp +theorem I_mul_im (z : ℂ) : (I * z).im = z.re := by simp set_option linter.uppercaseLean3 false in -#align complex.I_mul_im Complex.i_mul_im +#align complex.I_mul_im Complex.I_mul_im @[simp] theorem equivRealProd_symm_apply (p : ℝ × ℝ) : equivRealProd.symm p = p.1 + p.2 * I := by @@ -469,14 +469,14 @@ theorem coe_imAddGroupHom : (imAddGroupHom : ℂ → ℝ) = im := section set_option linter.deprecated false @[simp] -theorem i_pow_bit0 (n : ℕ) : I ^ bit0 n = (-1) ^ n := by rw [pow_bit0', Complex.i_mul_I] +theorem I_pow_bit0 (n : ℕ) : I ^ bit0 n = (-1) ^ n := by rw [pow_bit0', Complex.I_mul_I] set_option linter.uppercaseLean3 false in -#align complex.I_pow_bit0 Complex.i_pow_bit0 +#align complex.I_pow_bit0 Complex.I_pow_bit0 @[simp] -theorem i_pow_bit1 (n : ℕ) : I ^ bit1 n = (-1) ^ n * I := by rw [pow_bit1', Complex.i_mul_I] +theorem I_pow_bit1 (n : ℕ) : I ^ bit1 n = (-1) ^ n * I := by rw [pow_bit1', Complex.I_mul_I] set_option linter.uppercaseLean3 false in -#align complex.I_pow_bit1 Complex.i_pow_bit1 +#align complex.I_pow_bit1 Complex.I_pow_bit1 end /-! ### Complex conjugation -/ @@ -681,9 +681,9 @@ theorem ofReal_eq_coe (r : ℝ) : ofReal r = r := #align complex.of_real_eq_coe Complex.ofReal_eq_coe @[simp] -theorem i_sq : I ^ 2 = -1 := by rw [sq, i_mul_I] +theorem I_sq : I ^ 2 = -1 := by rw [sq, I_mul_I] set_option linter.uppercaseLean3 false in -#align complex.I_sq Complex.i_sq +#align complex.I_sq Complex.I_sq @[simp] theorem sub_re (z w : ℂ) : (z - w).re = z.re - w.re := @@ -759,14 +759,14 @@ noncomputable instance : Field ℂ := { section set_option linter.deprecated false @[simp] -theorem i_zpow_bit0 (n : ℤ) : I ^ bit0 n = (-1) ^ n := by rw [zpow_bit0', i_mul_I] +theorem I_zpow_bit0 (n : ℤ) : I ^ bit0 n = (-1) ^ n := by rw [zpow_bit0', I_mul_I] set_option linter.uppercaseLean3 false in -#align complex.I_zpow_bit0 Complex.i_zpow_bit0 +#align complex.I_zpow_bit0 Complex.I_zpow_bit0 @[simp] -theorem i_zpow_bit1 (n : ℤ) : I ^ bit1 n = (-1) ^ n * I := by rw [zpow_bit1', i_mul_I] +theorem I_zpow_bit1 (n : ℤ) : I ^ bit1 n = (-1) ^ n * I := by rw [zpow_bit1', I_mul_I] set_option linter.uppercaseLean3 false in -#align complex.I_zpow_bit1 Complex.i_zpow_bit1 +#align complex.I_zpow_bit1 Complex.I_zpow_bit1 end @@ -794,7 +794,7 @@ theorem ofReal_zpow (r : ℝ) (n : ℤ) : ((r ^ n : ℝ) : ℂ) = (r : ℂ) ^ n @[simp] theorem div_I (z : ℂ) : z / I = -(z * I) := - (div_eq_iff_mul_eq i_ne_zero).2 <| by simp [mul_assoc] + (div_eq_iff_mul_eq I_ne_zero).2 <| by simp [mul_assoc] set_option linter.uppercaseLean3 false in #align complex.div_I Complex.div_I @@ -885,7 +885,7 @@ theorem re_eq_add_conj (z : ℂ) : (z.re : ℂ) = (z + conj z) / 2 := by theorem im_eq_sub_conj (z : ℂ) : (z.im : ℂ) = (z - conj z) / (2 * I) := by have : (↑2 : ℝ ) * I = 2 * I := by rfl simp only [sub_conj, ofReal_mul, ofReal_one, ofReal_bit0, mul_right_comm, this, - mul_div_cancel_left _ (mul_ne_zero two_ne_zero i_ne_zero : 2 * I ≠ 0)] + mul_div_cancel_left _ (mul_ne_zero two_ne_zero I_ne_zero : 2 * I ≠ 0)] #align complex.im_eq_sub_conj Complex.im_eq_sub_conj From a5cded87d240458bccfcdbc369c4560174b7f684 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 16:00:18 +0000 Subject: [PATCH 08/23] progress --- Mathlib/Data/Complex/Basic.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 405bdfac1b500..692eb27e1ec85 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -478,6 +478,18 @@ theorem I_pow_bit1 (n : ℕ) : I ^ bit1 n = (-1) ^ n * I := by rw [pow_bit1', Co set_option linter.uppercaseLean3 false in #align complex.I_pow_bit1 Complex.I_pow_bit1 +--Porting note: new ugly theorem +@[simp] +theorem ofReal_two : ((2 : ℝ) : ℂ) = 2 := rfl + +--Porting note: new ugly theorem +@[simp] +theorem ofReal_three : ((3 : ℝ) : ℂ) = 3 := rfl + +--Porting note: new ugly theorem +@[simp] +theorem ofReal_four : ((4 : ℝ) : ℂ) = 4 := rfl + end /-! ### Complex conjugation -/ From 3c23fa1628a40f48c4fd318d5644807f34aa1be5 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 16:10:30 +0000 Subject: [PATCH 09/23] stuck at buggy lean --- Mathlib/Data/Complex/Exponential.lean | 648 +++++++++++++------------- 1 file changed, 332 insertions(+), 316 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index cf78db73a8725..82b7da8bb0109 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -399,7 +399,8 @@ def exp' (z : ℂ) : CauSeq ℂ Complex.abs := /-- The complex exponential function, defined via its Taylor series -/ --@[pp_nodot] Porting note: removed -irreducible_def exp (z : ℂ) : ℂ := +--Porting note: removed `irreducible` attribute, so I can prove things +def exp (z : ℂ) : ℂ := CauSeq.lim (exp' z) #align complex.exp Complex.exp @@ -442,54 +443,58 @@ def tanh (z : ℂ) : ℂ := end end Complex -#print Complex.exp + namespace Real ---open Complex +open Complex + +noncomputable section /-- The real exponential function, defined as the real part of the complex exponential -/ --@[pp_nodot] Porting note: removed -def exp (x : ℝ) : ℝ := - (Complex.exp x).re +nonrec def exp (x : ℝ) : ℝ := + (exp x).re #align real.exp Real.exp /-- The real sine function, defined as the real part of the complex sine -/ --@[pp_nodot] Porting note: removed -def sin (x : ℝ) : ℝ := +nonrec def sin (x : ℝ) : ℝ := (sin x).re #align real.sin Real.sin /-- The real cosine function, defined as the real part of the complex cosine -/ --@[pp_nodot] Porting note: removed -def cos (x : ℝ) : ℝ := +nonrec def cos (x : ℝ) : ℝ := (cos x).re #align real.cos Real.cos /-- The real tangent function, defined as the real part of the complex tangent -/ --@[pp_nodot] Porting note: removed -def tan (x : ℝ) : ℝ := +nonrec def tan (x : ℝ) : ℝ := (tan x).re #align real.tan Real.tan /-- The real hypebolic sine function, defined as the real part of the complex hyperbolic sine -/ --@[pp_nodot] Porting note: removed -def sinh (x : ℝ) : ℝ := +nonrec def sinh (x : ℝ) : ℝ := (sinh x).re #align real.sinh Real.sinh /-- The real hypebolic cosine function, defined as the real part of the complex hyperbolic cosine -/ --@[pp_nodot] Porting note: removed -def cosh (x : ℝ) : ℝ := +nonrec def cosh (x : ℝ) : ℝ := (cosh x).re #align real.cosh Real.cosh /-- The real hypebolic tangent function, defined as the real part of the complex hyperbolic tangent -/ --@[pp_nodot] Porting note: removed -def tanh (x : ℝ) : ℝ := +nonrec def tanh (x : ℝ) : ℝ := (tanh x).re #align real.tanh Real.tanh +end + end Real namespace Complex @@ -501,24 +506,22 @@ theorem exp_zero : exp 0 = 1 := by rw [exp] refine' lim_eq_of_equiv_const fun ε ε0 => ⟨1, fun j hj => _⟩ convert ε0 - cases j + cases' j with j j · exact absurd hj (not_le_of_gt zero_lt_one) · dsimp [exp'] induction' j with j ih - · dsimp [exp'] <;> simp - · rw [← ih (by decide)] + · dsimp [exp']; simp [show Nat.succ 0 = 1 from rfl] + · rw [← ih (by simp [show 1 = Nat.succ 0 from rfl, Nat.succ_le_succ])] simp only [sum_range_succ, pow_succ] simp #align complex.exp_zero Complex.exp_zero theorem exp_add : exp (x + y) = exp x * exp y := by - 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)!) := - by + have hj : ∀ j : ℕ, (∑ m in range j, (x + y) ^ m / m.factorial) = + ∑ i in range j, ∑ k in range (i + 1), x ^ k / k.factorial * + (y ^ (i - k) / (i - k).factorial) := by intro j - refine' Finset.sum_congr rfl fun m hm => _ + refine' Finset.sum_congr rfl fun m _ => _ rw [add_pow, div_eq_mul_inv, sum_mul] refine' Finset.sum_congr rfl fun I hi => _ have h₁ : (m.choose I : ℂ) ≠ 0 := @@ -532,33 +535,40 @@ theorem exp_add : exp (x + y) = exp x * exp y := by 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) + exact cauchy_product (isCauSeq_abs_exp x) (isCauSeq_exp y) #align complex.exp_add Complex.exp_add -theorem exp_list_sum (l : List ℂ) : exp l.Sum = (l.map exp).Prod := - @MonoidHom.map_list_prod (Multiplicative ℂ) ℂ _ _ ⟨exp, exp_zero, exp_add⟩ l +--Porting note: New definition +/-- the exponential function as a monoid hom from `Multiplicative ℂ` to `ℂ` -/ +noncomputable def expMonoidHom : MonoidHom (Multiplicative ℂ) ℂ := + { toFun := fun z => exp (Multiplicative.toAdd z), + map_one' := by simp, + map_mul' := by simp [exp_add] } + +theorem exp_list_sum (l : List ℂ) : exp l.sum = (l.map exp).prod := + @MonoidHom.map_list_prod (Multiplicative ℂ) ℂ _ _ expMonoidHom l #align complex.exp_list_sum Complex.exp_list_sum -theorem exp_multiset_sum (s : Multiset ℂ) : exp s.Sum = (s.map exp).Prod := - @MonoidHom.map_multiset_prod (Multiplicative ℂ) ℂ _ _ ⟨exp, exp_zero, exp_add⟩ s +theorem exp_multiset_sum (s : Multiset ℂ) : exp s.sum = (s.map exp).prod := + @MonoidHom.map_multiset_prod (Multiplicative ℂ) ℂ _ _ expMonoidHom s #align complex.exp_multiset_sum Complex.exp_multiset_sum theorem exp_sum {α : Type _} (s : Finset α) (f : α → ℂ) : exp (∑ x in s, f x) = ∏ x in s, exp (f x) := - @MonoidHom.map_prod (Multiplicative ℂ) α ℂ _ _ ⟨exp, exp_zero, exp_add⟩ f s + @map_prod (Multiplicative ℂ) α ℂ _ _ _ _ expMonoidHom f s #align complex.exp_sum Complex.exp_sum theorem exp_nat_mul (x : ℂ) : ∀ n : ℕ, exp (n * x) = exp x ^ n | 0 => by rw [Nat.cast_zero, zero_mul, exp_zero, pow_zero] - | Nat.succ n => by rw [pow_succ', Nat.cast_add_one, add_mul, exp_add, ← exp_nat_mul, one_mul] + | Nat.succ n => by rw [pow_succ', Nat.cast_add_one, add_mul, exp_add, ← exp_nat_mul _ n, one_mul] #align complex.exp_nat_mul Complex.exp_nat_mul theorem exp_ne_zero : exp x ≠ 0 := fun h => - zero_ne_one <| by rw [← exp_zero, ← add_neg_self x, exp_add, h] <;> simp + zero_ne_one <| by rw [← exp_zero, ← add_neg_self x, exp_add, h]; simp #align complex.exp_ne_zero Complex.exp_ne_zero theorem exp_neg : exp (-x) = (exp x)⁻¹ := by - rw [← mul_right_inj' (exp_ne_zero x), ← exp_add] <;> simp [mul_inv_cancel (exp_ne_zero x)] + rw [← mul_right_inj' (exp_ne_zero x), ← exp_add]; simp [mul_inv_cancel (exp_ne_zero x)] #align complex.exp_neg Complex.exp_neg theorem exp_sub : exp (x - y) = exp x / exp y := by @@ -567,38 +577,38 @@ theorem exp_sub : exp (x - y) = exp x / exp y := by theorem exp_int_mul (z : ℂ) (n : ℤ) : Complex.exp (n * z) = Complex.exp z ^ n := by cases n - · apply Complex.exp_nat_mul - · simpa [Complex.exp_neg, add_comm, ← neg_mul] using Complex.exp_nat_mul (-z) (1 + n) + . simp [exp_nat_mul] + . simp [exp_add, add_mul, pow_add, exp_neg, exp_nat_mul] #align complex.exp_int_mul Complex.exp_int_mul @[simp] theorem exp_conj : exp (conj x) = conj (exp x) := by dsimp [exp] rw [← lim_conj] - refine' congr_arg limUnder (CauSeq.ext fun _ => _) - dsimp [exp', Function.comp, cau_seq_conj] + refine' congr_arg CauSeq.lim (CauSeq.ext fun _ => _) + dsimp [exp', Function.comp, isCauSeq_conj, cauSeqConj] rw [(starRingEnd _).map_sum] - refine' sum_congr rfl fun n hn => _ - rw [map_div₀, map_pow, ← of_real_nat_cast, conj_of_real] + refine' sum_congr rfl fun n _ => _ + rw [map_div₀, map_pow, ← ofReal_nat_cast, conj_ofReal] #align complex.exp_conj Complex.exp_conj @[simp] -theorem of_real_exp_of_real_re (x : ℝ) : ((exp x).re : ℂ) = exp x := - eq_conj_iff_re.1 <| by rw [← exp_conj, conj_of_real] -#align complex.of_real_exp_of_real_re Complex.of_real_exp_of_real_re +theorem ofReal_exp_ofReal_re (x : ℝ) : ((exp x).re : ℂ) = exp x := + eq_conj_iff_re.1 <| by rw [← exp_conj, conj_ofReal] +#align complex.of_real_exp_of_real_re Complex.ofReal_exp_ofReal_re @[simp, norm_cast] -theorem of_real_exp (x : ℝ) : (Real.exp x : ℂ) = exp x := - of_real_exp_of_real_re _ -#align complex.of_real_exp Complex.of_real_exp +theorem ofReal_exp (x : ℝ) : (Real.exp x : ℂ) = exp x := + ofReal_exp_ofReal_re _ +#align complex.of_real_exp Complex.ofReal_exp @[simp] -theorem exp_of_real_im (x : ℝ) : (exp x).im = 0 := by rw [← of_real_exp_of_real_re, of_real_im] -#align complex.exp_of_real_im Complex.exp_of_real_im +theorem exp_ofReal_im (x : ℝ) : (exp x).im = 0 := by rw [← ofReal_exp_ofReal_re, ofReal_im] +#align complex.exp_of_real_im Complex.exp_ofReal_im -theorem exp_of_real_re (x : ℝ) : (exp x).re = Real.exp x := +theorem exp_ofReal_re (x : ℝ) : (exp x).re = Real.exp x := rfl -#align complex.exp_of_real_re Complex.exp_of_real_re +#align complex.exp_of_real_re Complex.exp_ofReal_re theorem two_sinh : 2 * sinh x = exp x - exp (-x) := mul_div_cancel' _ two_ne_zero @@ -618,7 +628,6 @@ theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _). private theorem sinh_add_aux {a b c d : ℂ} : (a - b) * (c + d) + (a + b) * (c - d) = 2 * (a * c - b * d) := by ring -#align complex.sinh_add_aux complex.sinh_add_aux theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← @@ -637,7 +646,6 @@ theorem cosh_neg : cosh (-x) = cosh x := by simp [add_comm, cosh, exp_neg] private theorem cosh_add_aux {a b c d : ℂ} : (a + b) * (c + d) + (a - b) * (c - d) = 2 * (a * c + b * d) := by ring -#align complex.cosh_add_aux complex.cosh_add_aux theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, exp_add, neg_add, exp_add, eq_comm, mul_add, ← @@ -655,22 +663,23 @@ theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by #align complex.cosh_sub Complex.cosh_sub theorem sinh_conj : sinh (conj x) = conj (sinh x) := by - rw [sinh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_sub, sinh, map_div₀, conj_bit0, - RingHom.map_one] + rw [sinh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_sub, sinh, map_div₀] + --Porting note: not nice + simp [← one_add_one_eq_two] #align complex.sinh_conj Complex.sinh_conj @[simp] -theorem of_real_sinh_of_real_re (x : ℝ) : ((sinh x).re : ℂ) = sinh x := - eq_conj_iff_re.1 <| by rw [← sinh_conj, conj_of_real] -#align complex.of_real_sinh_of_real_re Complex.of_real_sinh_of_real_re +theorem ofReal_sinh_ofReal_re (x : ℝ) : ((sinh x).re : ℂ) = sinh x := + eq_conj_iff_re.1 <| by rw [← sinh_conj, conj_ofReal] +#align complex.of_real_sinh_of_real_re Complex.ofReal_sinh_ofReal_re @[simp, norm_cast] -theorem of_real_sinh (x : ℝ) : (Real.sinh x : ℂ) = sinh x := - of_real_sinh_of_real_re _ -#align complex.of_real_sinh Complex.of_real_sinh +theorem ofReal_sinh (x : ℝ) : (Real.sinh x : ℂ) = sinh x := + ofReal_sinh_ofReal_re _ +#align complex.of_real_sinh Complex.ofReal_sinh @[simp] -theorem sinh_of_real_im (x : ℝ) : (sinh x).im = 0 := by rw [← of_real_sinh_of_real_re, of_real_im] +theorem sinh_of_real_im (x : ℝ) : (sinh x).im = 0 := by rw [← ofReal_sinh_ofReal_re, ofReal_im] #align complex.sinh_of_real_im Complex.sinh_of_real_im theorem sinh_of_real_re (x : ℝ) : (sinh x).re = Real.sinh x := @@ -678,27 +687,28 @@ theorem sinh_of_real_re (x : ℝ) : (sinh x).re = Real.sinh x := #align complex.sinh_of_real_re Complex.sinh_of_real_re theorem cosh_conj : cosh (conj x) = conj (cosh x) := by - rw [cosh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_add, cosh, map_div₀, conj_bit0, - RingHom.map_one] + rw [cosh, ← RingHom.map_neg, exp_conj, exp_conj, ← RingHom.map_add, cosh, map_div₀] + --Porting note: not nice + simp [← one_add_one_eq_two] #align complex.cosh_conj Complex.cosh_conj -theorem of_real_cosh_of_real_re (x : ℝ) : ((cosh x).re : ℂ) = cosh x := - eq_conj_iff_re.1 <| by rw [← cosh_conj, conj_of_real] -#align complex.of_real_cosh_of_real_re Complex.of_real_cosh_of_real_re +theorem ofReal_cosh_ofReal_re (x : ℝ) : ((cosh x).re : ℂ) = cosh x := + eq_conj_iff_re.1 <| by rw [← cosh_conj, conj_ofReal] +#align complex.of_real_cosh_of_real_re Complex.ofReal_cosh_ofReal_re @[simp, norm_cast] -theorem of_real_cosh (x : ℝ) : (Real.cosh x : ℂ) = cosh x := - of_real_cosh_of_real_re _ -#align complex.of_real_cosh Complex.of_real_cosh +theorem ofReal_cosh (x : ℝ) : (Real.cosh x : ℂ) = cosh x := + ofReal_cosh_ofReal_re _ +#align complex.of_real_cosh Complex.ofReal_cosh @[simp] -theorem cosh_of_real_im (x : ℝ) : (cosh x).im = 0 := by rw [← of_real_cosh_of_real_re, of_real_im] -#align complex.cosh_of_real_im Complex.cosh_of_real_im +theorem cosh_ofReal_im (x : ℝ) : (cosh x).im = 0 := by rw [← ofReal_cosh_ofReal_re, ofReal_im] +#align complex.cosh_of_real_im Complex.cosh_ofReal_im @[simp] -theorem cosh_of_real_re (x : ℝ) : (cosh x).re = Real.cosh x := +theorem cosh_ofReal_re (x : ℝ) : (cosh x).re = Real.cosh x := rfl -#align complex.cosh_of_real_re Complex.cosh_of_real_re +#align complex.cosh_of_real_re Complex.cosh_ofReal_re theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := rfl @@ -717,22 +727,22 @@ theorem tanh_conj : tanh (conj x) = conj (tanh x) := by #align complex.tanh_conj Complex.tanh_conj @[simp] -theorem of_real_tanh_of_real_re (x : ℝ) : ((tanh x).re : ℂ) = tanh x := - eq_conj_iff_re.1 <| by rw [← tanh_conj, conj_of_real] -#align complex.of_real_tanh_of_real_re Complex.of_real_tanh_of_real_re +theorem ofReal_tanh_ofReal_re (x : ℝ) : ((tanh x).re : ℂ) = tanh x := + eq_conj_iff_re.1 <| by rw [← tanh_conj, conj_ofReal] +#align complex.of_real_tanh_of_real_re Complex.ofReal_tanh_ofReal_re @[simp, norm_cast] -theorem of_real_tanh (x : ℝ) : (Real.tanh x : ℂ) = tanh x := - of_real_tanh_of_real_re _ -#align complex.of_real_tanh Complex.of_real_tanh +theorem ofReal_tanh (x : ℝ) : (Real.tanh x : ℂ) = tanh x := + ofReal_tanh_ofReal_re _ +#align complex.of_real_tanh Complex.ofReal_tanh @[simp] -theorem tanh_of_real_im (x : ℝ) : (tanh x).im = 0 := by rw [← of_real_tanh_of_real_re, of_real_im] -#align complex.tanh_of_real_im Complex.tanh_of_real_im +theorem tanh_ofReal_im (x : ℝ) : (tanh x).im = 0 := by rw [← ofReal_tanh_ofReal_re, ofReal_im] +#align complex.tanh_of_real_im Complex.tanh_ofReal_im -theorem tanh_of_real_re (x : ℝ) : (tanh x).re = Real.tanh x := +theorem tanh_ofReal_re (x : ℝ) : (tanh x).re = Real.tanh x := rfl -#align complex.tanh_of_real_re Complex.tanh_of_real_re +#align complex.tanh_of_real_re Complex.tanh_ofReal_re @[simp] theorem cosh_add_sinh : cosh x + sinh x = exp x := by @@ -820,34 +830,41 @@ theorem two_cos : 2 * cos x = exp (x * I) + exp (-x * I) := mul_div_cancel' _ two_ne_zero #align complex.two_cos Complex.two_cos -theorem sinh_mul_i : sinh (x * I) = sin x * I := by +theorem sinh_mul_I : sinh (x * I) = sin x * I := by rw [← mul_right_inj' (two_ne_zero' ℂ), two_sinh, ← mul_assoc, two_sin, mul_assoc, I_mul_I, mul_neg_one, neg_sub, neg_mul_eq_neg_mul] -#align complex.sinh_mul_I Complex.sinh_mul_i +set_option linter.uppercaseLean3 false in +#align complex.sinh_mul_I Complex.sinh_mul_I -theorem cosh_mul_i : cosh (x * I) = cos x := by +theorem cosh_mul_I : cosh (x * I) = cos x := by rw [← mul_right_inj' (two_ne_zero' ℂ), two_cosh, two_cos, neg_mul_eq_neg_mul] -#align complex.cosh_mul_I Complex.cosh_mul_i +set_option linter.uppercaseLean3 false in +#align complex.cosh_mul_I Complex.cosh_mul_I -theorem tanh_mul_i : tanh (x * I) = tan x * I := by +theorem tanh_mul_I : tanh (x * I) = tan x * I := by rw [tanh_eq_sinh_div_cosh, cosh_mul_I, sinh_mul_I, mul_div_right_comm, tan] -#align complex.tanh_mul_I Complex.tanh_mul_i +set_option linter.uppercaseLean3 false in +#align complex.tanh_mul_I Complex.tanh_mul_I -theorem cos_mul_i : cos (x * I) = cosh x := by rw [← cosh_mul_I] <;> ring_nf <;> simp -#align complex.cos_mul_I Complex.cos_mul_i +theorem cos_mul_I : cos (x * I) = cosh x := by rw [← cosh_mul_I]; ring_nf; simp +set_option linter.uppercaseLean3 false in +#align complex.cos_mul_I Complex.cos_mul_I -theorem sin_mul_i : sin (x * I) = sinh x * I := by +theorem sin_mul_I : sin (x * I) = sinh x * I := by have h : I * sin (x * I) = -sinh x := by rw [mul_comm, ← sinh_mul_I] ring_nf simp - simpa only [neg_mul, div_I, neg_neg] using CancelFactors.cancel_factors_eq_div h I_ne_zero -#align complex.sin_mul_I Complex.sin_mul_i + rw [← neg_neg (sinh x), ← h] + ext <;> simp +set_option linter.uppercaseLean3 false in +#align complex.sin_mul_I Complex.sin_mul_I -theorem tan_mul_i : tan (x * I) = tanh x * I := by +theorem tan_mul_I : tan (x * I) = tanh x * I := by rw [tan, sin_mul_I, cos_mul_I, mul_div_right_comm, tanh_eq_sinh_div_cosh] -#align complex.tan_mul_I Complex.tan_mul_i +set_option linter.uppercaseLean3 false in +#align complex.tan_mul_I Complex.tan_mul_I theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := by rw [← mul_left_inj' I_ne_zero, ← sinh_mul_I, add_mul, add_mul, mul_right_comm, ← sinh_mul_I, @@ -864,7 +881,6 @@ theorem cos_neg : cos (-x) = cos x := by simp [cos, sub_eq_add_neg, exp_neg, add private theorem cos_add_aux {a b c d : ℂ} : (a + b) * (c + d) - (b - a) * (d - c) * -1 = 2 * (a * c + b * d) := by ring -#align complex.cos_add_aux complex.cos_add_aux theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by rw [← cosh_mul_I, add_mul, cosh_add, cosh_mul_I, cosh_mul_I, sinh_mul_I, sinh_mul_I, @@ -879,20 +895,22 @@ theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] #align complex.cos_sub Complex.cos_sub -theorem sin_add_mul_i (x y : ℂ) : sin (x + y * I) = sin x * cosh y + cos x * sinh y * I := by +theorem sin_add_mul_I (x y : ℂ) : sin (x + y * I) = sin x * cosh y + cos x * sinh y * I := by rw [sin_add, cos_mul_I, sin_mul_I, mul_assoc] -#align complex.sin_add_mul_I Complex.sin_add_mul_i +set_option linter.uppercaseLean3 false in +#align complex.sin_add_mul_I Complex.sin_add_mul_I theorem sin_eq (z : ℂ) : sin z = sin z.re * cosh z.im + cos z.re * sinh z.im * I := by - convert sin_add_mul_I z.re z.im <;> exact (re_add_im z).symm + convert sin_add_mul_I z.re z.im; exact (re_add_im z).symm #align complex.sin_eq Complex.sin_eq -theorem cos_add_mul_i (x y : ℂ) : cos (x + y * I) = cos x * cosh y - sin x * sinh y * I := by +theorem cos_add_mul_I (x y : ℂ) : cos (x + y * I) = cos x * cosh y - sin x * sinh y * I := by rw [cos_add, cos_mul_I, sin_mul_I, mul_assoc] -#align complex.cos_add_mul_I Complex.cos_add_mul_i +set_option linter.uppercaseLean3 false in +#align complex.cos_add_mul_I Complex.cos_add_mul_I theorem cos_eq (z : ℂ) : cos z = cos z.re * cosh z.im - sin z.re * sinh z.im * I := by - convert cos_add_mul_I z.re z.im <;> exact (re_add_im z).symm + convert cos_add_mul_I z.re z.im; exact (re_add_im z).symm #align complex.cos_eq Complex.cos_eq theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := by @@ -916,14 +934,14 @@ theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := by have h2 : (2 : ℂ) ≠ 0 := by norm_num calc - cos x + cos y = cos ((x + y) / 2 + (x - y) / 2) + cos ((x + y) / 2 - (x - y) / 2) := _ + cos x + cos y = cos ((x + y) / 2 + (x - y) / 2) + cos ((x + y) / 2 - (x - y) / 2) := ?_ _ = cos ((x + y) / 2) * cos ((x - y) / 2) - sin ((x + y) / 2) * sin ((x - y) / 2) + (cos ((x + y) / 2) * cos ((x - y) / 2) + sin ((x + y) / 2) * sin ((x - y) / 2)) := - _ - _ = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := _ + ?_ + _ = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := ?_ - · congr <;> field_simp [h2] <;> ring + · congr <;> field_simp [h2] · rw [cos_add, cos_sub] ring #align complex.cos_add_cos Complex.cos_add_cos @@ -934,44 +952,44 @@ theorem sin_conj : sin (conj x) = conj (sin x) := by #align complex.sin_conj Complex.sin_conj @[simp] -theorem of_real_sin_of_real_re (x : ℝ) : ((sin x).re : ℂ) = sin x := - eq_conj_iff_re.1 <| by rw [← sin_conj, conj_of_real] -#align complex.of_real_sin_of_real_re Complex.of_real_sin_of_real_re +theorem ofReal_sin_ofReal_re (x : ℝ) : ((sin x).re : ℂ) = sin x := + eq_conj_iff_re.1 <| by rw [← sin_conj, conj_ofReal] +#align complex.of_real_sin_of_real_re Complex.ofReal_sin_ofReal_re @[simp, norm_cast] -theorem of_real_sin (x : ℝ) : (Real.sin x : ℂ) = sin x := - of_real_sin_of_real_re _ -#align complex.of_real_sin Complex.of_real_sin +theorem ofReal_sin (x : ℝ) : (Real.sin x : ℂ) = sin x := + ofReal_sin_ofReal_re _ +#align complex.of_real_sin Complex.ofReal_sin @[simp] -theorem sin_of_real_im (x : ℝ) : (sin x).im = 0 := by rw [← of_real_sin_of_real_re, of_real_im] -#align complex.sin_of_real_im Complex.sin_of_real_im +theorem sin_ofReal_im (x : ℝ) : (sin x).im = 0 := by rw [← ofReal_sin_ofReal_re, ofReal_im] +#align complex.sin_of_real_im Complex.sin_ofReal_im -theorem sin_of_real_re (x : ℝ) : (sin x).re = Real.sin x := +theorem sin_ofReal_re (x : ℝ) : (sin x).re = Real.sin x := rfl -#align complex.sin_of_real_re Complex.sin_of_real_re +#align complex.sin_of_real_re Complex.sin_ofReal_re theorem cos_conj : cos (conj x) = conj (cos x) := by rw [← cosh_mul_I, ← conj_neg_I, ← RingHom.map_mul, ← cosh_mul_I, cosh_conj, mul_neg, cosh_neg] #align complex.cos_conj Complex.cos_conj @[simp] -theorem of_real_cos_of_real_re (x : ℝ) : ((cos x).re : ℂ) = cos x := - eq_conj_iff_re.1 <| by rw [← cos_conj, conj_of_real] -#align complex.of_real_cos_of_real_re Complex.of_real_cos_of_real_re +theorem ofReal_cos_ofReal_re (x : ℝ) : ((cos x).re : ℂ) = cos x := + eq_conj_iff_re.1 <| by rw [← cos_conj, conj_ofReal] +#align complex.of_real_cos_of_real_re Complex.ofReal_cos_ofReal_re @[simp, norm_cast] -theorem of_real_cos (x : ℝ) : (Real.cos x : ℂ) = cos x := - of_real_cos_of_real_re _ -#align complex.of_real_cos Complex.of_real_cos +theorem ofReal_cos (x : ℝ) : (Real.cos x : ℂ) = cos x := + ofReal_cos_ofReal_re _ +#align complex.of_real_cos Complex.ofReal_cos @[simp] -theorem cos_of_real_im (x : ℝ) : (cos x).im = 0 := by rw [← of_real_cos_of_real_re, of_real_im] -#align complex.cos_of_real_im Complex.cos_of_real_im +theorem cos_ofReal_im (x : ℝ) : (cos x).im = 0 := by rw [← ofReal_cos_ofReal_re, ofReal_im] +#align complex.cos_of_real_im Complex.cos_ofReal_im -theorem cos_of_real_re (x : ℝ) : (cos x).re = Real.cos x := +theorem cos_ofReal_re (x : ℝ) : (cos x).re = Real.cos x := rfl -#align complex.cos_of_real_re Complex.cos_of_real_re +#align complex.cos_of_real_re Complex.cos_ofReal_re @[simp] theorem tan_zero : tan 0 = 0 := by simp [tan] @@ -993,30 +1011,32 @@ theorem tan_conj : tan (conj x) = conj (tan x) := by rw [tan, sin_conj, cos_conj #align complex.tan_conj Complex.tan_conj @[simp] -theorem of_real_tan_of_real_re (x : ℝ) : ((tan x).re : ℂ) = tan x := - eq_conj_iff_re.1 <| by rw [← tan_conj, conj_of_real] -#align complex.of_real_tan_of_real_re Complex.of_real_tan_of_real_re +theorem ofReal_tan_ofReal_re (x : ℝ) : ((tan x).re : ℂ) = tan x := + eq_conj_iff_re.1 <| by rw [← tan_conj, conj_ofReal] +#align complex.of_real_tan_of_real_re Complex.ofReal_tan_ofReal_re @[simp, norm_cast] -theorem of_real_tan (x : ℝ) : (Real.tan x : ℂ) = tan x := - of_real_tan_of_real_re _ -#align complex.of_real_tan Complex.of_real_tan +theorem ofReal_tan (x : ℝ) : (Real.tan x : ℂ) = tan x := + ofReal_tan_ofReal_re _ +#align complex.of_real_tan Complex.ofReal_tan @[simp] -theorem tan_of_real_im (x : ℝ) : (tan x).im = 0 := by rw [← of_real_tan_of_real_re, of_real_im] +theorem tan_of_real_im (x : ℝ) : (tan x).im = 0 := by rw [← ofReal_tan_ofReal_re, ofReal_im] #align complex.tan_of_real_im Complex.tan_of_real_im theorem tan_of_real_re (x : ℝ) : (tan x).re = Real.tan x := rfl #align complex.tan_of_real_re Complex.tan_of_real_re -theorem cos_add_sin_i : cos x + sin x * I = exp (x * I) := by +theorem cos_add_sin_I : cos x + sin x * I = exp (x * I) := by rw [← cosh_add_sinh, sinh_mul_I, cosh_mul_I] -#align complex.cos_add_sin_I Complex.cos_add_sin_i +set_option linter.uppercaseLean3 false in +#align complex.cos_add_sin_I Complex.cos_add_sin_I -theorem cos_sub_sin_i : cos x - sin x * I = exp (-x * I) := by +theorem cos_sub_sin_I : cos x - sin x * I = exp (-x * I) := by rw [neg_mul, ← cosh_sub_sinh, sinh_mul_I, cosh_mul_I] -#align complex.cos_sub_sin_I Complex.cos_sub_sin_i +set_option linter.uppercaseLean3 false in +#align complex.cos_sub_sin_I Complex.cos_sub_sin_I @[simp] theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := @@ -1079,12 +1099,14 @@ theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by ring #align complex.sin_three_mul Complex.sin_three_mul -theorem exp_mul_i : exp (x * I) = cos x + sin x * I := - (cos_add_sin_i _).symm -#align complex.exp_mul_I Complex.exp_mul_i +theorem exp_mul_I : exp (x * I) = cos x + sin x * I := + (cos_add_sin_I _).symm +set_option linter.uppercaseLean3 false in +#align complex.exp_mul_I Complex.exp_mul_I -theorem exp_add_mul_i : exp (x + y * I) = exp x * (cos y + sin y * I) := by rw [exp_add, exp_mul_I] -#align complex.exp_add_mul_I Complex.exp_add_mul_i +theorem exp_add_mul_I : exp (x + y * I) = exp x * (cos y + sin y * I) := by rw [exp_add, exp_mul_I] +set_option linter.uppercaseLean3 false in +#align complex.exp_add_mul_I Complex.exp_add_mul_I theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im * I) := by rw [← exp_add_mul_I, re_add_im] @@ -1092,32 +1114,35 @@ theorem exp_eq_exp_re_mul_sin_add_cos : exp x = exp x.re * (cos x.im + sin x.im theorem exp_re : (exp x).re = Real.exp x.re * Real.cos x.im := by rw [exp_eq_exp_re_mul_sin_add_cos] - simp [exp_of_real_re, cos_of_real_re] + simp [exp_ofReal_re, cos_ofReal_re] #align complex.exp_re Complex.exp_re theorem exp_im : (exp x).im = Real.exp x.re * Real.sin x.im := by rw [exp_eq_exp_re_mul_sin_add_cos] - simp [exp_of_real_re, sin_of_real_re] + simp [exp_ofReal_re, sin_ofReal_re] #align complex.exp_im Complex.exp_im @[simp] -theorem exp_of_real_mul_i_re (x : ℝ) : (exp (x * I)).re = Real.cos x := by - simp [exp_mul_I, cos_of_real_re] -#align complex.exp_of_real_mul_I_re Complex.exp_of_real_mul_i_re +theorem exp_ofReal_mul_I_re (x : ℝ) : (exp (x * I)).re = Real.cos x := by + simp [exp_mul_I, cos_ofReal_re] +set_option linter.uppercaseLean3 false in +#align complex.exp_of_real_mul_I_re Complex.exp_ofReal_mul_I_re @[simp] -theorem exp_of_real_mul_i_im (x : ℝ) : (exp (x * I)).im = Real.sin x := by - simp [exp_mul_I, sin_of_real_re] -#align complex.exp_of_real_mul_I_im Complex.exp_of_real_mul_i_im +theorem exp_ofReal_mul_I_im (x : ℝ) : (exp (x * I)).im = Real.sin x := by + simp [exp_mul_I, sin_ofReal_re] +set_option linter.uppercaseLean3 false in +#align complex.exp_of_real_mul_I_im Complex.exp_ofReal_mul_I_im /-- **De Moivre's formula** -/ -theorem cos_add_sin_mul_i_pow (n : ℕ) (z : ℂ) : +theorem cos_add_sin_mul_I_pow (n : ℕ) (z : ℂ) : (cos z + sin z * I) ^ n = cos (↑n * z) + sin (↑n * z) * I := by rw [← exp_mul_I, ← exp_mul_I] induction' n with n ih · rw [pow_zero, Nat.cast_zero, zero_mul, zero_mul, exp_zero] · rw [pow_succ', ih, Nat.cast_succ, add_mul, add_mul, one_mul, exp_add] -#align complex.cos_add_sin_mul_I_pow Complex.cos_add_sin_mul_i_pow +set_option linter.uppercaseLean3 false in +#align complex.cos_add_sin_mul_I_pow Complex.cos_add_sin_mul_I_pow end Complex @@ -1131,33 +1156,39 @@ variable (x y : ℝ) theorem exp_zero : exp 0 = 1 := by simp [Real.exp] #align real.exp_zero Real.exp_zero -theorem exp_add : exp (x + y) = exp x * exp y := by simp [exp_add, exp] +nonrec theorem exp_add : exp (x + y) = exp x * exp y := by simp [exp_add, exp] #align real.exp_add Real.exp_add -theorem exp_list_sum (l : List ℝ) : exp l.Sum = (l.map exp).Prod := - @MonoidHom.map_list_prod (Multiplicative ℝ) ℝ _ _ ⟨exp, exp_zero, exp_add⟩ l +--Porting note: New definition +/-- the exponential function as a monoid hom from `Multiplicative ℝ` to `ℝ` -/ +noncomputable def expMonoidHom : MonoidHom (Multiplicative ℝ) ℝ := + { toFun := fun x => exp (Multiplicative.toAdd x), + map_one' := by simp, + map_mul' := by simp [exp_add] } + +theorem exp_list_sum (l : List ℝ) : exp l.sum = (l.map exp).prod := + @MonoidHom.map_list_prod (Multiplicative ℝ) ℝ _ _ expMonoidHom l #align real.exp_list_sum Real.exp_list_sum -theorem exp_multiset_sum (s : Multiset ℝ) : exp s.Sum = (s.map exp).Prod := - @MonoidHom.map_multiset_prod (Multiplicative ℝ) ℝ _ _ ⟨exp, exp_zero, exp_add⟩ s +theorem exp_multiset_sum (s : Multiset ℝ) : exp s.sum = (s.map exp).prod := + @MonoidHom.map_multiset_prod (Multiplicative ℝ) ℝ _ _ expMonoidHom s #align real.exp_multiset_sum Real.exp_multiset_sum theorem exp_sum {α : Type _} (s : Finset α) (f : α → ℝ) : exp (∑ x in s, f x) = ∏ x in s, exp (f x) := - @MonoidHom.map_prod (Multiplicative ℝ) α ℝ _ _ ⟨exp, exp_zero, exp_add⟩ f s + @map_prod (Multiplicative ℝ) α ℝ _ _ _ _ expMonoidHom f s #align real.exp_sum Real.exp_sum -theorem exp_nat_mul (x : ℝ) : ∀ n : ℕ, exp (n * x) = exp x ^ n - | 0 => by rw [Nat.cast_zero, zero_mul, exp_zero, pow_zero] - | Nat.succ n => by rw [pow_succ', Nat.cast_add_one, add_mul, exp_add, ← exp_nat_mul, one_mul] +nonrec theorem exp_nat_mul (x : ℝ) (n : ℕ) : exp (n * x) = exp x ^ n := + ofReal_injective (by simp [exp_nat_mul]) #align real.exp_nat_mul Real.exp_nat_mul -theorem exp_ne_zero : exp x ≠ 0 := fun h => - exp_ne_zero x <| by rw [exp, ← of_real_inj] at h <;> simp_all +nonrec theorem exp_ne_zero : exp x ≠ 0 := fun h => + exp_ne_zero x <| by rw [exp, ← ofReal_inj] at h; simp_all #align real.exp_ne_zero Real.exp_ne_zero -theorem exp_neg : exp (-x) = (exp x)⁻¹ := by - rw [← of_real_inj, exp, of_real_exp_of_real_re, of_real_neg, exp_neg, of_real_inv, of_real_exp] +nonrec theorem exp_neg : exp (-x) = (exp x)⁻¹ := + ofReal_injective <| by simp [exp_neg] #align real.exp_neg Real.exp_neg theorem exp_sub : exp (x - y) = exp x / exp y := by @@ -1172,8 +1203,8 @@ theorem sin_zero : sin 0 = 0 := by simp [sin] theorem sin_neg : sin (-x) = -sin x := by simp [sin, exp_neg, (neg_div _ _).symm, add_mul] #align real.sin_neg Real.sin_neg -theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := by - rw [← of_real_inj] <;> simp [sin, sin_add] +nonrec theorem sin_add : sin (x + y) = sin x * cos y + cos x * sin y := + ofReal_injective <| by simp [sin_add] #align real.sin_add Real.sin_add @[simp] @@ -1186,11 +1217,11 @@ theorem cos_neg : cos (-x) = cos x := by simp [cos, exp_neg] @[simp] theorem cos_abs : cos (|x|) = cos x := by - cases le_total x 0 <;> simp only [*, abs_of_nonneg, abs_of_nonpos, cos_neg] + cases le_total x 0 <;> simp only [*, _root_.abs_of_nonneg, abs_of_nonpos, cos_neg] #align real.cos_abs Real.cos_abs -theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := by - rw [← of_real_inj] <;> simp [cos, cos_add] +nonrec theorem cos_add : cos (x + y) = cos x * cos y - sin x * sin y := + ofReal_injective <| by simp [cos_add] #align real.cos_add Real.cos_add theorem sin_sub : sin (x - y) = sin x * cos y - cos x * sin y := by @@ -1201,30 +1232,20 @@ theorem cos_sub : cos (x - y) = cos x * cos y + sin x * sin y := by simp [sub_eq_add_neg, cos_add, sin_neg, cos_neg] #align real.cos_sub Real.cos_sub -theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := by - rw [← of_real_inj] - simp only [sin, cos, of_real_sin_of_real_re, of_real_sub, of_real_add, of_real_div, of_real_mul, - of_real_one, of_real_bit0] - convert sin_sub_sin _ _ <;> norm_cast +nonrec theorem sin_sub_sin : sin x - sin y = 2 * sin ((x - y) / 2) * cos ((x + y) / 2) := + ofReal_injective <| by simp [sin_sub_sin] #align real.sin_sub_sin Real.sin_sub_sin -theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := by - rw [← of_real_inj] - simp only [cos, neg_mul, of_real_sin, of_real_sub, of_real_add, of_real_cos_of_real_re, - of_real_div, of_real_mul, of_real_one, of_real_neg, of_real_bit0] - convert cos_sub_cos _ _ - ring +nonrec theorem cos_sub_cos : cos x - cos y = -2 * sin ((x + y) / 2) * sin ((x - y) / 2) := + ofReal_injective <| by simp [cos_sub_cos] #align real.cos_sub_cos Real.cos_sub_cos -theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := by - rw [← of_real_inj] - simp only [cos, of_real_sub, of_real_add, of_real_cos_of_real_re, of_real_div, of_real_mul, - of_real_one, of_real_bit0] - convert cos_add_cos _ _ <;> norm_cast +nonrec theorem cos_add_cos : cos x + cos y = 2 * cos ((x + y) / 2) * cos ((x - y) / 2) := + ofReal_injective <| by simp [cos_add_cos] #align real.cos_add_cos Real.cos_add_cos -theorem tan_eq_sin_div_cos : tan x = sin x / cos x := by - rw [← of_real_inj, of_real_tan, tan_eq_sin_div_cos, of_real_div, of_real_sin, of_real_cos] +nonrec theorem tan_eq_sin_div_cos : tan x = sin x / cos x := + ofReal_injective <| by simp [tan_eq_sin_div_cos] #align real.tan_eq_sin_div_cos Real.tan_eq_sin_div_cos theorem tan_mul_cos {x : ℝ} (hx : cos x ≠ 0) : tan x * cos x = sin x := by @@ -1240,8 +1261,8 @@ theorem tan_neg : tan (-x) = -tan x := by simp [tan, neg_div] #align real.tan_neg Real.tan_neg @[simp] -theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := - of_real_inj.1 <| by simp +nonrec theorem sin_sq_add_cos_sq : sin x ^ 2 + cos x ^ 2 = 1 := + ofReal_injective (by simp [sin_sq_add_cos_sq]) #align real.sin_sq_add_cos_sq Real.sin_sq_add_cos_sq @[simp] @@ -1249,11 +1270,11 @@ theorem cos_sq_add_sin_sq : cos x ^ 2 + sin x ^ 2 = 1 := by rw [add_comm, sin_sq #align real.cos_sq_add_sin_sq Real.cos_sq_add_sin_sq theorem sin_sq_le_one : sin x ^ 2 ≤ 1 := by - rw [← sin_sq_add_cos_sq x] <;> exact le_add_of_nonneg_right (sq_nonneg _) + rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_right (sq_nonneg _) #align real.sin_sq_le_one Real.sin_sq_le_one theorem cos_sq_le_one : cos x ^ 2 ≤ 1 := by - rw [← sin_sq_add_cos_sq x] <;> exact le_add_of_nonneg_left (sq_nonneg _) + rw [← sin_sq_add_cos_sq x]; exact le_add_of_nonneg_left (sq_nonneg _) #align real.cos_sq_le_one Real.cos_sq_le_one theorem abs_sin_le_one : |sin x| ≤ 1 := @@ -1280,20 +1301,20 @@ theorem neg_one_le_cos : -1 ≤ cos x := (abs_le.1 (abs_cos_le_one _)).1 #align real.neg_one_le_cos Real.neg_one_le_cos -theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := by - rw [← of_real_inj] <;> simp [cos_two_mul] +nonrec theorem cos_two_mul : cos (2 * x) = 2 * cos x ^ 2 - 1 := + ofReal_injective <| by simp [cos_two_mul] #align real.cos_two_mul Real.cos_two_mul -theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := by - rw [← of_real_inj] <;> simp [cos_two_mul'] +nonrec theorem cos_two_mul' : cos (2 * x) = cos x ^ 2 - sin x ^ 2 := + ofReal_injective <| by simp [cos_two_mul'] #align real.cos_two_mul' Real.cos_two_mul' -theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := by - rw [← of_real_inj] <;> simp [sin_two_mul] +nonrec theorem sin_two_mul : sin (2 * x) = 2 * sin x * cos x := + ofReal_injective <| by simp [sin_two_mul] #align real.sin_two_mul Real.sin_two_mul -theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := - of_real_inj.1 <| by simpa using cos_sq x +nonrec theorem cos_sq : cos x ^ 2 = 1 / 2 + cos (2 * x) / 2 := + ofReal_injective <| by simp [cos_sq] #align real.cos_sq Real.cos_sq theorem cos_sq' : cos x ^ 2 = 1 - sin x ^ 2 := by rw [← sin_sq_add_cos_sq x, add_sub_cancel'] @@ -1313,7 +1334,7 @@ theorem abs_cos_eq_sqrt_one_sub_sin_sq (x : ℝ) : |cos x| = sqrt (1 - sin x ^ 2 theorem inv_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : (1 + tan x ^ 2)⁻¹ = cos x ^ 2 := have : Complex.cos x ≠ 0 := mt (congr_arg re) hx - of_real_inj.1 <| by simpa using Complex.inv_one_add_tan_sq this + ofReal_inj.1 <| by simpa using Complex.inv_one_add_tan_sq this #align real.inv_one_add_tan_sq Real.inv_one_add_tan_sq theorem tan_sq_div_one_add_tan_sq {x : ℝ} (hx : cos x ≠ 0) : @@ -1330,19 +1351,17 @@ theorem tan_div_sqrt_one_add_tan_sq {x : ℝ} (hx : 0 < cos x) : rw [← tan_mul_cos hx.ne', ← inv_sqrt_one_add_tan_sq hx, div_eq_mul_inv] #align real.tan_div_sqrt_one_add_tan_sq Real.tan_div_sqrt_one_add_tan_sq -theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by - rw [← of_real_inj] <;> simp [cos_three_mul] +nonrec theorem cos_three_mul : cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x := by + rw [← ofReal_inj]; simp [cos_three_mul] #align real.cos_three_mul Real.cos_three_mul -theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by - rw [← of_real_inj] <;> simp [sin_three_mul] +nonrec theorem sin_three_mul : sin (3 * x) = 3 * sin x - 4 * sin x ^ 3 := by + rw [← ofReal_inj]; simp [sin_three_mul] #align real.sin_three_mul Real.sin_three_mul /-- The definition of `sinh` in terms of `exp`. -/ -theorem sinh_eq (x : ℝ) : sinh x = (exp x - exp (-x)) / 2 := - eq_div_of_mul_eq two_ne_zero <| by - rw [sinh, exp, exp, Complex.of_real_neg, Complex.sinh, mul_two, ← Complex.add_re, ← mul_two, - div_mul_cancel _ (two_ne_zero' ℂ), Complex.sub_re] +nonrec theorem sinh_eq (x : ℝ) : sinh x = (exp x - exp (-x)) / 2 := + ofReal_injective <| by simp [Complex.sinh] #align real.sinh_eq Real.sinh_eq @[simp] @@ -1353,14 +1372,14 @@ theorem sinh_zero : sinh 0 = 0 := by simp [sinh] theorem sinh_neg : sinh (-x) = -sinh x := by simp [sinh, exp_neg, (neg_div _ _).symm, add_mul] #align real.sinh_neg Real.sinh_neg -theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by - rw [← of_real_inj] <;> simp [sinh_add] +nonrec theorem sinh_add : sinh (x + y) = sinh x * cosh y + cosh x * sinh y := by + rw [← ofReal_inj]; simp [sinh_add] #align real.sinh_add Real.sinh_add /-- The definition of `cosh` in terms of `exp`. -/ theorem cosh_eq (x : ℝ) : cosh x = (exp x + exp (-x)) / 2 := eq_div_of_mul_eq two_ne_zero <| by - rw [cosh, exp, exp, Complex.of_real_neg, Complex.cosh, mul_two, ← Complex.add_re, ← mul_two, + rw [cosh, exp, exp, Complex.ofReal_neg, Complex.cosh, mul_two, ← Complex.add_re, ← mul_two, div_mul_cancel _ (two_ne_zero' ℂ), Complex.add_re] #align real.cosh_eq Real.cosh_eq @@ -1370,16 +1389,16 @@ theorem cosh_zero : cosh 0 = 1 := by simp [cosh] @[simp] theorem cosh_neg : cosh (-x) = cosh x := - of_real_inj.1 <| by simp + ofReal_inj.1 <| by simp #align real.cosh_neg Real.cosh_neg @[simp] theorem cosh_abs : cosh (|x|) = cosh x := by - cases le_total x 0 <;> simp [*, abs_of_nonneg, abs_of_nonpos] + cases le_total x 0 <;> simp [*, _root_.abs_of_nonneg, abs_of_nonpos] #align real.cosh_abs Real.cosh_abs -theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by - rw [← of_real_inj] <;> simp [cosh_add] +nonrec theorem cosh_add : cosh (x + y) = cosh x * cosh y + sinh x * sinh y := by + rw [← ofReal_inj]; simp [cosh_add] #align real.cosh_add Real.cosh_add theorem sinh_sub : sinh (x - y) = sinh x * cosh y - cosh x * sinh y := by @@ -1390,8 +1409,8 @@ theorem cosh_sub : cosh (x - y) = cosh x * cosh y - sinh x * sinh y := by simp [sub_eq_add_neg, cosh_add, sinh_neg, cosh_neg] #align real.cosh_sub Real.cosh_sub -theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := - of_real_inj.1 <| by simp [tanh_eq_sinh_div_cosh] +nonrec theorem tanh_eq_sinh_div_cosh : tanh x = sinh x / cosh x := + ofReal_inj.1 <| by simp [tanh_eq_sinh_div_cosh] #align real.tanh_eq_sinh_div_cosh Real.tanh_eq_sinh_div_cosh @[simp] @@ -1403,7 +1422,7 @@ theorem tanh_neg : tanh (-x) = -tanh x := by simp [tanh, neg_div] #align real.tanh_neg Real.tanh_neg @[simp] -theorem cosh_add_sinh : cosh x + sinh x = exp x := by rw [← of_real_inj] <;> simp +theorem cosh_add_sinh : cosh x + sinh x = exp x := by rw [← ofReal_inj]; simp #align real.cosh_add_sinh Real.cosh_add_sinh @[simp] @@ -1422,7 +1441,7 @@ theorem exp_sub_sinh : exp x - sinh x = cosh x := @[simp] theorem cosh_sub_sinh : cosh x - sinh x = exp (-x) := by - rw [← of_real_inj] + rw [← ofReal_inj] simp #align real.cosh_sub_sinh Real.cosh_sub_sinh @@ -1431,56 +1450,56 @@ theorem sinh_sub_cosh : sinh x - cosh x = -exp (-x) := by rw [← neg_sub, cosh_ #align real.sinh_sub_cosh Real.sinh_sub_cosh @[simp] -theorem cosh_sq_sub_sinh_sq (x : ℝ) : cosh x ^ 2 - sinh x ^ 2 = 1 := by rw [← of_real_inj] <;> simp +theorem cosh_sq_sub_sinh_sq (x : ℝ) : cosh x ^ 2 - sinh x ^ 2 = 1 := by rw [← ofReal_inj]; simp #align real.cosh_sq_sub_sinh_sq Real.cosh_sq_sub_sinh_sq -theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by rw [← of_real_inj] <;> simp [cosh_sq] +nonrec theorem cosh_sq : cosh x ^ 2 = sinh x ^ 2 + 1 := by rw [← ofReal_inj]; simp [cosh_sq] #align real.cosh_sq Real.cosh_sq theorem cosh_sq' : cosh x ^ 2 = 1 + sinh x ^ 2 := (cosh_sq x).trans (add_comm _ _) #align real.cosh_sq' Real.cosh_sq' -theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by rw [← of_real_inj] <;> simp [sinh_sq] +nonrec theorem sinh_sq : sinh x ^ 2 = cosh x ^ 2 - 1 := by rw [← ofReal_inj]; simp [sinh_sq] #align real.sinh_sq Real.sinh_sq -theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by - rw [← of_real_inj] <;> simp [cosh_two_mul] +nonrec theorem cosh_two_mul : cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2 := by + rw [← ofReal_inj]; simp [cosh_two_mul] #align real.cosh_two_mul Real.cosh_two_mul -theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by - rw [← of_real_inj] <;> simp [sinh_two_mul] +nonrec theorem sinh_two_mul : sinh (2 * x) = 2 * sinh x * cosh x := by + rw [← ofReal_inj]; simp [sinh_two_mul] #align real.sinh_two_mul Real.sinh_two_mul -theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by - rw [← of_real_inj] <;> simp [cosh_three_mul] +nonrec theorem cosh_three_mul : cosh (3 * x) = 4 * cosh x ^ 3 - 3 * cosh x := by + rw [← ofReal_inj]; simp [cosh_three_mul] #align real.cosh_three_mul Real.cosh_three_mul -theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by - rw [← of_real_inj] <;> simp [sinh_three_mul] +nonrec theorem sinh_three_mul : sinh (3 * x) = 4 * sinh x ^ 3 + 3 * sinh x := by + rw [← ofReal_inj]; simp [sinh_three_mul] #align real.sinh_three_mul Real.sinh_three_mul open IsAbsoluteValue -/-- This is an intermediate result that is later replaced by `real.add_one_le_exp`; use that lemma +/-- This is an intermediate result that is later replaced by `Real.add_one_le_exp`; use that lemma instead. -/ theorem add_one_le_exp_of_nonneg {x : ℝ} (hx : 0 ≤ x) : x + 1 ≤ exp x := calc - x + 1 ≤ limUnder (⟨fun n : ℕ => ((exp' x) n).re, isCauSeq_re (exp' x)⟩ : CauSeq ℝ Abs.abs) := + x + 1 ≤ CauSeq.lim (⟨fun n : ℕ => ((exp' x) n).re, isCauSeq_re (exp' x)⟩ : CauSeq ℝ Abs.abs) := le_lim (CauSeq.le_of_exists ⟨2, fun j hj => - show x + (1 : ℝ) ≤ (∑ m in range j, (x ^ m / m ! : ℂ)).re - by - have h₁ : (((fun m : ℕ => (x ^ m / m ! : ℂ)) ∘ Nat.succ) 0).re = x := by simp - have h₂ : ((x : ℂ) ^ 0 / 0!).re = 1 := by simp - rw [← tsub_add_cancel_of_le hj, sum_range_succ', sum_range_succ', add_re, add_re, h₁, - h₂, add_assoc, ← coe_re_add_group_hom, re_add_group_hom.map_sum, - coe_re_add_group_hom] - refine' le_add_of_nonneg_of_le (sum_nonneg fun m hm => _) le_rfl - rw [← of_real_pow, ← of_real_nat_cast, ← of_real_div, of_real_re] + show x + (1 : ℝ) ≤ (∑ m in range j, ((x : ℂ) ^ m / m.factorial)).re by + have h₁ : (((fun m : ℕ => ((x : ℂ) ^ m / m.factorial)) ∘ Nat.succ) 0).re = x := + by simp [show Nat.succ 0 = 1 from rfl, Complex.ofReal_re] + have h₂ : ((x : ℂ) ^ 0 / (Nat.factorial 0)).re = 1 := by simp + erw [← tsub_add_cancel_of_le hj, sum_range_succ', sum_range_succ', add_re, add_re, h₁, + h₂, add_assoc, ← coe_reAddGroupHom, reAddGroupHom.map_sum, + coe_reAddGroupHom] + refine' le_add_of_nonneg_of_le (sum_nonneg fun m _ => _) le_rfl + rw [← ofReal_pow, ← ofReal_nat_cast, ← ofReal_div, ofReal_re] exact div_nonneg (pow_nonneg hx _) (Nat.cast_nonneg _)⟩) - _ = exp x := by rw [exp, Complex.exp, ← cau_seq_re, lim_re] + _ = exp x := by rw [exp, Complex.exp, ← cauSeqRe, lim_re] #align real.add_one_le_exp_of_nonneg Real.add_one_le_exp_of_nonneg @@ -1489,8 +1508,8 @@ theorem one_le_exp {x : ℝ} (hx : 0 ≤ x) : 1 ≤ exp x := by linarith [add_on theorem exp_pos (x : ℝ) : 0 < exp x := (le_total 0 x).elim (lt_of_lt_of_le zero_lt_one ∘ one_le_exp) fun h => by - rw [← neg_neg x, Real.exp_neg] <;> - exact inv_pos.2 (lt_of_lt_of_le zero_lt_one (one_le_exp (neg_nonneg.2 h))) + rw [← neg_neg x, Real.exp_neg] + exact inv_pos.2 (lt_of_lt_of_le zero_lt_one (one_le_exp (neg_nonneg.2 h))) #align real.exp_pos Real.exp_pos @[simp] @@ -1500,15 +1519,14 @@ theorem abs_exp (x : ℝ) : |exp x| = exp x := @[mono] theorem exp_strictMono : StrictMono exp := fun x y h => by - rw [← sub_add_cancel y x, Real.exp_add] <;> - exact - (lt_mul_iff_one_lt_left (exp_pos _)).2 - (lt_of_lt_of_le (by linarith) (add_one_le_exp_of_nonneg (by linarith))) + rw [← sub_add_cancel y x, Real.exp_add] + exact (lt_mul_iff_one_lt_left (exp_pos _)).2 + (lt_of_lt_of_le (by linarith) (add_one_le_exp_of_nonneg (by linarith))) #align real.exp_strict_mono Real.exp_strictMono @[mono] theorem exp_monotone : Monotone exp := - exp_strictMono.Monotone + exp_strictMono.monotone #align real.exp_monotone Real.exp_monotone @[simp] @@ -1522,7 +1540,7 @@ theorem exp_le_exp {x y : ℝ} : exp x ≤ exp y ↔ x ≤ y := #align real.exp_le_exp Real.exp_le_exp theorem exp_injective : Function.Injective exp := - exp_strictMono.Injective + exp_strictMono.injective #align real.exp_injective Real.exp_injective @[simp] @@ -1567,51 +1585,49 @@ end Real namespace Complex theorem sum_div_factorial_le {α : Type _} [LinearOrderedField α] (n j : ℕ) (hn : 0 < n) : - (∑ m in Filter (fun k => n ≤ k) (range j), (1 / m ! : α)) ≤ n.succ / (n ! * n) := + (∑ m in filter (fun k => n ≤ k) (range j), + (1 / m.factorial : α)) ≤ n.succ / (n.factorial * n) := calc - (∑ m in Filter (fun k => n ≤ k) (range j), (1 / m ! : α)) = - ∑ m in range (j - n), 1 / (m + n)! := + (∑ m in filter (fun k => n ≤ k) (range j), (1 / m.factorial : α)) = + ∑ m in range (j - n), (1 / ((m + n).factorial : α)) := sum_bij (fun m _ => m - n) (fun m hm => mem_range.2 <| - (tsub_lt_tsub_iff_right (by simp at hm <;> tauto)).2 (by simp at hm <;> tauto)) - (fun m hm => by rw [tsub_add_cancel_of_le] <;> simp at * <;> tauto) + (tsub_lt_tsub_iff_right (by simp at hm; tauto)).2 (by simp at hm; tauto)) + (fun m hm => by rw [tsub_add_cancel_of_le]; simp at *; tauto) (fun a₁ a₂ ha₁ ha₂ h => by rwa [tsub_eq_iff_eq_add_of_le, tsub_add_eq_add_tsub, eq_comm, tsub_eq_iff_eq_add_of_le, - add_left_inj, eq_comm] at h <;> - simp at * <;> - tauto) + add_left_inj, eq_comm] at h <;> + simp at * <;> aesop) fun b hb => ⟨b + n, mem_filter.2 ⟨mem_range.2 <| lt_tsub_iff_right.mp (mem_range.1 hb), Nat.le_add_left _ _⟩, - by rw [add_tsub_cancel_right]⟩ - _ ≤ ∑ m in range (j - n), (n ! * n.succ ^ m)⁻¹ := + by dsimp; rw [add_tsub_cancel_right]⟩ + _ ≤ ∑ m in range (j - n), ((n.factorial : α) * (n.succ : α) ^ m)⁻¹ := by refine' sum_le_sum fun m n => _ rw [one_div, inv_le_inv] · rw [← Nat.cast_pow, ← Nat.cast_mul, Nat.cast_le, add_comm] exact Nat.factorial_mul_pow_le_factorial · exact Nat.cast_pos.2 (Nat.factorial_pos _) - · - exact - mul_pos (Nat.cast_pos.2 (Nat.factorial_pos _)) + · exact mul_pos (Nat.cast_pos.2 (Nat.factorial_pos _)) (pow_pos (Nat.cast_pos.2 (Nat.succ_pos _)) _) - _ = n !⁻¹ * ∑ m in range (j - n), n.succ⁻¹ ^ m := by + _ = (n.factorial : α)⁻¹ * ∑ m in range (j - n), (n.succ : α)⁻¹ ^ m := by simp [mul_inv, mul_sum.symm, sum_mul.symm, -Nat.factorial_succ, mul_comm, inv_pow] - _ = (n.succ - n.succ * n.succ⁻¹ ^ (j - n)) / (n ! * n) := - by - have h₁ : (n.succ : α) ≠ 1 := - @Nat.cast_one α _ ▸ mt Nat.cast_inj.1 (mt Nat.succ.inj (pos_iff_ne_zero.1 hn)) - have h₂ : (n.succ : α) ≠ 0 := Nat.cast_ne_zero.2 (Nat.succ_ne_zero _) - have h₃ : (n ! * n : α) ≠ 0 := - mul_ne_zero (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 (Nat.factorial_pos _))) - (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 hn)) - have h₄ : (n.succ - 1 : α) = n := by simp - rw [geom_sum_inv h₁ h₂, eq_div_iff_mul_eq h₃, mul_comm _ (n ! * n : α), ← - mul_assoc (n !⁻¹ : α), ← mul_inv_rev, h₄, ← mul_assoc (n ! * n : α), mul_comm (n : α) n !, - mul_inv_cancel h₃] <;> + _ = ((n.succ : α) - n.succ * (n.succ : α)⁻¹ ^ (j - n)) / (n.factorial * n) := + by + have h₁ : (n.succ : α) ≠ 1 := + @Nat.cast_one α _ ▸ mt Nat.cast_inj.1 (mt Nat.succ.inj (pos_iff_ne_zero.1 hn)) + have h₂ : (n.succ : α) ≠ 0 := Nat.cast_ne_zero.2 (Nat.succ_ne_zero _) + have h₃ : (n.factorial * n : α) ≠ 0 := + mul_ne_zero (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 (Nat.factorial_pos _))) + (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 hn)) + have h₄ : (n.succ - 1 : α) = n := by simp + rw [geom_sum_inv h₁ h₂, eq_div_iff_mul_eq h₃, mul_comm _ (n.factorial * n : α), + ← mul_assoc (n.factorial⁻¹ : α), ← mul_inv_rev, h₄, ← mul_assoc (n.factorial * n : α), + mul_comm (n : α) n.factorial, mul_inv_cancel h₃] simp [mul_add, add_mul, mul_assoc, mul_comm] - _ ≤ n.succ / (n ! * n) := + _ ≤ n.succ / (n.factorial * n) := by refine' Iff.mpr (div_le_div_right (mul_pos _ _)) _ exact Nat.cast_pos.2 (Nat.factorial_pos _) @@ -1619,28 +1635,28 @@ theorem sum_div_factorial_le {α : Type _} [LinearOrderedField α] (n j : ℕ) ( exact sub_le_self _ (mul_nonneg (Nat.cast_nonneg _) (pow_nonneg (inv_nonneg.2 (Nat.cast_nonneg _)) _)) - #align complex.sum_div_factorial_le Complex.sum_div_factorial_le theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : - abs (exp x - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n * (n.succ * (n ! * n)⁻¹) := by + abs (exp x - ∑ m in range n, x ^ m / m.factorial) ≤ + abs x ^ n * (n.succ * (n.factorial * n)⁻¹) := by rw [← lim_const (∑ m in range n, _), exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] refine' lim_le (CauSeq.le_of_exists ⟨n, fun j hj => _⟩) simp_rw [← sub_eq_add_neg] show - abs ((∑ m in range j, x ^ m / m !) - ∑ m in range n, x ^ m / m !) ≤ - abs x ^ n * (n.succ * (n ! * n)⁻¹) + abs ((∑ m in range j, x ^ m / m.factorial) - ∑ m in range n, x ^ m / m.factorial) ≤ + abs x ^ n * (n.succ * (n.factorial * n)⁻¹) rw [sum_range_sub_sum_range hj] calc - abs (∑ m in (range j).filterₓ fun k => n ≤ k, (x ^ m / m ! : ℂ)) = - abs (∑ m in (range j).filterₓ fun k => n ≤ k, (x ^ n * (x ^ (m - n) / m !) : ℂ)) := + abs (∑ m in (range j).filterₓ fun k => n ≤ k, (x ^ m / m.factorial : ℂ)) = + abs (∑ m in (range j).filterₓ fun k => n ≤ k, (x ^ n * (x ^ (m - n) / m.factorial) : ℂ)) := by refine' congr_arg abs (sum_congr rfl fun m hm => _) rw [mem_filter, mem_range] at hm rw [← mul_div_assoc, ← pow_add, add_tsub_cancel_of_le hm.2] - _ ≤ ∑ m in Filter (fun k => n ≤ k) (range j), abs (x ^ n * (_ / m !)) := + _ ≤ ∑ m in filter (fun k => n ≤ k) (range j), abs (x ^ n * (_ / m.factorial)) := (abv_sum_le_sum_abv _ _) - _ ≤ ∑ m in Filter (fun k => n ≤ k) (range j), abs x ^ n * (1 / m !) := + _ ≤ ∑ m in filter (fun k => n ≤ k) (range j), abs x ^ n * (1 / m.factorial) := by refine' sum_le_sum fun m hm => _ rw [map_mul, map_pow, map_div₀, abs_cast_nat] @@ -1649,19 +1665,19 @@ theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : · rw [abv_pow abs] exact pow_le_one _ (abs.nonneg _) hx · exact pow_nonneg (abs.nonneg _) _ - _ = abs x ^ n * ∑ m in (range j).filterₓ fun k => n ≤ k, (1 / m ! : ℝ) := by + _ = abs x ^ n * ∑ m in (range j).filterₓ fun k => n ≤ k, (1 / m.factorial : ℝ) := by simp [abs_mul, abv_pow abs, abs_div, mul_sum.symm] - _ ≤ abs x ^ n * (n.succ * (n ! * n)⁻¹) := + _ ≤ abs x ^ n * (n.succ * (n.factorial * n)⁻¹) := mul_le_mul_of_nonneg_left (sum_div_factorial_le _ _ hn) (pow_nonneg (abs.nonneg _) _) #align complex.exp_bound Complex.exp_bound theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : - abs (exp x - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n / n ! * 2 := by + abs (exp x - ∑ m in range n, x ^ m / m.factorial) ≤ abs x ^ n / n.factorial * 2 := by rw [← lim_const (∑ m in range n, _), exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] refine' lim_le (CauSeq.le_of_exists ⟨n, fun j hj => _⟩) simp_rw [← sub_eq_add_neg] - show abs ((∑ m in range j, x ^ m / m !) - ∑ m in range n, x ^ m / m !) ≤ abs x ^ n / n ! * 2 + show abs ((∑ m in range j, x ^ m / m.factorial) - ∑ m in range n, x ^ m / m.factorial) ≤ abs x ^ n / n.factorial * 2 let k := j - n have hj : j = n + k := (add_tsub_cancel_of_le hj).symm rw [hj, sum_range_add_sub_sum_range] @@ -1671,9 +1687,9 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : abv_sum_le_sum_abv _ _ _ ≤ ∑ I : ℕ in range k, abs x ^ (n + I) / (n + I)! := by simp only [Complex.abs_cast_nat, map_div₀, abv_pow abs] - _ ≤ ∑ I : ℕ in range k, abs x ^ (n + I) / (n ! * n.succ ^ I) := _ - _ = ∑ I : ℕ in range k, abs x ^ n / n ! * (abs x ^ I / n.succ ^ I) := _ - _ ≤ abs x ^ n / ↑n ! * 2 := _ + _ ≤ ∑ I : ℕ in range k, abs x ^ (n + I) / (n.factorial * n.succ ^ I) := _ + _ = ∑ I : ℕ in range k, abs x ^ n / n.factorial * (abs x ^ I / n.succ ^ I) := _ + _ ≤ abs x ^ n / ↑n.factorial * 2 := _ · refine' sum_le_sum fun m hm => div_le_div (pow_nonneg (abs.nonneg x) (n + m)) le_rfl _ _ · exact_mod_cast mul_pos n.factorial_pos (pow_pos n.succ_pos _) @@ -1692,12 +1708,12 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : (pow_nonneg (add_nonneg n.cast_nonneg zero_le_one) k) · linarith · linarith - · exact div_nonneg (pow_nonneg (abs.nonneg x) n) (Nat.cast_nonneg n !) + · exact div_nonneg (pow_nonneg (abs.nonneg x) n) (Nat.cast_nonneg n.factorial) #align complex.exp_bound' Complex.exp_bound' theorem abs_exp_sub_one_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1) ≤ 2 * abs x := calc - abs (exp x - 1) = abs (exp x - ∑ m in range 1, x ^ m / m !) := by simp [sum_range_succ] + abs (exp x - 1) = abs (exp x - ∑ m in range 1, x ^ m / m.factorial) := by simp [sum_range_succ] _ ≤ abs x ^ 1 * (Nat.succ 1 * (1! * (1 : ℕ))⁻¹) := (exp_bound hx (by decide)) _ = 2 * abs x := by simp [two_mul, mul_two, mul_add, mul_comm] @@ -1705,7 +1721,7 @@ theorem abs_exp_sub_one_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1) ≤ 2 theorem abs_exp_sub_one_sub_id_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1 - x) ≤ abs x ^ 2 := calc - abs (exp x - 1 - x) = abs (exp x - ∑ m in range 2, x ^ m / m !) := by + abs (exp x - 1 - x) = abs (exp x - ∑ m in range 2, x ^ m / m.factorial) := by simp [sub_eq_add_neg, sum_range_succ_comm, add_assoc] _ ≤ abs x ^ 2 * (Nat.succ 2 * (2! * (2 : ℕ))⁻¹) := (exp_bound hx (by decide)) _ ≤ abs x ^ 2 * 1 := (mul_le_mul_of_nonneg_left (by norm_num) (sq_nonneg (abs x))) @@ -1720,13 +1736,13 @@ namespace Real open Complex Finset theorem exp_bound {x : ℝ} (hx : |x| ≤ 1) {n : ℕ} (hn : 0 < n) : - |exp x - ∑ m in range n, x ^ m / m !| ≤ |x| ^ n * (n.succ / (n ! * n)) := by + |exp x - ∑ m in range n, x ^ m / m.factorial| ≤ |x| ^ n * (n.succ / (n.factorial * n)) := by have hxc : Complex.abs x ≤ 1 := by exact_mod_cast hx convert exp_bound hxc hn <;> norm_cast #align real.exp_bound Real.exp_bound theorem exp_bound' {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) {n : ℕ} (hn : 0 < n) : - Real.exp x ≤ (∑ m in Finset.range n, x ^ m / m !) + x ^ n * (n + 1) / (n ! * n) := by + Real.exp x ≤ (∑ m in Finset.range n, x ^ m / m.factorial) + x ^ n * (n + 1) / (n.factorial * n) := by have h3 : |x| = x := by simpa have h4 : |x| ≤ 1 := by rwa [h3] have h' := Real.exp_bound h4 hn @@ -1752,7 +1768,7 @@ For fixed `n` this is just a linear map wrt `r`, and each map is a simple linear of the previous (see `exp_near_succ`), with `exp_near n x r ⟶ exp x` as `n ⟶ ∞`, for any `r`. -/ def expNear (n : ℕ) (x r : ℝ) : ℝ := - (∑ m in range n, x ^ m / m !) + x ^ n / n ! * r + (∑ m in range n, x ^ m / m.factorial) + x ^ n / n.factorial * r #align real.exp_near Real.expNear @[simp] @@ -1766,12 +1782,12 @@ theorem expNear_succ (n x r) : expNear (n + 1) x r = expNear n x (1 + x / (n + 1 ac_rfl #align real.exp_near_succ Real.expNear_succ -theorem expNear_sub (n x r₁ r₂) : expNear n x r₁ - expNear n x r₂ = x ^ n / n ! * (r₁ - r₂) := by +theorem expNear_sub (n x r₁ r₂) : expNear n x r₁ - expNear n x r₂ = x ^ n / n.factorial * (r₁ - r₂) := by simp [exp_near, mul_sub] #align real.exp_near_sub Real.expNear_sub theorem exp_approx_end (n m : ℕ) (x : ℝ) (e₁ : n + 1 = m) (h : |x| ≤ 1) : - |exp x - expNear m x 0| ≤ |x| ^ m / m ! * ((m + 1) / m) := by + |exp x - expNear m x 0| ≤ |x| ^ m / m.factorial * ((m + 1) / m) := by simp [exp_near] convert exp_bound h _ using 1 field_simp [mul_comm] @@ -1780,8 +1796,8 @@ theorem exp_approx_end (n m : ℕ) (x : ℝ) (e₁ : n + 1 = m) (h : |x| ≤ 1) theorem exp_approx_succ {n} {x a₁ b₁ : ℝ} (m : ℕ) (e₁ : n + 1 = m) (a₂ b₂ : ℝ) (e : |1 + x / m * a₂ - a₁| ≤ b₁ - |x| / m * b₂) - (h : |exp x - expNear m x a₂| ≤ |x| ^ m / m ! * b₂) : - |exp x - expNear n x a₁| ≤ |x| ^ n / n ! * b₁ := by + (h : |exp x - expNear m x a₂| ≤ |x| ^ m / m.factorial * b₂) : + |exp x - expNear n x a₁| ≤ |x| ^ n / n.factorial * b₁ := by refine' (abs_sub_le _ _ _).trans ((add_le_add_right h _).trans _) subst e₁; rw [exp_near_succ, exp_near_sub, abs_mul] convert mul_le_mul_of_nonneg_left (le_sub_iff_add_le'.1 e) _ @@ -1792,13 +1808,13 @@ theorem exp_approx_succ {n} {x a₁ b₁ : ℝ} (m : ℕ) (e₁ : n + 1 = m) (a theorem exp_approx_end' {n} {x a b : ℝ} (m : ℕ) (e₁ : n + 1 = m) (rm : ℝ) (er : ↑m = rm) (h : |x| ≤ 1) (e : |1 - a| ≤ b - |x| / rm * ((rm + 1) / rm)) : - |exp x - expNear n x a| ≤ |x| ^ n / n ! * b := by + |exp x - expNear n x a| ≤ |x| ^ n / n.factorial * b := by subst er <;> exact exp_approx_succ _ e₁ _ _ (by simpa using e) (exp_approx_end _ _ _ e₁ h) #align real.exp_approx_end' Real.exp_approx_end' theorem exp_1_approx_succ_eq {n} {a₁ b₁ : ℝ} {m : ℕ} (en : n + 1 = m) {rm : ℝ} (er : ↑m = rm) - (h : |exp 1 - expNear m 1 ((a₁ - 1) * rm)| ≤ |1| ^ m / m ! * (b₁ * rm)) : - |exp 1 - expNear n 1 a₁| ≤ |1| ^ n / n ! * b₁ := by + (h : |exp 1 - expNear m 1 ((a₁ - 1) * rm)| ≤ |1| ^ m / m.factorial * (b₁ * rm)) : + |exp 1 - expNear n 1 a₁| ≤ |1| ^ n / n.factorial * b₁ := by subst er refine' exp_approx_succ _ en _ _ _ h field_simp [show (m : ℝ) ≠ 0 by norm_cast <;> linarith] @@ -1816,8 +1832,8 @@ theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| simp [Complex.cos, sub_div, add_div, neg_div, div_self (two_ne_zero' ℂ)] _ = abs - (((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) + - (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !)) / + (((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) + + (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial)) / 2) := (congr_arg abs (congr_arg (fun x : ℂ => x / 2) @@ -1826,12 +1842,12 @@ theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| simp [pow_succ] apply Complex.ext <;> simp [div_eq_mul_inv, norm_sq] <;> ring))) _ ≤ - abs ((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) / 2) + - abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) / 2) := + abs ((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2) + + abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) / 2) := by rw [add_div] <;> exact complex.abs.add_le _ _ _ = - abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) / 2 + - abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) / 2 := + abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2 + + abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) / 2 := by simp [map_div₀] _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + @@ -1852,8 +1868,8 @@ theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| show (3 : ℂ) * 2 = 6 by norm_num] _ = abs - (((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) - - (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !)) * + (((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) - + (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial)) * I / 2) := (congr_arg abs @@ -1863,12 +1879,12 @@ theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| simp [pow_succ] apply Complex.ext <;> simp [div_eq_mul_inv, norm_sq] <;> ring))) _ ≤ - abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) * I / 2) + - abs (-((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) * I) / 2) := + abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) * I / 2) + + abs (-((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) * I) / 2) := by rw [sub_mul, sub_eq_add_neg, add_div] <;> exact complex.abs.add_le _ _ _ = - abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m !) / 2 + - abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m !) / 2 := + abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2 + + abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) / 2 := by simp [add_comm, map_div₀] _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + @@ -2060,7 +2076,7 @@ end Tactic namespace Complex @[simp] -theorem abs_cos_add_sin_mul_i (x : ℝ) : abs (cos x + sin x * I) = 1 := by +theorem abs_cos_add_sin_mul_I (x : ℝ) : abs (cos x + sin x * I) = 1 := by have := Real.sin_sq_add_cos_sq x simp_all [add_comm, abs, norm_sq, sq, sin_of_real_re, cos_of_real_re, mul_re] #align complex.abs_cos_add_sin_mul_I Complex.abs_cos_add_sin_mul_i @@ -2071,7 +2087,7 @@ theorem abs_exp_of_real (x : ℝ) : abs (exp x) = Real.exp x := by #align complex.abs_exp_of_real Complex.abs_exp_of_real @[simp] -theorem abs_exp_of_real_mul_i (x : ℝ) : abs (exp (x * I)) = 1 := by +theorem abs_exp_of_real_mul_I (x : ℝ) : abs (exp (x * I)) = 1 := by rw [exp_mul_I, abs_cos_add_sin_mul_I] #align complex.abs_exp_of_real_mul_I Complex.abs_exp_of_real_mul_i From 17f691f2375b170fdfb5ce57ef98765d4d1b67a7 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Fri, 10 Mar 2023 16:26:15 +0000 Subject: [PATCH 10/23] minor improvements --- Mathlib/Data/Complex/Exponential.lean | 49 ++++++++++++++------------- 1 file changed, 26 insertions(+), 23 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 82b7da8bb0109..0152f520cddcd 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1639,42 +1639,44 @@ theorem sum_div_factorial_le {α : Type _} [LinearOrderedField α] (n j : ℕ) ( theorem exp_bound {x : ℂ} (hx : abs x ≤ 1) {n : ℕ} (hn : 0 < n) : abs (exp x - ∑ m in range n, x ^ m / m.factorial) ≤ - abs x ^ n * (n.succ * (n.factorial * n)⁻¹) := by - rw [← lim_const (∑ m in range n, _), exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] + abs x ^ n * ((n.succ : ℝ) * (n.factorial * n : ℝ)⁻¹) := by + rw [← lim_const (abv := Complex.abs) (∑ m in range n, _), exp, sub_eq_add_neg, + ← lim_neg, lim_add, ← lim_abs] refine' lim_le (CauSeq.le_of_exists ⟨n, fun j hj => _⟩) simp_rw [← sub_eq_add_neg] show abs ((∑ m in range j, x ^ m / m.factorial) - ∑ m in range n, x ^ m / m.factorial) ≤ - abs x ^ n * (n.succ * (n.factorial * n)⁻¹) + abs x ^ n * ((n.succ : ℝ) * (n.factorial * n : ℝ)⁻¹) rw [sum_range_sub_sum_range hj] calc - abs (∑ m in (range j).filterₓ fun k => n ≤ k, (x ^ m / m.factorial : ℂ)) = - abs (∑ m in (range j).filterₓ fun k => n ≤ k, (x ^ n * (x ^ (m - n) / m.factorial) : ℂ)) := + abs (∑ m in (range j).filter fun k => n ≤ k, (x ^ m / m.factorial : ℂ)) = + abs (∑ m in (range j).filter fun k => n ≤ k, (x ^ n * (x ^ (m - n) / m.factorial) : ℂ)) := by refine' congr_arg abs (sum_congr rfl fun m hm => _) rw [mem_filter, mem_range] at hm rw [← mul_div_assoc, ← pow_add, add_tsub_cancel_of_le hm.2] - _ ≤ ∑ m in filter (fun k => n ≤ k) (range j), abs (x ^ n * (_ / m.factorial)) := - (abv_sum_le_sum_abv _ _) + _ ≤ ∑ m in filter (fun k => n ≤ k) (range j), abs (x ^ n * (x ^ (m - n) / m.factorial)) := + (abv_sum_le_sum_abv (abv := Complex.abs) _ _) _ ≤ ∑ m in filter (fun k => n ≤ k) (range j), abs x ^ n * (1 / m.factorial) := by - refine' sum_le_sum fun m hm => _ + refine' sum_le_sum fun m _ => _ rw [map_mul, map_pow, map_div₀, abs_cast_nat] refine' mul_le_mul_of_nonneg_left ((div_le_div_right _).2 _) _ · exact Nat.cast_pos.2 (Nat.factorial_pos _) · rw [abv_pow abs] exact pow_le_one _ (abs.nonneg _) hx · exact pow_nonneg (abs.nonneg _) _ - _ = abs x ^ n * ∑ m in (range j).filterₓ fun k => n ≤ k, (1 / m.factorial : ℝ) := by + _ = abs x ^ n * ∑ m in (range j).filter fun k => n ≤ k, (1 / m.factorial : ℝ) := by simp [abs_mul, abv_pow abs, abs_div, mul_sum.symm] - _ ≤ abs x ^ n * (n.succ * (n.factorial * n)⁻¹) := + _ ≤ abs x ^ n * (n.succ * (n.factorial * n : ℝ)⁻¹) := mul_le_mul_of_nonneg_left (sum_div_factorial_le _ _ hn) (pow_nonneg (abs.nonneg _) _) #align complex.exp_bound Complex.exp_bound theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : abs (exp x - ∑ m in range n, x ^ m / m.factorial) ≤ abs x ^ n / n.factorial * 2 := by - rw [← lim_const (∑ m in range n, _), exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] + rw [← lim_const (abv := Complex.abs) (∑ m in range n, _), + exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] refine' lim_le (CauSeq.le_of_exists ⟨n, fun j hj => _⟩) simp_rw [← sub_eq_add_neg] show abs ((∑ m in range j, x ^ m / m.factorial) - ∑ m in range n, x ^ m / m.factorial) ≤ abs x ^ n / n.factorial * 2 @@ -1682,15 +1684,14 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : have hj : j = n + k := (add_tsub_cancel_of_le hj).symm rw [hj, sum_range_add_sub_sum_range] calc - abs (∑ I : ℕ in range k, x ^ (n + I) / ((n + I)! : ℂ)) ≤ - ∑ I : ℕ in range k, abs (x ^ (n + I) / ((n + I)! : ℂ)) := + abs (∑ i : ℕ in range k, x ^ (n + i) / ((n + i).factorial : ℂ)) ≤ + ∑ i : ℕ in range k, abs (x ^ (n + i) / ((n + i).factorial : ℂ)) := abv_sum_le_sum_abv _ _ - _ ≤ ∑ I : ℕ in range k, abs x ^ (n + I) / (n + I)! := by - simp only [Complex.abs_cast_nat, map_div₀, abv_pow abs] - _ ≤ ∑ I : ℕ in range k, abs x ^ (n + I) / (n.factorial * n.succ ^ I) := _ - _ = ∑ I : ℕ in range k, abs x ^ n / n.factorial * (abs x ^ I / n.succ ^ I) := _ - _ ≤ abs x ^ n / ↑n.factorial * 2 := _ - + _ ≤ ∑ i : ℕ in range k, abs x ^ (n + i) / (n + i).factorial := by + simp [Complex.abs_cast_nat, map_div₀, abv_pow abs] + _ ≤ ∑ i : ℕ in range k, abs x ^ (n + i) / ((n.factorial : ℝ) * (n.succ : ℝ) ^ i) := ?_ + _ = ∑ i : ℕ in range k, abs x ^ n / n.factorial * (abs x ^ i / (n.succ : ℝ) ^ i) := ?_ + _ ≤ abs x ^ n / ↑n.factorial * 2 := ?_ · refine' sum_le_sum fun m hm => div_le_div (pow_nonneg (abs.nonneg x) (n + m)) le_rfl _ _ · exact_mod_cast mul_pos n.factorial_pos (pow_pos n.succ_pos _) · exact_mod_cast Nat.factorial_mul_pow_le_factorial @@ -1714,8 +1715,9 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : theorem abs_exp_sub_one_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1) ≤ 2 * abs x := calc abs (exp x - 1) = abs (exp x - ∑ m in range 1, x ^ m / m.factorial) := by simp [sum_range_succ] - _ ≤ abs x ^ 1 * (Nat.succ 1 * (1! * (1 : ℕ))⁻¹) := (exp_bound hx (by decide)) - _ = 2 * abs x := by simp [two_mul, mul_two, mul_add, mul_comm] + _ ≤ abs x ^ 1 * ((Nat.succ 1 : ℝ) * ((Nat.factorial 1) * (1 : ℕ) : ℝ)⁻¹) := + (exp_bound hx (by decide)) + _ = 2 * abs x := by simp [two_mul, mul_two, mul_add, mul_comm, add_mul] #align complex.abs_exp_sub_one_le Complex.abs_exp_sub_one_le @@ -1723,7 +1725,8 @@ theorem abs_exp_sub_one_sub_id_le {x : ℂ} (hx : abs x ≤ 1) : abs (exp x - 1 calc abs (exp x - 1 - x) = abs (exp x - ∑ m in range 2, x ^ m / m.factorial) := by simp [sub_eq_add_neg, sum_range_succ_comm, add_assoc] - _ ≤ abs x ^ 2 * (Nat.succ 2 * (2! * (2 : ℕ))⁻¹) := (exp_bound hx (by decide)) + _ ≤ abs x ^ 2 * ((Nat.succ 2 : ℝ) * (Nat.factorial 2 * (2 : ℕ) : ℝ)⁻¹) := + (exp_bound hx (by decide)) _ ≤ abs x ^ 2 * 1 := (mul_le_mul_of_nonneg_left (by norm_num) (sq_nonneg (abs x))) _ = abs x ^ 2 := by rw [mul_one] @@ -1735,7 +1738,7 @@ namespace Real open Complex Finset -theorem exp_bound {x : ℝ} (hx : |x| ≤ 1) {n : ℕ} (hn : 0 < n) : +nonrec theorem exp_bound {x : ℝ} (hx : |x| ≤ 1) {n : ℕ} (hn : 0 < n) : |exp x - ∑ m in range n, x ^ m / m.factorial| ≤ |x| ^ n * (n.succ / (n.factorial * n)) := by have hxc : Complex.abs x ≤ 1 := by exact_mod_cast hx convert exp_bound hxc hn <;> norm_cast From 0c4d7533a14e09281289a44b46819b45939639a1 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Sat, 11 Mar 2023 12:44:28 +0000 Subject: [PATCH 11/23] prove ofReal_ofNat --- Mathlib/Data/Complex/Basic.lean | 15 ++++----------- 1 file changed, 4 insertions(+), 11 deletions(-) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 692eb27e1ec85..13781de832b07 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -478,17 +478,10 @@ theorem I_pow_bit1 (n : ℕ) : I ^ bit1 n = (-1) ^ n * I := by rw [pow_bit1', Co set_option linter.uppercaseLean3 false in #align complex.I_pow_bit1 Complex.I_pow_bit1 ---Porting note: new ugly theorem -@[simp] -theorem ofReal_two : ((2 : ℝ) : ℂ) = 2 := rfl - ---Porting note: new ugly theorem -@[simp] -theorem ofReal_three : ((3 : ℝ) : ℂ) = 3 := rfl - ---Porting note: new ugly theorem -@[simp] -theorem ofReal_four : ((4 : ℝ) : ℂ) = 4 := rfl +--Porting note: new theorem +@[simp, norm_cast] +theorem ofReal_ofNat (n : ℕ) [n.AtLeastTwo] : ((OfNat.ofNat n : ℝ) : ℂ) = OfNat.ofNat n := + rfl end /-! ### Complex conjugation -/ From e393967617dcc2a1c9903e94531c0679a37970a3 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Sat, 11 Mar 2023 13:27:17 +0000 Subject: [PATCH 12/23] slight tidy --- Mathlib/Data/Complex/Exponential.lean | 27 ++++++++++++++------------- 1 file changed, 14 insertions(+), 13 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 0152f520cddcd..ed871a7fb3044 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1615,19 +1615,20 @@ theorem sum_div_factorial_le {α : Type _} [LinearOrderedField α] (n j : ℕ) ( _ = (n.factorial : α)⁻¹ * ∑ m in range (j - n), (n.succ : α)⁻¹ ^ m := by simp [mul_inv, mul_sum.symm, sum_mul.symm, -Nat.factorial_succ, mul_comm, inv_pow] _ = ((n.succ : α) - n.succ * (n.succ : α)⁻¹ ^ (j - n)) / (n.factorial * n) := - by - have h₁ : (n.succ : α) ≠ 1 := - @Nat.cast_one α _ ▸ mt Nat.cast_inj.1 (mt Nat.succ.inj (pos_iff_ne_zero.1 hn)) - have h₂ : (n.succ : α) ≠ 0 := Nat.cast_ne_zero.2 (Nat.succ_ne_zero _) - have h₃ : (n.factorial * n : α) ≠ 0 := - mul_ne_zero (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 (Nat.factorial_pos _))) - (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 hn)) - have h₄ : (n.succ - 1 : α) = n := by simp - rw [geom_sum_inv h₁ h₂, eq_div_iff_mul_eq h₃, mul_comm _ (n.factorial * n : α), - ← mul_assoc (n.factorial⁻¹ : α), ← mul_inv_rev, h₄, ← mul_assoc (n.factorial * n : α), - mul_comm (n : α) n.factorial, mul_inv_cancel h₃] - simp [mul_add, add_mul, mul_assoc, mul_comm] - _ ≤ n.succ / (n.factorial * n) := + by + have h₁ : (n.succ : α) ≠ 1 := + @Nat.cast_one α _ ▸ mt Nat.cast_inj.1 (mt Nat.succ.inj (pos_iff_ne_zero.1 hn)) + have h₂ : (n.succ : α) ≠ 0 := Nat.cast_ne_zero.2 (Nat.succ_ne_zero _) + have h₃ : (n.factorial * n : α) ≠ 0 := + mul_ne_zero (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 (Nat.factorial_pos _))) + (Nat.cast_ne_zero.2 (pos_iff_ne_zero.1 hn)) + have h₄ : (n.succ - 1 : α) = n := by simp + rw [geom_sum_inv h₁ h₂, eq_div_iff_mul_eq h₃, mul_comm _ (n.factorial * n : α), + ← mul_assoc (n.factorial⁻¹ : α), ← mul_inv_rev, h₄, ← mul_assoc (n.factorial * n : α), + mul_comm (n : α) n.factorial, mul_inv_cancel h₃] + simp [mul_add, add_mul, mul_assoc, mul_comm] + simp + _ ≤ n.succ / (n.factorial * n : α) := by refine' Iff.mpr (div_le_div_right (mul_pos _ _)) _ exact Nat.cast_pos.2 (Nat.factorial_pos _) From 03253992e65a1864e46517763f5e96fe5047fcf9 Mon Sep 17 00:00:00 2001 From: Lukas Miaskiwskyi Date: Sun, 12 Mar 2023 18:03:47 +0100 Subject: [PATCH 13/23] mild fixes --- Mathlib/Data/Complex/Exponential.lean | 69 ++++++++++++++------------- 1 file changed, 35 insertions(+), 34 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index ed871a7fb3044..4968f5e5af3ac 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -21,7 +21,7 @@ hyperbolic sine, hyperbolic cosine, and hyperbolic tangent functions. -/ --- mathport name: exprabs' +@[inherit_doc] local notation "abs'" => Abs.abs open IsAbsoluteValue @@ -1603,9 +1603,8 @@ theorem sum_div_factorial_le {α : Type _} [LinearOrderedField α] (n j : ℕ) ( ⟨b + n, mem_filter.2 ⟨mem_range.2 <| lt_tsub_iff_right.mp (mem_range.1 hb), Nat.le_add_left _ _⟩, by dsimp; rw [add_tsub_cancel_right]⟩ - _ ≤ ∑ m in range (j - n), ((n.factorial : α) * (n.succ : α) ^ m)⁻¹ := - by - refine' sum_le_sum fun m n => _ + _ ≤ ∑ m in range (j - n), ((n.factorial : α) * (n.succ : α) ^ m)⁻¹ := by + refine' sum_le_sum fun m _ => _ rw [one_div, inv_le_inv] · rw [← Nat.cast_pow, ← Nat.cast_mul, Nat.cast_le, add_comm] exact Nat.factorial_mul_pow_le_factorial @@ -1625,9 +1624,7 @@ theorem sum_div_factorial_le {α : Type _} [LinearOrderedField α] (n j : ℕ) ( have h₄ : (n.succ - 1 : α) = n := by simp rw [geom_sum_inv h₁ h₂, eq_div_iff_mul_eq h₃, mul_comm _ (n.factorial * n : α), ← mul_assoc (n.factorial⁻¹ : α), ← mul_inv_rev, h₄, ← mul_assoc (n.factorial * n : α), - mul_comm (n : α) n.factorial, mul_inv_cancel h₃] - simp [mul_add, add_mul, mul_assoc, mul_comm] - simp + mul_comm (n : α) n.factorial, mul_inv_cancel h₃, one_mul, mul_comm] _ ≤ n.succ / (n.factorial * n : α) := by refine' Iff.mpr (div_le_div_right (mul_pos _ _)) _ @@ -1762,37 +1759,37 @@ theorem abs_exp_sub_one_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1| ≤ 2 * |x| #align real.abs_exp_sub_one_le Real.abs_exp_sub_one_le theorem abs_exp_sub_one_sub_id_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1 - x| ≤ x ^ 2 := by - rw [← sq_abs] + rw [← _root_.sq_abs] have : Complex.abs x ≤ 1 := by exact_mod_cast hx exact_mod_cast Complex.abs_exp_sub_one_sub_id_le this #align real.abs_exp_sub_one_sub_id_le Real.abs_exp_sub_one_sub_id_le /-- A finite initial segment of the exponential series, followed by an arbitrary tail. For fixed `n` this is just a linear map wrt `r`, and each map is a simple linear function -of the previous (see `exp_near_succ`), with `exp_near n x r ⟶ exp x` as `n ⟶ ∞`, +of the previous (see `expNear_succ`), with `expNear n x r ⟶ exp x` as `n ⟶ ∞`, for any `r`. -/ def expNear (n : ℕ) (x r : ℝ) : ℝ := (∑ m in range n, x ^ m / m.factorial) + x ^ n / n.factorial * r #align real.exp_near Real.expNear @[simp] -theorem expNear_zero (x r) : expNear 0 x r = r := by simp [exp_near] +theorem expNear_zero (x r) : expNear 0 x r = r := by simp [expNear] #align real.exp_near_zero Real.expNear_zero @[simp] theorem expNear_succ (n x r) : expNear (n + 1) x r = expNear n x (1 + x / (n + 1) * r) := by - simp [exp_near, range_succ, mul_add, add_left_comm, add_assoc, pow_succ, div_eq_mul_inv, - mul_inv] <;> - ac_rfl + simp [expNear, range_succ, mul_add, add_left_comm, add_assoc, pow_succ, div_eq_mul_inv, + mul_inv] + ac_rfl #align real.exp_near_succ Real.expNear_succ theorem expNear_sub (n x r₁ r₂) : expNear n x r₁ - expNear n x r₂ = x ^ n / n.factorial * (r₁ - r₂) := by - simp [exp_near, mul_sub] + simp [expNear, mul_sub] #align real.exp_near_sub Real.expNear_sub theorem exp_approx_end (n m : ℕ) (x : ℝ) (e₁ : n + 1 = m) (h : |x| ≤ 1) : |exp x - expNear m x 0| ≤ |x| ^ m / m.factorial * ((m + 1) / m) := by - simp [exp_near] + simp [expNear] convert exp_bound h _ using 1 field_simp [mul_comm] linarith @@ -1803,7 +1800,7 @@ theorem exp_approx_succ {n} {x a₁ b₁ : ℝ} (m : ℕ) (e₁ : n + 1 = m) (a (h : |exp x - expNear m x a₂| ≤ |x| ^ m / m.factorial * b₂) : |exp x - expNear n x a₁| ≤ |x| ^ n / n.factorial * b₁ := by refine' (abs_sub_le _ _ _).trans ((add_le_add_right h _).trans _) - subst e₁; rw [exp_near_succ, exp_near_sub, abs_mul] + subst e₁; rw [expNear_succ, expNear_sub, abs_mul] convert mul_le_mul_of_nonneg_left (le_sub_iff_add_le'.1 e) _ · simp [mul_add, pow_succ', div_eq_mul_inv, abs_mul, abs_inv, ← pow_abs, mul_inv] ac_rfl @@ -1813,7 +1810,8 @@ theorem exp_approx_succ {n} {x a₁ b₁ : ℝ} (m : ℕ) (e₁ : n + 1 = m) (a theorem exp_approx_end' {n} {x a b : ℝ} (m : ℕ) (e₁ : n + 1 = m) (rm : ℝ) (er : ↑m = rm) (h : |x| ≤ 1) (e : |1 - a| ≤ b - |x| / rm * ((rm + 1) / rm)) : |exp x - expNear n x a| ≤ |x| ^ n / n.factorial * b := by - subst er <;> exact exp_approx_succ _ e₁ _ _ (by simpa using e) (exp_approx_end _ _ _ e₁ h) + subst er + exact exp_approx_succ _ e₁ _ _ (by simpa using e) (exp_approx_end _ _ _ e₁ h) #align real.exp_approx_end' Real.exp_approx_end' theorem exp_1_approx_succ_eq {n} {a₁ b₁ : ℝ} {m : ℕ} (en : n + 1 = m) {rm : ℝ} (er : ↑m = rm) @@ -1821,7 +1819,7 @@ theorem exp_1_approx_succ_eq {n} {a₁ b₁ : ℝ} {m : ℕ} (en : n + 1 = m) {r |exp 1 - expNear n 1 a₁| ≤ |1| ^ n / n.factorial * b₁ := by subst er refine' exp_approx_succ _ en _ _ _ h - field_simp [show (m : ℝ) ≠ 0 by norm_cast <;> linarith] + field_simp [show (m : ℝ) ≠ 0 by norm_cast; linarith] #align real.exp_1_approx_succ_eq Real.exp_1_approx_succ_eq theorem exp_approx_start (x a b : ℝ) (h : |exp x - expNear 0 x a| ≤ |x| ^ 0 / 0! * b) : @@ -1885,7 +1883,7 @@ theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| _ ≤ abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) * I / 2) + abs (-((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) * I) / 2) := - by rw [sub_mul, sub_eq_add_neg, add_div] <;> exact complex.abs.add_le _ _ + by rw [sub_mul, sub_eq_add_neg, add_div] <;> exact Complex.abs.add_le _ _ _ = abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2 + abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) / 2 := @@ -1910,8 +1908,8 @@ theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := add_le_add (mul_le_mul_of_nonneg_right (pow_le_one _ (abs_nonneg _) hx) (by norm_num)) ((div_le_div_right (by norm_num)).2 (by - rw [sq, ← abs_mul_self, abs_mul] <;> - exact mul_le_one hx (abs_nonneg _) hx)) + rw [sq, ← abs_mul_self, abs_mul] + exact mul_le_one hx (abs_nonneg _) hx)) _ < 1 := by norm_num ) _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 @@ -1930,8 +1928,8 @@ theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < si (calc |x| ^ 4 ≤ |x| ^ 1 := pow_le_pow_of_le_one (abs_nonneg _) - (by rwa [abs_of_nonneg (le_of_lt hx0)]) (by decide) - _ = x := by simp [abs_of_nonneg (le_of_lt hx0)] + (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]) (by decide) + _ = x := by simp [_root_.abs_of_nonneg (le_of_lt hx0)] ) (by norm_num)) ((div_le_div_right (by norm_num)).2 @@ -1942,7 +1940,7 @@ theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < si _ < x := by linarith ) _ ≤ sin x := - sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [abs_of_nonneg (le_of_lt hx0)]))).2 + sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]))).2 #align real.sin_pos_of_pos_of_le_one Real.sin_pos_of_pos_of_le_one @@ -1951,7 +1949,7 @@ theorem sin_pos_of_pos_of_le_two {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 2) : 0 < si calc 0 < 2 * sin (x / 2) * cos (x / 2) := mul_pos (mul_pos (by norm_num) (sin_pos_of_pos_of_le_one (half_pos hx0) this)) - (cos_pos_of_le_one (by rwa [abs_of_nonneg (le_of_lt (half_pos hx0))])) + (cos_pos_of_le_one (by rwa [_root_.abs_of_nonneg (le_of_lt (half_pos hx0))])) _ = sin x := by rw [← sin_two_mul, two_mul, add_halves] #align real.sin_pos_of_pos_of_le_two Real.sin_pos_of_pos_of_le_two @@ -2043,7 +2041,7 @@ theorem add_one_le_exp_of_nonpos {x : ℝ} (h : x ≤ 0) : x + 1 ≤ Real.exp x #align real.add_one_le_exp_of_nonpos Real.add_one_le_exp_of_nonpos theorem add_one_le_exp (x : ℝ) : x + 1 ≤ Real.exp x := by - cases le_or_lt 0 x + cases' le_or_lt 0 x with h h · exact Real.add_one_le_exp_of_nonneg h exact add_one_le_exp_of_nonpos h.le #align real.add_one_le_exp Real.add_one_le_exp @@ -2082,21 +2080,24 @@ namespace Complex @[simp] theorem abs_cos_add_sin_mul_I (x : ℝ) : abs (cos x + sin x * I) = 1 := by have := Real.sin_sq_add_cos_sq x - simp_all [add_comm, abs, norm_sq, sq, sin_of_real_re, cos_of_real_re, mul_re] -#align complex.abs_cos_add_sin_mul_I Complex.abs_cos_add_sin_mul_i + simp_all [add_comm, abs, normSq, sq, sin_ofReal_re, cos_ofReal_re, mul_re] +set_option linter.uppercaseLean3 false in +#align complex.abs_cos_add_sin_mul_I Complex.abs_cos_add_sin_mul_I @[simp] -theorem abs_exp_of_real (x : ℝ) : abs (exp x) = Real.exp x := by - rw [← of_real_exp] <;> exact abs_of_nonneg (le_of_lt (Real.exp_pos _)) -#align complex.abs_exp_of_real Complex.abs_exp_of_real +theorem abs_exp_ofReal (x : ℝ) : abs (exp x) = Real.exp x := by + rw [← ofReal_exp] + exact abs_of_nonneg (le_of_lt (Real.exp_pos _)) +#align complex.abs_exp_of_real Complex.abs_exp_ofReal @[simp] -theorem abs_exp_of_real_mul_I (x : ℝ) : abs (exp (x * I)) = 1 := by +theorem abs_exp_ofReal_mul_I (x : ℝ) : abs (exp (x * I)) = 1 := by rw [exp_mul_I, abs_cos_add_sin_mul_I] -#align complex.abs_exp_of_real_mul_I Complex.abs_exp_of_real_mul_i +set_option linter.uppercaseLean3 false in +#align complex.abs_exp_of_real_mul_I Complex.abs_exp_ofReal_mul_I theorem abs_exp (z : ℂ) : abs (exp z) = Real.exp z.re := by - rw [exp_eq_exp_re_mul_sin_add_cos, map_mul, abs_exp_of_real, abs_cos_add_sin_mul_I, mul_one] + rw [exp_eq_exp_re_mul_sin_add_cos, map_mul, abs_exp_ofReal, abs_cos_add_sin_mul_I, mul_one] #align complex.abs_exp Complex.abs_exp theorem abs_exp_eq_iff_re_eq {x y : ℂ} : abs (exp x) = abs (exp y) ↔ x.re = y.re := by From eb8cfe717037957c25f8a9df77ab7f173846814d Mon Sep 17 00:00:00 2001 From: qawbecrdtey Date: Mon, 13 Mar 2023 03:24:10 +0000 Subject: [PATCH 14/23] Fixed exp_bound' --- Mathlib/Data/Complex/Exponential.lean | 26 +++++++++++++++++++++++--- 1 file changed, 23 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 4968f5e5af3ac..095d4857f09ce 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1700,13 +1700,33 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : · simp_rw [← div_pow] rw [geom_sum_eq, div_le_iff_of_neg] · trans (-1 : ℝ) - · linarith + · -- Porting note: was linarith + simp [Nat.succ_eq_add_one] at hx + have : (↑(abs x) / ((n : ℝ) + 1) - 1 ≤ - (1 : ℝ) / 2) := by + norm_num + simp [hx] + rw [mul_comm, ← le_div_iff] + simp [hx] + . norm_num [this, hx] + simp [hx] + . exact zero_lt_two · simp only [neg_le_sub_iff_le_add, div_pow, Nat.cast_succ, le_add_iff_nonneg_left] exact div_nonneg (pow_nonneg (abs.nonneg x) k) (pow_nonneg (add_nonneg n.cast_nonneg zero_le_one) k) - · linarith - · linarith + · -- Porting note: was linarith + simp [Nat.succ_eq_add_one] at hx + have : (↑(abs x) / ((n : ℝ) + 1) - 1 ≤ - (1 : ℝ) / 2) := by + norm_num + simp [hx] + simp + apply lt_of_le_of_lt hx + norm_num + · -- Porting note: was linarith + intro h + simp at h + simp [h] at hx + norm_num at hx · exact div_nonneg (pow_nonneg (abs.nonneg x) n) (Nat.cast_nonneg n.factorial) #align complex.exp_bound' Complex.exp_bound' From 13511133a029a6fb6a6474d425921b7ab4661375 Mon Sep 17 00:00:00 2001 From: qawbecrdtey Date: Mon, 13 Mar 2023 03:31:01 +0000 Subject: [PATCH 15/23] Indent fix. --- Mathlib/Data/Complex/Exponential.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 095d4857f09ce..df36f14af190f 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1690,7 +1690,7 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : _ ≤ ∑ i : ℕ in range k, abs x ^ (n + i) / ((n.factorial : ℝ) * (n.succ : ℝ) ^ i) := ?_ _ = ∑ i : ℕ in range k, abs x ^ n / n.factorial * (abs x ^ i / (n.succ : ℝ) ^ i) := ?_ _ ≤ abs x ^ n / ↑n.factorial * 2 := ?_ - · refine' sum_le_sum fun m hm => div_le_div (pow_nonneg (abs.nonneg x) (n + m)) le_rfl _ _ + · refine' sum_le_sum fun m _ => div_le_div (pow_nonneg (abs.nonneg x) (n + m)) le_rfl _ _ · exact_mod_cast mul_pos n.factorial_pos (pow_pos n.succ_pos _) · exact_mod_cast Nat.factorial_mul_pow_le_factorial · refine' Finset.sum_congr rfl fun _ _ => _ @@ -1717,8 +1717,8 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : · -- Porting note: was linarith simp [Nat.succ_eq_add_one] at hx have : (↑(abs x) / ((n : ℝ) + 1) - 1 ≤ - (1 : ℝ) / 2) := by - norm_num - simp [hx] + norm_num + simp [hx] simp apply lt_of_le_of_lt hx norm_num From 53398c8cc28c8085d05f99498256a682bee73cfc Mon Sep 17 00:00:00 2001 From: qawbecrdtey Date: Mon, 13 Mar 2023 03:41:43 +0000 Subject: [PATCH 16/23] Fixed one_sub_le_exp_minus_of_pos --- Mathlib/Data/Complex/Exponential.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index df36f14af190f..73c1a352c2b98 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -2044,10 +2044,10 @@ theorem one_sub_le_exp_minus_of_pos {y : ℝ} (h : 0 ≤ y) : 1 - y ≤ Real.exp · have h'' : (1 - y) * y.exp ≤ 0 := by rw [mul_nonpos_iff] right - exact ⟨h_1, y.exp_pos.le⟩ + exact ⟨by assumption, y.exp_pos.le⟩ linarith have hy1 : y < 1 := by linarith - rw [← le_div_iff' h_1] + rw [← le_div_iff' ‹0 < 1 - y›] exact exp_bound_div_one_sub_of_interval h hy1 rw [inv_eq_one_div] rw [le_div_iff' y.exp_pos] From 2cf9b2e6a7204bd9ce2710f134c35e276de151c7 Mon Sep 17 00:00:00 2001 From: qawbecrdtey Date: Mon, 13 Mar 2023 03:51:29 +0000 Subject: [PATCH 17/23] Fixed exp_bound_div_one_sub_of_interval --- Mathlib/Data/Complex/Exponential.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 73c1a352c2b98..703dd6aced34a 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -2025,12 +2025,13 @@ theorem exp_bound_div_one_sub_of_interval {x : ℝ} (h1 : 0 ≤ x) (h2 : x < 1) by norm_num [Finset.sum] have h1x : 0 < 1 - x := by simpa - rw [le_div_iff h1x] + rw [inv_eq_one_div, le_div_iff h1x] norm_num [← add_assoc, mul_sub_left_distrib, mul_one, add_mul, sub_add_eq_sub_sub, pow_succ' x 2] have hx3 : 0 ≤ x ^ 3 := by norm_num - exact h1 + simp [h1] + simp [Finset.sum] linarith (exp_bound' h1 h2.le <| by linarith).trans ((exp_bound_div_one_sub_of_interval_approx h1 h2.le).trans h) From 5b76f9be5f23d5fdb924035904eb14bad3eae49c Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Mon, 13 Mar 2023 13:56:53 +0000 Subject: [PATCH 18/23] small fixes --- Mathlib/Data/Complex/Exponential.lean | 23 +++++++++++++++++------ 1 file changed, 17 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 703dd6aced34a..a08ddfedc1011 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1759,7 +1759,10 @@ open Complex Finset nonrec theorem exp_bound {x : ℝ} (hx : |x| ≤ 1) {n : ℕ} (hn : 0 < n) : |exp x - ∑ m in range n, x ^ m / m.factorial| ≤ |x| ^ n * (n.succ / (n.factorial * n)) := by have hxc : Complex.abs x ≤ 1 := by exact_mod_cast hx - convert exp_bound hxc hn <;> norm_cast + convert exp_bound hxc hn <;> + --Porting note: was `norm_cast` + simp only [← abs_ofReal, ← ofReal_sub, ← ofReal_exp, ← ofReal_sum, ← ofReal_pow, + ← ofReal_div, ← ofReal_nat_cast] #align real.exp_bound Real.exp_bound theorem exp_bound' {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) {n : ℕ} (hn : 0 < n) : @@ -1774,21 +1777,29 @@ theorem exp_bound' {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) {n : ℕ} (hn : 0 < n #align real.exp_bound' Real.exp_bound' theorem abs_exp_sub_one_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1| ≤ 2 * |x| := by - have : Complex.abs x ≤ 1 := by exact_mod_cast hx - exact_mod_cast Complex.abs_exp_sub_one_le this + have : abs' x ≤ 1 := by exact_mod_cast hx + --Porting note: was + --exact_mod_cast Complex.abs_exp_sub_one_le (x := x) this + have := Complex.abs_exp_sub_one_le (x := x) (by simpa using this) + rw [← ofReal_exp, ← ofReal_one, ← ofReal_sub, abs_ofReal, abs_ofReal] at this + exact this #align real.abs_exp_sub_one_le Real.abs_exp_sub_one_le theorem abs_exp_sub_one_sub_id_le {x : ℝ} (hx : |x| ≤ 1) : |exp x - 1 - x| ≤ x ^ 2 := by rw [← _root_.sq_abs] + --Porting note: was + --exact_mod_cast Complex.abs_exp_sub_one_sub_id_le this have : Complex.abs x ≤ 1 := by exact_mod_cast hx - exact_mod_cast Complex.abs_exp_sub_one_sub_id_le this + have := Complex.abs_exp_sub_one_sub_id_le this + rw [← ofReal_one, ← ofReal_exp, ← ofReal_sub, ← ofReal_sub, abs_ofReal, abs_ofReal] at this + exact this #align real.abs_exp_sub_one_sub_id_le Real.abs_exp_sub_one_sub_id_le /-- A finite initial segment of the exponential series, followed by an arbitrary tail. For fixed `n` this is just a linear map wrt `r`, and each map is a simple linear function of the previous (see `expNear_succ`), with `expNear n x r ⟶ exp x` as `n ⟶ ∞`, for any `r`. -/ -def expNear (n : ℕ) (x r : ℝ) : ℝ := +noncomputable def expNear (n : ℕ) (x r : ℝ) : ℝ := (∑ m in range n, x ^ m / m.factorial) + x ^ n / n.factorial * r #align real.exp_near Real.expNear @@ -1810,7 +1821,7 @@ theorem expNear_sub (n x r₁ r₂) : expNear n x r₁ - expNear n x r₂ = x ^ theorem exp_approx_end (n m : ℕ) (x : ℝ) (e₁ : n + 1 = m) (h : |x| ≤ 1) : |exp x - expNear m x 0| ≤ |x| ^ m / m.factorial * ((m + 1) / m) := by simp [expNear] - convert exp_bound h _ using 1 + convert exp_bound (n := m) h ?_ using 1 field_simp [mul_comm] linarith #align real.exp_approx_end Real.exp_approx_end From f3717a65f1b36f1f613363c270ca78f10e8e5d16 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Mon, 13 Mar 2023 14:00:53 +0000 Subject: [PATCH 19/23] more fixes --- Mathlib/Data/Complex/Exponential.lean | 15 ++++++++------- 1 file changed, 8 insertions(+), 7 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index a08ddfedc1011..aaa7246fa7e54 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1832,7 +1832,8 @@ theorem exp_approx_succ {n} {x a₁ b₁ : ℝ} (m : ℕ) (e₁ : n + 1 = m) (a |exp x - expNear n x a₁| ≤ |x| ^ n / n.factorial * b₁ := by refine' (abs_sub_le _ _ _).trans ((add_le_add_right h _).trans _) subst e₁; rw [expNear_succ, expNear_sub, abs_mul] - convert mul_le_mul_of_nonneg_left (le_sub_iff_add_le'.1 e) _ + convert mul_le_mul_of_nonneg_left (a := abs' x ^ n / ↑(Nat.factorial n)) + (le_sub_iff_add_le'.1 e) ?_ · simp [mul_add, pow_succ', div_eq_mul_inv, abs_mul, abs_inv, ← pow_abs, mul_inv] ac_rfl · simp [div_nonneg, abs_nonneg] @@ -1853,13 +1854,13 @@ theorem exp_1_approx_succ_eq {n} {a₁ b₁ : ℝ} {m : ℕ} (en : n + 1 = m) {r field_simp [show (m : ℝ) ≠ 0 by norm_cast; linarith] #align real.exp_1_approx_succ_eq Real.exp_1_approx_succ_eq -theorem exp_approx_start (x a b : ℝ) (h : |exp x - expNear 0 x a| ≤ |x| ^ 0 / 0! * b) : +theorem exp_approx_start (x a b : ℝ) (h : |exp x - expNear 0 x a| ≤ |x| ^ 0 / Nat.factorial 0 * b) : |exp x - a| ≤ b := by simpa using h #align real.exp_approx_start Real.exp_approx_start theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| ^ 4 * (5 / 96) := calc - |cos x - (1 - x ^ 2 / 2)| = abs (Complex.cos x - (1 - x ^ 2 / 2)) := by + |cos x - (1 - x ^ 2 / 2)| = abs' (Complex.cos x - (1 - x ^ 2 / 2)) := by rw [← abs_of_real] <;> simp [of_real_bit0, of_real_one, of_real_inv] _ = abs ((Complex.exp (x * I) + Complex.exp (-x * I) - (2 - x ^ 2)) / 2) := by simp [Complex.cos, sub_div, add_div, neg_div, div_self (two_ne_zero' ℂ)] @@ -1868,12 +1869,12 @@ theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| (((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) + (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial)) / 2) := - (congr_arg abs + (congr_arg abs' (congr_arg (fun x : ℂ => x / 2) (by simp only [sum_range_succ] simp [pow_succ] - apply Complex.ext <;> simp [div_eq_mul_inv, norm_sq] <;> ring))) + apply Complex.ext <;> simp [div_eq_mul_inv, normSq] <;> ring))) _ ≤ abs ((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2) + abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) / 2) := @@ -1883,8 +1884,8 @@ theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) / 2 := by simp [map_div₀] _ ≤ - Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + - Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 := + Complex.abs (x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ))⁻¹) / 2 + + Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ))⁻¹) / 2 := (add_le_add ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide))) ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide)))) _ ≤ |x| ^ 4 * (5 / 96) := by From ed9c891e912031e4148e11c75bc806d99b388ff4 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Mon, 13 Mar 2023 14:47:26 +0000 Subject: [PATCH 20/23] fix build --- Mathlib/Data/Complex/Basic.lean | 8 ++ Mathlib/Data/Complex/Exponential.lean | 151 +++++++++++--------------- 2 files changed, 73 insertions(+), 86 deletions(-) diff --git a/Mathlib/Data/Complex/Basic.lean b/Mathlib/Data/Complex/Basic.lean index 13781de832b07..4419f4d0987f0 100644 --- a/Mathlib/Data/Complex/Basic.lean +++ b/Mathlib/Data/Complex/Basic.lean @@ -483,6 +483,14 @@ set_option linter.uppercaseLean3 false in theorem ofReal_ofNat (n : ℕ) [n.AtLeastTwo] : ((OfNat.ofNat n : ℝ) : ℂ) = OfNat.ofNat n := rfl +@[simp] +theorem re_ofNat (n : ℕ) [n.AtLeastTwo] : (OfNat.ofNat n : ℂ).re = OfNat.ofNat n := + rfl + +@[simp] +theorem im_ofNat (n : ℕ) [n.AtLeastTwo] : (OfNat.ofNat n : ℂ).im = 0 := + rfl + end /-! ### Complex conjugation -/ diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index aaa7246fa7e54..d06f75653983e 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1860,79 +1860,64 @@ theorem exp_approx_start (x a b : ℝ) (h : |exp x - expNear 0 x a| ≤ |x| ^ 0 theorem cos_bound {x : ℝ} (hx : |x| ≤ 1) : |cos x - (1 - x ^ 2 / 2)| ≤ |x| ^ 4 * (5 / 96) := calc - |cos x - (1 - x ^ 2 / 2)| = abs' (Complex.cos x - (1 - x ^ 2 / 2)) := by - rw [← abs_of_real] <;> simp [of_real_bit0, of_real_one, of_real_inv] - _ = abs ((Complex.exp (x * I) + Complex.exp (-x * I) - (2 - x ^ 2)) / 2) := by + |cos x - (1 - x ^ 2 / 2)| = Complex.abs (Complex.cos x - (1 - (x : ℂ) ^ 2 / 2)) := by + rw [← abs_ofReal]; simp + _ = Complex.abs ((Complex.exp (x * I) + Complex.exp (-x * I) - (2 - (x : ℂ) ^ 2)) / 2) := by simp [Complex.cos, sub_div, add_div, neg_div, div_self (two_ne_zero' ℂ)] - _ = - abs + _ = abs (((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) + - (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial)) / - 2) := - (congr_arg abs' + (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial)) / 2) := + (congr_arg Complex.abs (congr_arg (fun x : ℂ => x / 2) (by simp only [sum_range_succ] simp [pow_succ] - apply Complex.ext <;> simp [div_eq_mul_inv, normSq] <;> ring))) - _ ≤ - abs ((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2) + + apply Complex.ext <;> simp [div_eq_mul_inv, normSq] <;> ring_nf + ))) + _ ≤ abs ((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2) + abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) / 2) := - by rw [add_div] <;> exact complex.abs.add_le _ _ - _ = - abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2 + + by rw [add_div]; exact Complex.abs.add_le _ _ + _ = abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2 + abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) / 2 := by simp [map_div₀] - _ ≤ - Complex.abs (x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ))⁻¹) / 2 + - Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ))⁻¹) / 2 := + _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 + + Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * ((Nat.factorial 4) * (4 : ℕ) : ℝ)⁻¹) / 2 := (add_le_add ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide))) ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide)))) - _ ≤ |x| ^ 4 * (5 / 96) := by - norm_num <;> simp [mul_assoc, mul_comm, mul_left_comm, mul_div_assoc] - + _ ≤ |x| ^ 4 * (5 / 96) := by norm_num #align real.cos_bound Real.cos_bound theorem sin_bound {x : ℝ} (hx : |x| ≤ 1) : |sin x - (x - x ^ 3 / 6)| ≤ |x| ^ 4 * (5 / 96) := calc - |sin x - (x - x ^ 3 / 6)| = abs (Complex.sin x - (x - x ^ 3 / 6)) := by - rw [← abs_of_real] <;> simp [of_real_bit0, of_real_one, of_real_inv] - _ = abs (((Complex.exp (-x * I) - Complex.exp (x * I)) * I - (2 * x - x ^ 3 / 3)) / 2) := by + |sin x - (x - x ^ 3 / 6)| = Complex.abs (Complex.sin x - (x - x ^ 3 / 6 : ℝ)) := by + rw [← abs_ofReal]; simp + _ = Complex.abs (((Complex.exp (-x * I) - Complex.exp (x * I)) * I - + (2 * x - x ^ 3 / 3 : ℝ)) / 2) := by simp [Complex.sin, sub_div, add_div, neg_div, mul_div_cancel_left _ (two_ne_zero' ℂ), div_div, show (3 : ℂ) * 2 = 6 by norm_num] - _ = - abs - (((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) - - (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial)) * - I / - 2) := - (congr_arg abs + _ = Complex.abs (((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) - + (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial)) * I / 2) := + (congr_arg Complex.abs (congr_arg (fun x : ℂ => x / 2) (by simp only [sum_range_succ] simp [pow_succ] - apply Complex.ext <;> simp [div_eq_mul_inv, norm_sq] <;> ring))) - _ ≤ - abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) * I / 2) + + apply Complex.ext <;> simp [div_eq_mul_inv, normSq]; ring))) + _ ≤ abs ((Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) * I / 2) + abs (-((Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) * I) / 2) := - by rw [sub_mul, sub_eq_add_neg, add_div] <;> exact Complex.abs.add_le _ _ - _ = - abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2 + + by rw [sub_mul, sub_eq_add_neg, add_div]; exact Complex.abs.add_le _ _ + _ = abs (Complex.exp (x * I) - ∑ m in range 4, (x * I) ^ m / m.factorial) / 2 + abs (Complex.exp (-x * I) - ∑ m in range 4, (-x * I) ^ m / m.factorial) / 2 := by simp [add_comm, map_div₀] - _ ≤ - Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 + - Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * (4! * (4 : ℕ))⁻¹) / 2 := + _ ≤ Complex.abs (x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 + + Complex.abs (-x * I) ^ 4 * (Nat.succ 4 * (Nat.factorial 4 * (4 : ℕ) : ℝ)⁻¹) / 2 := (add_le_add ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide))) ((div_le_div_right (by norm_num)).2 (Complex.exp_bound (by simpa) (by decide)))) - _ ≤ |x| ^ 4 * (5 / 96) := by - norm_num <;> simp [mul_assoc, mul_comm, mul_left_comm, mul_div_assoc] - + _ ≤ |x| ^ 4 * (5 / 96) := by norm_num #align real.sin_bound Real.sin_bound theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := - calc - 0 < 1 - x ^ 2 / 2 - |x| ^ 4 * (5 / 96) := + calc 0 < 1 - x ^ 2 / 2 - |x| ^ 4 * (5 / 96) := sub_pos.2 <| lt_sub_iff_add_lt.2 (calc @@ -1942,17 +1927,13 @@ theorem cos_pos_of_le_one {x : ℝ} (hx : |x| ≤ 1) : 0 < cos x := (by rw [sq, ← abs_mul_self, abs_mul] exact mul_le_one hx (abs_nonneg _) hx)) - _ < 1 := by norm_num - ) + _ < 1 := by norm_num) _ ≤ cos x := sub_le_comm.1 (abs_sub_le_iff.1 (cos_bound hx)).2 - #align real.cos_pos_of_le_one Real.cos_pos_of_le_one theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < sin x := - calc - 0 < x - x ^ 3 / 6 - |x| ^ 4 * (5 / 96) := - sub_pos.2 <| - lt_sub_iff_add_lt.2 + calc 0 < x - x ^ 3 / 6 - |x| ^ 4 * (5 / 96) := + sub_pos.2 <| lt_sub_iff_add_lt.2 (calc |x| ^ 4 * (5 / 96) + x ^ 3 / 6 ≤ x * (5 / 96) + x / 6 := add_le_add @@ -1969,11 +1950,11 @@ theorem sin_pos_of_pos_of_le_one {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 1) : 0 < si x ^ 3 ≤ x ^ 1 := pow_le_pow_of_le_one (le_of_lt hx0) hx (by decide) _ = x := pow_one _ )) - _ < x := by linarith - ) + --Porting note : was `_ < x := by linarith` + _ = x * (7 / 32) := by ring + _ < x := (mul_lt_iff_lt_one_right hx0).2 (by norm_num)) _ ≤ sin x := sub_le_comm.1 (abs_sub_le_iff.1 (sin_bound (by rwa [_root_.abs_of_nonneg (le_of_lt hx0)]))).2 - #align real.sin_pos_of_pos_of_le_one Real.sin_pos_of_pos_of_le_one theorem sin_pos_of_pos_of_le_two {x : ℝ} (hx0 : 0 < x) (hx : x ≤ 2) : 0 < sin x := @@ -1999,8 +1980,7 @@ theorem cos_one_pos : 0 < cos 1 := #align real.cos_one_pos Real.cos_one_pos theorem cos_two_neg : cos 2 < 0 := - calc - cos 2 = cos (2 * 1) := congr_arg cos (mul_one _).symm + calc cos 2 = cos (2 * 1) := congr_arg cos (mul_one _).symm _ = _ := (Real.cos_two_mul 1) _ ≤ 2 * (2 / 3) ^ 2 - 1 := (sub_le_sub_right @@ -2008,28 +1988,26 @@ theorem cos_two_neg : cos 2 < 0 := (by rw [sq, sq] exact mul_self_le_mul_self (le_of_lt cos_one_pos) cos_one_le) - zero_le_two) - _) + zero_le_two) _) _ < 0 := by norm_num - #align real.cos_two_neg Real.cos_two_neg -theorem exp_bound_div_one_sub_of_interval_approx {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) : +--Porting note: removed `(h1 : 0 ≤ x)` because it is no longer used +theorem exp_bound_div_one_sub_of_interval_approx {x : ℝ} (h2 : x ≤ 1) : (∑ j : ℕ in Finset.range 3, x ^ j / j.factorial) + x ^ 3 * ((3 : ℕ) + 1) / ((3 : ℕ).factorial * (3 : ℕ)) ≤ - ∑ j in Finset.range 3, x ^ j := by - norm_num [Finset.sum] - rw [add_assoc, add_comm (x + 1) (x ^ 3 * 4 / 18), ← add_assoc, add_le_add_iff_right, ← - add_le_add_iff_left (-(x ^ 2 / 2)), ← add_assoc, CommRing.add_left_neg (x ^ 2 / 2), zero_add, - neg_add_eq_sub, sub_half, sq, pow_succ, sq] - have i1 : x * 4 / 18 ≤ 1 / 2 := by linarith - have i2 : 0 ≤ x * 4 / 18 := by linarith - have i3 := mul_le_mul h1 h1 le_rfl h1 - rw [zero_mul] at i3 - have t := mul_le_mul le_rfl i1 i2 i3 - rw [← mul_assoc] - rwa [mul_one_div, ← mul_div_assoc, ← mul_assoc] at t -#align real.exp_bound_div_one_sub_of_interval_approx Real.exp_bound_div_one_sub_of_interval_approx + ∑ j in Finset.range 3, x ^ j := + calc + (∑ j : ℕ in Finset.range 3, x ^ j / j.factorial) + + x ^ 3 * ((3 : ℕ) + 1) / ((3 : ℕ).factorial * (3 : ℕ)) + = (2 / 9) * x ^ 3 + x ^ 2 / 2 + x + 1 := by simp [Finset.sum]; ring + _ ≤ x ^ 2 + x + 1 := sub_nonneg.1 <| + calc 0 ≤ x^2 * (2 / 9) * (9 / 4 - x) := + mul_nonneg (mul_nonneg (pow_two_nonneg _) (by norm_num : (0 : ℝ) ≤ 2 / 9)) + (sub_nonneg.2 (le_trans h2 (by norm_num))) + _ = _ := by ring + _ = _ := by simp [Finset.sum]; ring +#align real.exp_bound_div_one_sub_of_interval_approx Real.exp_bound_div_one_sub_of_interval_approxₓ theorem exp_bound_div_one_sub_of_interval {x : ℝ} (h1 : 0 ≤ x) (h2 : x < 1) : Real.exp x ≤ 1 / (1 - x) := @@ -2046,7 +2024,7 @@ theorem exp_bound_div_one_sub_of_interval {x : ℝ} (h1 : 0 ≤ x) (h2 : x < 1) simp [Finset.sum] linarith (exp_bound' h1 h2.le <| by linarith).trans - ((exp_bound_div_one_sub_of_interval_approx h1 h2.le).trans h) + ((exp_bound_div_one_sub_of_interval_approx h2.le).trans h) #align real.exp_bound_div_one_sub_of_interval Real.exp_bound_div_one_sub_of_interval theorem one_sub_le_exp_minus_of_pos {y : ℝ} (h : 0 ≤ y) : 1 - y ≤ Real.exp (-y) := by @@ -2083,30 +2061,31 @@ theorem one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t rcases eq_or_ne n 0 with (rfl | hn) · simp rwa [Nat.cast_zero] at ht' - convert pow_le_pow_of_le_left _ (add_one_le_exp (-(t / n))) n + convert pow_le_pow_of_le_left ?_ (add_one_le_exp (-(t / n))) n · abel · rw [← Real.exp_nat_mul] congr 1 - field_simp [nat.cast_ne_zero.mpr hn] - ring + field_simp [(Nat.cast_ne_zero (R := ℝ)).mpr hn] + ring_nf · rwa [add_comm, ← sub_eq_add_neg, sub_nonneg, div_le_one] positivity #align real.one_sub_div_pow_le_exp_neg Real.one_sub_div_pow_le_exp_neg end Real -namespace Tactic +--Porting note: TODO: write this extension +-- namespace Tactic -open Positivity Real +-- open Positivity Real -/-- Extension for the `positivity` tactic: `real.exp` is always positive. -/ -@[positivity] -unsafe def positivity_exp : expr → tactic strictness - | q(Real.exp $(a)) => positive <$> mk_app `real.exp_pos [a] - | e => pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `real.exp r`" -#align tactic.positivity_exp tactic.positivity_exp +-- /-- Extension for the `positivity` tactic: `real.exp` is always positive. -/ +-- @[positivity] +-- unsafe def positivity_exp : expr → tactic strictness +-- | q(Real.exp $(a)) => positive <$> mk_app `real.exp_pos [a] +-- | e => pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `real.exp r`" +-- #align tactic.positivity_exp tactic.positivity_exp -end Tactic +-- end Tactic namespace Complex From ac5fd49a7a7e334d0a9bcff325fadda35d585eac Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Mon, 13 Mar 2023 14:49:24 +0000 Subject: [PATCH 21/23] fix lint --- Mathlib/Data/Complex/Exponential.lean | 6 ------ 1 file changed, 6 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index d06f75653983e..9688a0367c31a 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1702,9 +1702,6 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : · trans (-1 : ℝ) · -- Porting note: was linarith simp [Nat.succ_eq_add_one] at hx - have : (↑(abs x) / ((n : ℝ) + 1) - 1 ≤ - (1 : ℝ) / 2) := by - norm_num - simp [hx] rw [mul_comm, ← le_div_iff] simp [hx] . norm_num [this, hx] @@ -1716,9 +1713,6 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : (pow_nonneg (add_nonneg n.cast_nonneg zero_le_one) k) · -- Porting note: was linarith simp [Nat.succ_eq_add_one] at hx - have : (↑(abs x) / ((n : ℝ) + 1) - 1 ≤ - (1 : ℝ) / 2) := by - norm_num - simp [hx] simp apply lt_of_le_of_lt hx norm_num From 24ecaf424e344cdf760c04c879dc0eb5ef375179 Mon Sep 17 00:00:00 2001 From: ChrisHughes24 Date: Mon, 13 Mar 2023 14:51:22 +0000 Subject: [PATCH 22/23] line length --- Mathlib/Data/Complex/Exponential.lean | 9 ++++++--- 1 file changed, 6 insertions(+), 3 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 9688a0367c31a..f24f4a1ee9b61 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -1677,7 +1677,8 @@ theorem exp_bound' {x : ℂ} {n : ℕ} (hx : abs x / n.succ ≤ 1 / 2) : exp, sub_eq_add_neg, ← lim_neg, lim_add, ← lim_abs] refine' lim_le (CauSeq.le_of_exists ⟨n, fun j hj => _⟩) simp_rw [← sub_eq_add_neg] - show abs ((∑ m in range j, x ^ m / m.factorial) - ∑ m in range n, x ^ m / m.factorial) ≤ abs x ^ n / n.factorial * 2 + show abs ((∑ m in range j, x ^ m / m.factorial) - ∑ m in range n, x ^ m / m.factorial) ≤ + abs x ^ n / n.factorial * 2 let k := j - n have hj : j = n + k := (add_tsub_cancel_of_le hj).symm rw [hj, sum_range_add_sub_sum_range] @@ -1760,7 +1761,8 @@ nonrec theorem exp_bound {x : ℝ} (hx : |x| ≤ 1) {n : ℕ} (hn : 0 < n) : #align real.exp_bound Real.exp_bound theorem exp_bound' {x : ℝ} (h1 : 0 ≤ x) (h2 : x ≤ 1) {n : ℕ} (hn : 0 < n) : - Real.exp x ≤ (∑ m in Finset.range n, x ^ m / m.factorial) + x ^ n * (n + 1) / (n.factorial * n) := by + Real.exp x ≤ (∑ m in Finset.range n, x ^ m / m.factorial) + + x ^ n * (n + 1) / (n.factorial * n) := by have h3 : |x| = x := by simpa have h4 : |x| ≤ 1 := by rwa [h3] have h' := Real.exp_bound h4 hn @@ -1808,7 +1810,8 @@ theorem expNear_succ (n x r) : expNear (n + 1) x r = expNear n x (1 + x / (n + 1 ac_rfl #align real.exp_near_succ Real.expNear_succ -theorem expNear_sub (n x r₁ r₂) : expNear n x r₁ - expNear n x r₂ = x ^ n / n.factorial * (r₁ - r₂) := by +theorem expNear_sub (n x r₁ r₂) : expNear n x r₁ - + expNear n x r₂ = x ^ n / n.factorial * (r₁ - r₂) := by simp [expNear, mul_sub] #align real.exp_near_sub Real.expNear_sub From 218da65a06bc325c8537802203f518c773e1f50a Mon Sep 17 00:00:00 2001 From: David Renshaw Date: Mon, 13 Mar 2023 14:40:05 -0400 Subject: [PATCH 23/23] implement Real.exp positivity extenstion --- Mathlib/Data/Complex/Exponential.lean | 19 ++++++++----------- 1 file changed, 8 insertions(+), 11 deletions(-) diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index f24f4a1ee9b61..051397b54a134 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -2070,19 +2070,16 @@ theorem one_sub_div_pow_le_exp_neg {n : ℕ} {t : ℝ} (ht' : t ≤ n) : (1 - t end Real ---Porting note: TODO: write this extension --- namespace Tactic +namespace Tactic +open Lean.Meta Qq --- open Positivity Real +/-- Extension for the `positivity` tactic: `real.exp` is always positive. -/ +@[positivity Real.exp _] +def evalExp : Mathlib.Meta.Positivity.PositivityExt where eval {_ _} _ _ e := do + let (.app _ (a : Q(ℝ))) ← withReducible (whnf e) | throwError "not Real.exp" + pure (.positive (q(Real.exp_pos $a) : Lean.Expr)) --- /-- Extension for the `positivity` tactic: `real.exp` is always positive. -/ --- @[positivity] --- unsafe def positivity_exp : expr → tactic strictness --- | q(Real.exp $(a)) => positive <$> mk_app `real.exp_pos [a] --- | e => pp e >>= fail ∘ format.bracket "The expression `" "` isn't of the form `real.exp r`" --- #align tactic.positivity_exp tactic.positivity_exp - --- end Tactic +end Tactic namespace Complex