From e67cdd34614843ad8875d812bd4e114331b8bff7 Mon Sep 17 00:00:00 2001 From: mehbark Date: Wed, 22 Feb 2023 22:17:27 -0500 Subject: [PATCH] =?UTF-8?q?tan=C2=B2=20=CE=B8=20+=201=20=3D=20sec=C2=B2=20?= =?UTF-8?q?=CE=B8=20but=20only=20if=20cos=C2=B2=20=CE=B8=20=E2=89=A0=200?= =?UTF-8?q?=3F=3F?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- TrigonometricAnalysis.lean | 89 ++++++++++++++++++++++++++++++++++++++ lake-manifest.json | 27 ++++++++++++ lakefile.lean | 14 ++++++ lean-toolchain | 1 + 4 files changed, 131 insertions(+) create mode 100644 TrigonometricAnalysis.lean create mode 100644 lake-manifest.json create mode 100644 lakefile.lean create mode 100644 lean-toolchain diff --git a/TrigonometricAnalysis.lean b/TrigonometricAnalysis.lean new file mode 100644 index 0000000..a7477dd --- /dev/null +++ b/TrigonometricAnalysis.lean @@ -0,0 +1,89 @@ +-- actually it's analytical trigonometry +import Mathlib.Data.Real.Basic +import Mathlib.Tactic.LibrarySearch + +axiom sin : Real → Real +axiom csc : Real → Real + +axiom sin_eq_reciprocal_of_csc : ∀x, sin x = 1 / (csc x) +axiom csc_eq_reciprocal_of_sin : ∀x, csc x = 1 / (sin x) + +axiom cos : Real → Real +axiom sec : Real → Real + +axiom cos_eq_reciprocal_of_sec : ∀x, cos x = (sec x)⁻¹ +axiom sec_eq_reciprocal_of_cos : ∀x, sec x = (cos x)⁻¹ + +@[simp] +noncomputable def tan (x : Real) : Real := sin x / cos x +theorem tan_def : ∀x, tan x = sin x / cos x := by simp +@[simp] +noncomputable def cot (x : Real) : Real := cos x / sin x + +theorem tan_eq_sin_div_cos : ∀x, tan x = sin x / cos x := by simp +theorem cot_eq_cos_div_sin : ∀x, cot x = cos x / sin x := by simp + +theorem tan_eq_reciprocal_of_cot : ∀x, tan x = 1 / (cot x) := by simp +theorem cot_eq_reciprocal_of_tan : ∀x, cot x = 1 / (tan x) := by simp + +noncomputable def sin2 (x : Real) : Real := sin x * sin x +noncomputable def cos2 (x : Real) : Real := cos x * cos x +noncomputable def tan2 (x : Real) : Real := tan x * tan x +theorem tan2_def : ∀x, tan2 x = sin2 x / cos2 x := by + intro x + simp [tan2, sin2, cos2] + exact div_mul_div_comm (sin x) (cos x) (sin x) (cos x) + +noncomputable def csc2 (x : Real) : Real := csc x * csc x +noncomputable def sec2 (x : Real) : Real := sec x * sec x +theorem sec2_eq_reciprocal_of_cos2 : ∀x, sec2 x = (cos2 x)⁻¹ := by + intro x + rw [sec2, cos2] + simp + rw [sec_eq_reciprocal_of_cos] + +noncomputable def cot2 (x : Real) : Real := cot x * cot x + +noncomputable def sin3 (x : Real) : Real := sin2 x * sin x +noncomputable def cos3 (x : Real) : Real := cos2 x * cos x +noncomputable def tan3 (x : Real) : Real := tan2 x * tan x + +noncomputable def csc3 (x : Real) : Real := csc2 x * csc x +noncomputable def sec3 (x : Real) : Real := sec2 x * sec x +noncomputable def cot3 (x : Real) : Real := cot2 x * cot x + +@[simp] +axiom pyth_id : ∀x, sin2 x + cos2 x = 1 + +theorem sin_squared_eq {x : Real} : sin2 x = 1 - cos2 x := by + rw [← add_left_inj (cos2 x)] + rw [sub_add_cancel] + exact pyth_id x + +theorem cos_squared_eq {x : Real} : cos2 x = 1 - sin2 x := by + rw [← add_left_inj (sin2 x)] + rw [sub_add_cancel] + rw [add_comm] + exact pyth_id x + +theorem bleh {a b c : Real} (d : Real) : a + b = c → a / d + b / d = c / d := by + intro h + rw [← add_div, h] + +def tan_pyth_id : ∀x, cos2 x ≠ 0 → tan2 x + 1 = sec2 x := by + intro x + intro hnz + have h := pyth_id x + have h : sin2 x / cos2 x + cos2 x / cos2 x = 1 / cos2 x := + bleh (cos2 x) h + + rw [← tan2_def] at h + simp at h + rw [← sec2_eq_reciprocal_of_cos2, div_self] at h + exact h + exact hnz + +example (x : Real) : 2 * (1 / (csc x)) = 2 * (sin x) := by + rw [sin_eq_reciprocal_of_csc] + +--todo evens and odds and the sin/cos (a ± b) ones diff --git a/lake-manifest.json b/lake-manifest.json new file mode 100644 index 0000000..af47f84 --- /dev/null +++ b/lake-manifest.json @@ -0,0 +1,27 @@ +{"version": 4, + "packagesDir": "lake-packages", + "packages": + [{"git": + {"url": "https://github.com/leanprover-community/mathlib4.git", + "subDir?": null, + "rev": "076ab7d4b2f10fcbb8d8030c608caa254a3f77af", + "name": "mathlib", + "inputRev?": null}}, + {"git": + {"url": "https://github.com/gebner/quote4", + "subDir?": null, + "rev": "7ac99aa3fac487bec1d5860e751b99c7418298cf", + "name": "Qq", + "inputRev?": "master"}}, + {"git": + {"url": "https://github.com/JLimperg/aesop", + "subDir?": null, + "rev": "ba61f7fec6174d8c7d2796457da5a8d0b0da44c6", + "name": "aesop", + "inputRev?": "master"}}, + {"git": + {"url": "https://github.com/leanprover/std4", + "subDir?": null, + "rev": "de7e2a79905a3f87cad1ad5bf57045206f9738c7", + "name": "std", + "inputRev?": "main"}}]} diff --git a/lakefile.lean b/lakefile.lean new file mode 100644 index 0000000..4f7d3e3 --- /dev/null +++ b/lakefile.lean @@ -0,0 +1,14 @@ +import Lake +open Lake DSL + +package «trigonometric-analysis» { + -- add any package configuration options here +} + +require mathlib from git + "https://github.com/leanprover-community/mathlib4.git" + +@[default_target] +lean_lib «TrigonometricAnalysis» { + -- add any library configuration options here +} diff --git a/lean-toolchain b/lean-toolchain new file mode 100644 index 0000000..5bf01da --- /dev/null +++ b/lean-toolchain @@ -0,0 +1 @@ +leanprover/lean4:nightly-2023-02-10