Documentation

Mathlib.Geometry.Euclidean.Angle.Oriented.RightAngle

Oriented angles in right-angled triangles. #

This file proves basic geometrical results about distances and oriented angles in (possibly degenerate) right-angled triangles in real inner product spaces and Euclidean affine spaces.

theorem Orientation.oangle_add_right_eq_arccos_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle x (x + y) = (x / x + y).arccos

An angle in a right-angled triangle expressed using arccos.

theorem Orientation.oangle_add_left_eq_arccos_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle (x + y) y = (y / x + y).arccos

An angle in a right-angled triangle expressed using arccos.

theorem Orientation.oangle_add_right_eq_arcsin_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle x (x + y) = (y / x + y).arcsin

An angle in a right-angled triangle expressed using arcsin.

theorem Orientation.oangle_add_left_eq_arcsin_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle (x + y) y = (x / x + y).arcsin

An angle in a right-angled triangle expressed using arcsin.

theorem Orientation.oangle_add_right_eq_arctan_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle x (x + y) = (y / x).arctan

An angle in a right-angled triangle expressed using arctan.

theorem Orientation.oangle_add_left_eq_arctan_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle (x + y) y = (x / y).arctan

An angle in a right-angled triangle expressed using arctan.

theorem Orientation.cos_oangle_add_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle x (x + y)).cos = x / x + y

The cosine of an angle in a right-angled triangle as a ratio of sides.

theorem Orientation.cos_oangle_add_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x + y) y).cos = y / x + y

The cosine of an angle in a right-angled triangle as a ratio of sides.

theorem Orientation.sin_oangle_add_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle x (x + y)).sin = y / x + y

The sine of an angle in a right-angled triangle as a ratio of sides.

theorem Orientation.sin_oangle_add_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x + y) y).sin = x / x + y

The sine of an angle in a right-angled triangle as a ratio of sides.

theorem Orientation.tan_oangle_add_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle x (x + y)).tan = y / x

The tangent of an angle in a right-angled triangle as a ratio of sides.

theorem Orientation.tan_oangle_add_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x + y) y).tan = x / y

The tangent of an angle in a right-angled triangle as a ratio of sides.

theorem Orientation.cos_oangle_add_right_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle x (x + y)).cos * x + y = x

The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the adjacent side.

theorem Orientation.cos_oangle_add_left_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x + y) y).cos * x + y = y

The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the adjacent side.

theorem Orientation.sin_oangle_add_right_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle x (x + y)).sin * x + y = y

The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the opposite side.

theorem Orientation.sin_oangle_add_left_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x + y) y).sin * x + y = x

The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the opposite side.

theorem Orientation.tan_oangle_add_right_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle x (x + y)).tan * x = y

The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals the opposite side.

theorem Orientation.tan_oangle_add_left_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x + y) y).tan * y = x

The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals the opposite side.

theorem Orientation.norm_div_cos_oangle_add_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
x / (o.oangle x (x + y)).cos = x + y

A side of a right-angled triangle divided by the cosine of the adjacent angle equals the hypotenuse.

theorem Orientation.norm_div_cos_oangle_add_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
y / (o.oangle (x + y) y).cos = x + y

A side of a right-angled triangle divided by the cosine of the adjacent angle equals the hypotenuse.

theorem Orientation.norm_div_sin_oangle_add_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
y / (o.oangle x (x + y)).sin = x + y

A side of a right-angled triangle divided by the sine of the opposite angle equals the hypotenuse.

theorem Orientation.norm_div_sin_oangle_add_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
x / (o.oangle (x + y) y).sin = x + y

A side of a right-angled triangle divided by the sine of the opposite angle equals the hypotenuse.

theorem Orientation.norm_div_tan_oangle_add_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
y / (o.oangle x (x + y)).tan = x

A side of a right-angled triangle divided by the tangent of the opposite angle equals the adjacent side.

theorem Orientation.norm_div_tan_oangle_add_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
x / (o.oangle (x + y) y).tan = y

A side of a right-angled triangle divided by the tangent of the opposite angle equals the adjacent side.

theorem Orientation.oangle_sub_right_eq_arccos_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle y (y - x) = (y / y - x).arccos

An angle in a right-angled triangle expressed using arccos, version subtracting vectors.

