Documentation

Mathlib.Geometry.Euclidean.Angle.Unoriented.RightAngle

Right-angled triangles #

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

Implementation notes #

Results in this file are generally given in a form with only those non-degeneracy conditions needed for the particular result, rather than requiring affine independence of the points of a triangle unnecessarily.

References #

Pythagorean theorem, vector angle form.

Pythagorean theorem, subtracting vectors, if-and-only-if vector angle form.

Pythagorean theorem, subtracting vectors, vector angle form.

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

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

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

theorem InnerProductGeometry.angle_add_pos_of_inner_eq_zero {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] {x : V} {y : V} (h : x, y⟫_ = 0) (h0 : x = 0 y 0) :

An angle in a non-degenerate right-angled triangle is positive.

An angle in a right-angled triangle is at most π / 2.

An angle in a non-degenerate right-angled triangle is less than π / 2.

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

theorem InnerProductGeometry.sin_angle_add_of_inner_eq_zero {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] {x : V} {y : V} (h : x, y⟫_ = 0) (h0 : x 0 y 0) :

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

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

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

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

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

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

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

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

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

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

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

theorem InnerProductGeometry.angle_sub_pos_of_inner_eq_zero {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] {x : V} {y : V} (h : x, y⟫_ = 0) (h0 : x = 0 y 0) :

An angle in a non-degenerate right-angled triangle is positive, version subtracting vectors.

An angle in a right-angled triangle is at most π / 2, version subtracting vectors.

An angle in a non-degenerate right-angled triangle is less than π / 2, version subtracting vectors.

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

theorem InnerProductGeometry.sin_angle_sub_of_inner_eq_zero {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] {x : V} {y : V} (h : x, y⟫_ = 0) (h0 : x 0 y 0) :

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

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

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

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

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

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

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

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

theorem EuclideanGeometry.dist_sq_eq_dist_sq_add_dist_sq_iff_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] (p1 : P) (p2 : P) (p3 : P) :
dist p1 p3 * dist p1 p3 = dist p1 p2 * dist p1 p2 + dist p3 p2 * dist p3 p2 EuclideanGeometry.angle p1 p2 p3 = Real.pi / 2

Pythagorean theorem, if-and-only-if angle-at-point form.

theorem EuclideanGeometry.angle_eq_arccos_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) :
EuclideanGeometry.angle p₂ p₃ p₁ = Real.arccos (dist p₃ p₂ / dist p₁ p₃)

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

theorem EuclideanGeometry.angle_eq_arcsin_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) (h0 : p₁ p₂ p₃ p₂) :
EuclideanGeometry.angle p₂ p₃ p₁ = Real.arcsin (dist p₁ p₂ / dist p₁ p₃)

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

theorem EuclideanGeometry.angle_eq_arctan_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) (h0 : p₃ p₂) :
EuclideanGeometry.angle p₂ p₃ p₁ = Real.arctan (dist p₁ p₂ / dist p₃ p₂)

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

theorem EuclideanGeometry.angle_pos_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) (h0 : p₁ p₂ p₃ = p₂) :
0 < EuclideanGeometry.angle p₂ p₃ p₁

An angle in a non-degenerate right-angled triangle is positive.

theorem EuclideanGeometry.angle_le_pi_div_two_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) :

An angle in a right-angled triangle is at most π / 2.

theorem EuclideanGeometry.angle_lt_pi_div_two_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) (h0 : p₃ p₂) :

An angle in a non-degenerate right-angled triangle is less than π / 2.

theorem EuclideanGeometry.cos_angle_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) :
Real.cos (EuclideanGeometry.angle p₂ p₃ p₁) = dist p₃ p₂ / dist p₁ p₃

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

theorem EuclideanGeometry.sin_angle_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) (h0 : p₁ p₂ p₃ p₂) :
Real.sin (EuclideanGeometry.angle p₂ p₃ p₁) = dist p₁ p₂ / dist p₁ p₃

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

theorem EuclideanGeometry.tan_angle_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) :
Real.tan (EuclideanGeometry.angle p₂ p₃ p₁) = dist p₁ p₂ / dist p₃ p₂

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

theorem EuclideanGeometry.cos_angle_mul_dist_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) :
Real.cos (EuclideanGeometry.angle p₂ p₃ p₁) * 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_angle_mul_dist_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) :
Real.sin (EuclideanGeometry.angle p₂ p₃ p₁) * 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_angle_mul_dist_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) (h0 : p₁ = p₂ p₃ p₂) :
Real.tan (EuclideanGeometry.angle p₂ p₃ p₁) * 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_angle_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) (h0 : p₁ = p₂ p₃ p₂) :
dist p₃ p₂ / Real.cos (EuclideanGeometry.angle p₂ p₃ p₁) = 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_angle_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) (h0 : p₁ p₂ p₃ = p₂) :
dist p₁ p₂ / Real.sin (EuclideanGeometry.angle p₂ p₃ p₁) = 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_angle_of_angle_eq_pi_div_two {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p₁ : P} {p₂ : P} {p₃ : P} (h : EuclideanGeometry.angle p₁ p₂ p₃ = Real.pi / 2) (h0 : p₁ p₂ p₃ = p₂) :
dist p₁ p₂ / Real.tan (EuclideanGeometry.angle p₂ p₃ p₁) = dist p₃ p₂

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