The type of angles #
In this file we define Real.Angle
to be the quotient group ℝ/2πℤ
and prove a few simple lemmas
about trigonometric functions and angles.
Equations
Equations
- Real.Angle.instCoe = { coe := Real.Angle.coe }
An induction principle to deduce results for Angle
from those for ℝ
, used with
induction θ using Real.Angle.induction_on
.
@[simp]
The tangent of a Real.Angle
.
Instances For
The sign of a Real.Angle
is 0
if the angle is 0
or π
, 1
if the angle is strictly
between 0
and π
and -1
is the angle is strictly between -π
and 0
. It is defined as the
sign of the sine of the angle.
Equations
- θ.sign = SignType.sign θ.sin
Instances For
@[simp]
theorem
ContinuousOn.angle_sign_comp
{α : Type u_1}
[TopologicalSpace α]
{f : α → Real.Angle}
{s : Set α}
(hf : ContinuousOn f s)
(hs : ∀ z ∈ s, f z ≠ 0 ∧ f z ≠ ↑Real.pi)
:
ContinuousOn (Real.Angle.sign ∘ f) s
theorem
Real.Angle.sign_eq_of_continuousOn
{α : Type u_1}
[TopologicalSpace α]
{f : α → Angle}
{s : Set α}
{x y : α}
(hc : IsConnected s)
(hf : ContinuousOn f s)
(hs : ∀ z ∈ s, f z ≠ 0 ∧ f z ≠ ↑Real.pi)
(hx : x ∈ s)
(hy : y ∈ s)
:
Suppose a function to angles is continuous on a connected set and never takes the values 0
or π
on that set. Then the values of the function on that set all have the same sign.