theorem Orientation.oangle_sub_left_eq_arccos_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle (x - y) x = (x / x - y).arccos

An angle in a right-angled triangle expressed using arccos, version subtracting vectors.

theorem Orientation.oangle_sub_right_eq_arcsin_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle y (y - x) = (x / y - x).arcsin

An angle in a right-angled triangle expressed using arcsin, version subtracting vectors.

theorem Orientation.oangle_sub_left_eq_arcsin_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle (x - y) x = (y / x - y).arcsin

An angle in a right-angled triangle expressed using arcsin, version subtracting vectors.

theorem Orientation.oangle_sub_right_eq_arctan_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle y (y - x) = (x / y).arctan

An angle in a right-angled triangle expressed using arctan, version subtracting vectors.

theorem Orientation.oangle_sub_left_eq_arctan_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
o.oangle (x - y) x = (y / x).arctan

An angle in a right-angled triangle expressed using arctan, version subtracting vectors.

theorem Orientation.cos_oangle_sub_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle y (y - x)).cos = y / y - x

The cosine of an angle in a right-angled triangle as a ratio of sides, version subtracting vectors.

theorem Orientation.cos_oangle_sub_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x - y) x).cos = x / x - y

The cosine of an angle in a right-angled triangle as a ratio of sides, version subtracting vectors.

theorem Orientation.sin_oangle_sub_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle y (y - x)).sin = x / y - x

The sine of an angle in a right-angled triangle as a ratio of sides, version subtracting vectors.

theorem Orientation.sin_oangle_sub_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x - y) x).sin = y / x - y

The sine of an angle in a right-angled triangle as a ratio of sides, version subtracting vectors.

theorem Orientation.tan_oangle_sub_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle y (y - x)).tan = x / y

The tangent of an angle in a right-angled triangle as a ratio of sides, version subtracting vectors.

theorem Orientation.tan_oangle_sub_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x - y) x).tan = y / x

The tangent of an angle in a right-angled triangle as a ratio of sides, version subtracting vectors.

theorem Orientation.cos_oangle_sub_right_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle y (y - x)).cos * y - x = y

The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the adjacent side, version subtracting vectors.

theorem Orientation.cos_oangle_sub_left_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x - y) x).cos * x - y = x

The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the adjacent side, version subtracting vectors.

theorem Orientation.sin_oangle_sub_right_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle y (y - x)).sin * y - x = x

The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the opposite side, version subtracting vectors.

theorem Orientation.sin_oangle_sub_left_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x - y) x).sin * x - y = y

The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the opposite side, version subtracting vectors.

theorem Orientation.tan_oangle_sub_right_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle y (y - x)).tan * y = x

The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals the opposite side, version subtracting vectors.

theorem Orientation.tan_oangle_sub_left_mul_norm_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
(o.oangle (x - y) x).tan * x = y

The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals the opposite side, version subtracting vectors.

theorem Orientation.norm_div_cos_oangle_sub_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
y / (o.oangle y (y - x)).cos = y - x

A side of a right-angled triangle divided by the cosine of the adjacent angle equals the hypotenuse, version subtracting vectors.

theorem Orientation.norm_div_cos_oangle_sub_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
x / (o.oangle (x - y) x).cos = x - y

A side of a right-angled triangle divided by the cosine of the adjacent angle equals the hypotenuse, version subtracting vectors.

theorem Orientation.norm_div_sin_oangle_sub_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
x / (o.oangle y (y - x)).sin = y - x

A side of a right-angled triangle divided by the sine of the opposite angle equals the hypotenuse, version subtracting vectors.

theorem Orientation.norm_div_sin_oangle_sub_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
y / (o.oangle (x - y) x).sin = x - y

A side of a right-angled triangle divided by the sine of the opposite angle equals the hypotenuse, version subtracting vectors.

theorem Orientation.norm_div_tan_oangle_sub_right_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
x / (o.oangle y (y - x)).tan = y

A side of a right-angled triangle divided by the tangent of the opposite angle equals the adjacent side, version subtracting vectors.

theorem Orientation.norm_div_tan_oangle_sub_left_of_oangle_eq_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} {y : V} (h : o.oangle x y = (Real.pi / 2)) :
y / (o.oangle (x - y) x).tan = x

