Documentation

Mathlib.Geometry.Euclidean.Triangle

Triangles #

This file proves basic geometrical results about distances and angles in (possibly degenerate) triangles in real inner product spaces and Euclidean affine spaces. More specialized results, and results developed for simplices in general rather than just for triangles, are in separate files. Definitions and results that make sense in more general affine spaces rather than just in the Euclidean case go under LinearAlgebra.AffineSpace.

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 #

Geometrical results on triangles in real inner product spaces #

This section develops some results on (possibly degenerate) triangles in real inner product spaces, where those definitions and results can most conveniently be developed in terms of vectors and then used to deduce corresponding results for Euclidean affine spaces.

Pons asinorum, vector angle form.

Converse of pons asinorum, vector angle form.

The cosine of the sum of two angles in a possibly degenerate triangle (where two given sides are nonzero), vector angle form.

The sine of the sum of two angles in a possibly degenerate triangle (where two given sides are nonzero), vector angle form.

theorem InnerProductGeometry.cos_angle_add_angle_sub_add_angle_sub_eq_neg_one {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] {x y : V} (hx : x 0) (hy : y 0) :
Real.cos (angle x y + angle x (x - y) + angle y (y - x)) = -1

The cosine of the sum of the angles of a possibly degenerate triangle (where two given sides are nonzero), vector angle form.

theorem InnerProductGeometry.sin_angle_add_angle_sub_add_angle_sub_eq_zero {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] {x y : V} (hx : x 0) (hy : y 0) :
Real.sin (angle x y + angle x (x - y) + angle y (y - x)) = 0

The sine of the sum of the angles of a possibly degenerate triangle (where two given sides are nonzero), vector angle form.

theorem InnerProductGeometry.angle_add_angle_sub_add_angle_sub_eq_pi {V : Type u_1} [NormedAddCommGroup V] [InnerProductSpace V] {x y : V} (hx : x 0) (hy : y 0) :
angle x y + angle x (x - y) + angle y (y - x) = Real.pi

The sum of the angles of a possibly degenerate triangle (where the two given sides are nonzero), vector angle form.

Geometrical results on triangles in Euclidean affine spaces #

This section develops some geometrical definitions and results on (possibly degenerate) triangles in Euclidean affine spaces.

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

Law of cosines (cosine rule), angle-at-point form.

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

Alias of EuclideanGeometry.dist_sq_eq_dist_sq_add_dist_sq_sub_two_mul_dist_mul_dist_mul_cos_angle.


Law of cosines (cosine rule), angle-at-point form.

theorem EuclideanGeometry.angle_eq_angle_of_dist_eq {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p1 p2 p3 : P} (h : dist p1 p2 = dist p1 p3) :
angle p1 p2 p3 = angle p1 p3 p2

Isosceles Triangle Theorem: Pons asinorum, angle-at-point form.

theorem EuclideanGeometry.dist_eq_of_angle_eq_angle_of_angle_ne_pi {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p1 p2 p3 : P} (h : angle p1 p2 p3 = angle p1 p3 p2) (hpi : angle p2 p1 p3 Real.pi) :
dist p1 p2 = dist p1 p3

Converse of pons asinorum, angle-at-point form.

theorem EuclideanGeometry.angle_add_angle_add_angle_eq_pi {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] {p1 p2 p3 : P} (h2 : p2 p1) (h3 : p3 p1) :
angle p1 p2 p3 + angle p2 p3 p1 + angle p3 p1 p2 = Real.pi

The sum of the angles of a triangle (possibly degenerate, where the given vertex is distinct from the others), angle-at-point.

theorem EuclideanGeometry.oangle_add_oangle_add_oangle_eq_pi {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] [Module.Oriented V (Fin 2)] [Fact (Module.finrank V = 2)] {p1 p2 p3 : P} (h21 : p2 p1) (h32 : p3 p2) (h13 : p1 p3) :
oangle p1 p2 p3 + oangle p2 p3 p1 + oangle p3 p1 p2 = Real.pi

The sum of the angles of a triangle (possibly degenerate, where the triangle is a line), oriented angles at point.

theorem EuclideanGeometry.dist_sq_mul_dist_add_dist_sq_mul_dist {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] (a b c p : P) (h : angle b p c = Real.pi) :
dist a b ^ 2 * dist c p + dist a c ^ 2 * dist b p = dist b c * (dist a p ^ 2 + dist b p * dist c p)

Stewart's Theorem.

Apollonius's Theorem.

theorem EuclideanGeometry.dist_mul_of_eq_angle_of_dist_mul {V : Type u_1} {P : Type u_2} [NormedAddCommGroup V] [InnerProductSpace V] [MetricSpace P] [NormedAddTorsor V P] (a b c a' b' c' : P) (r : ) (h : angle a' b' c' = angle a b c) (hab : dist a' b' = r * dist a b) (hcb : dist c' b' = r * dist c b) :
dist a' c' = r * dist a c