A side of a right-angled triangle divided by the tangent of the opposite angle equals the adjacent side, version subtracting vectors.

theorem Orientation.oangle_add_right_smul_rotation_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} (h : x 0) (r : ) :
o.oangle x (x + r (o.rotation (Real.pi / 2)) x) = r.arctan

An angle in a right-angled triangle expressed using arctan, where one side is a multiple of a rotation of another by π / 2.

theorem Orientation.oangle_add_left_smul_rotation_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} (h : x 0) (r : ) :
o.oangle (x + r (o.rotation (Real.pi / 2)) x) (r (o.rotation (Real.pi / 2)) x) = r⁻¹.arctan

An angle in a right-angled triangle expressed using arctan, where one side is a multiple of a rotation of another by π / 2.

theorem Orientation.tan_oangle_add_right_smul_rotation_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} (h : x 0) (r : ) :
(o.oangle x (x + r (o.rotation (Real.pi / 2)) x)).tan = r

The tangent of an angle in a right-angled triangle, where one side is a multiple of a rotation of another by π / 2.

theorem Orientation.tan_oangle_add_left_smul_rotation_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} (h : x 0) (r : ) :
(o.oangle (x + r (o.rotation (Real.pi / 2)) x) (r (o.rotation (Real.pi / 2)) x)).tan = r⁻¹

The tangent of an angle in a right-angled triangle, where one side is a multiple of a rotation of another by π / 2.

theorem Orientation.oangle_sub_right_smul_rotation_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} (h : x 0) (r : ) :
o.oangle (r (o.rotation (Real.pi / 2)) x) (r (o.rotation (Real.pi / 2)) x - x) = r⁻¹.arctan

An angle in a right-angled triangle expressed using arctan, where one side is a multiple of a rotation of another by π / 2, version subtracting vectors.

theorem Orientation.oangle_sub_left_smul_rotation_pi_div_two {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] [hd2 : Fact (FiniteDimensional.finrank V = 2)] (o : Orientation V (Fin 2)) {x : V} (h : x 0) (r : ) :
o.oangle (x - r (o.rotation (Real.pi / 2)) x) x = r.arctan

An angle in a right-angled triangle expressed using arctan, where one side is a multiple of a rotation of another by π / 2, version subtracting vectors.

theorem EuclideanGeometry.oangle_right_eq_arccos_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
EuclideanGeometry.oangle p₂ p₃ p₁ = (dist p₃ p₂ / dist p₁ p₃).arccos

An angle in a right-angled triangle expressed using arccos.

theorem EuclideanGeometry.oangle_left_eq_arccos_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
EuclideanGeometry.oangle p₃ p₁ p₂ = (dist p₁ p₂ / dist p₁ p₃).arccos

An angle in a right-angled triangle expressed using arccos.

theorem EuclideanGeometry.oangle_right_eq_arcsin_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
EuclideanGeometry.oangle p₂ p₃ p₁ = (dist p₁ p₂ / dist p₁ p₃).arcsin

An angle in a right-angled triangle expressed using arcsin.

theorem EuclideanGeometry.oangle_left_eq_arcsin_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
EuclideanGeometry.oangle p₃ p₁ p₂ = (dist p₃ p₂ / dist p₁ p₃).arcsin

An angle in a right-angled triangle expressed using arcsin.

theorem EuclideanGeometry.oangle_right_eq_arctan_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
EuclideanGeometry.oangle p₂ p₃ p₁ = (dist p₁ p₂ / dist p₃ p₂).arctan

An angle in a right-angled triangle expressed using arctan.

theorem EuclideanGeometry.oangle_left_eq_arctan_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
EuclideanGeometry.oangle p₃ p₁ p₂ = (dist p₃ p₂ / dist p₁ p₂).arctan

An angle in a right-angled triangle expressed using arctan.

theorem EuclideanGeometry.cos_oangle_right_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₂ p₃ p₁).cos = dist p₃ p₂ / dist p₁ p₃

The cosine of an angle in a right-angled triangle as a ratio of sides.

theorem EuclideanGeometry.cos_oangle_left_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₃ p₁ p₂).cos = dist p₁ p₂ / dist p₁ p₃

The cosine of an angle in a right-angled triangle as a ratio of sides.

theorem EuclideanGeometry.sin_oangle_right_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₂ p₃ p₁).sin = dist p₁ p₂ / dist p₁ p₃

The sine of an angle in a right-angled triangle as a ratio of sides.

theorem EuclideanGeometry.sin_oangle_left_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₃ p₁ p₂).sin = dist p₃ p₂ / dist p₁ p₃

The sine of an angle in a right-angled triangle as a ratio of sides.

theorem EuclideanGeometry.tan_oangle_right_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₂ p₃ p₁).tan = dist p₁ p₂ / dist p₃ p₂

The tangent of an angle in a right-angled triangle as a ratio of sides.

theorem EuclideanGeometry.tan_oangle_left_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₃ p₁ p₂).tan = dist p₃ p₂ / dist p₁ p₂

The tangent of an angle in a right-angled triangle as a ratio of sides.

theorem EuclideanGeometry.cos_oangle_right_mul_dist_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₂ p₃ p₁).cos * dist p₁ p₃ = dist p₃ p₂

The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the adjacent side.

theorem EuclideanGeometry.cos_oangle_left_mul_dist_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₃ p₁ p₂).cos * dist p₁ p₃ = dist p₁ p₂

The cosine of an angle in a right-angled triangle multiplied by the hypotenuse equals the adjacent side.

theorem EuclideanGeometry.sin_oangle_right_mul_dist_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₂ p₃ p₁).sin * dist p₁ p₃ = dist p₁ p₂

The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the opposite side.

theorem EuclideanGeometry.sin_oangle_left_mul_dist_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₃ p₁ p₂).sin * dist p₁ p₃ = dist p₃ p₂

The sine of an angle in a right-angled triangle multiplied by the hypotenuse equals the opposite side.

theorem EuclideanGeometry.tan_oangle_right_mul_dist_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₂ p₃ p₁).tan * dist p₃ p₂ = dist p₁ p₂

The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals the opposite side.

theorem EuclideanGeometry.tan_oangle_left_mul_dist_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
(EuclideanGeometry.oangle p₃ p₁ p₂).tan * dist p₁ p₂ = dist p₃ p₂

The tangent of an angle in a right-angled triangle multiplied by the adjacent side equals the opposite side.

theorem EuclideanGeometry.dist_div_cos_oangle_right_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
dist p₃ p₂ / (EuclideanGeometry.oangle p₂ p₃ p₁).cos = dist p₁ p₃

A side of a right-angled triangle divided by the cosine of the adjacent angle equals the hypotenuse.

theorem EuclideanGeometry.dist_div_cos_oangle_left_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
dist p₁ p₂ / (EuclideanGeometry.oangle p₃ p₁ p₂).cos = dist p₁ p₃

A side of a right-angled triangle divided by the cosine of the adjacent angle equals the hypotenuse.

theorem EuclideanGeometry.dist_div_sin_oangle_right_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
dist p₁ p₂ / (EuclideanGeometry.oangle p₂ p₃ p₁).sin = dist p₁ p₃

A side of a right-angled triangle divided by the sine of the opposite angle equals the hypotenuse.

theorem EuclideanGeometry.dist_div_sin_oangle_left_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
dist p₃ p₂ / (EuclideanGeometry.oangle p₃ p₁ p₂).sin = dist p₁ p₃

A side of a right-angled triangle divided by the sine of the opposite angle equals the hypotenuse.

theorem EuclideanGeometry.dist_div_tan_oangle_right_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
dist p₁ p₂ / (EuclideanGeometry.oangle p₂ p₃ p₁).tan = dist p₃ p₂

A side of a right-angled triangle divided by the tangent of the opposite angle equals the adjacent side.

theorem EuclideanGeometry.dist_div_tan_oangle_left_of_oangle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [hd2 : Fact (FiniteDimensional.finrank V = 2)] [Module.Oriented V (Fin 2)] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.oangle p₁ p₂ p₃ = (Real.pi / 2)) :
dist p₃ p₂ / (EuclideanGeometry.oangle p₃ p₁ p₂).tan = dist p₁ p₂

A side of a right-angled triangle divided by the tangent of the opposite angle equals the adjacent side.