Documentation

Mathlib.Analysis.Normed.Affine.Simplex

Simplices in torsors over normed spaces. #

This file defines properties of simplices in a NormedAddTorsor.

Main definitions #

def Affine.Simplex.Scalene {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {n : } (s : Simplex R P n) :

A simplex is scalene if all the edge lengths are different.

Equations
Instances For
    theorem Affine.Simplex.Scalene.dist_ne {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {n : } {s : Simplex R P n} (hs : s.Scalene) {i₁ i₂ i₃ i₄ : Fin (n + 1)} (h₁₂ : i₁ i₂) (h₃₄ : i₃ i₄) (h₁₂₃₄ : ¬(i₁ = i₃ i₂ = i₄)) (h₁₂₄₃ : ¬(i₁ = i₄ i₂ = i₃)) :
    dist (s.points i₁) (s.points i₂) dist (s.points i₃) (s.points i₄)
    @[simp]
    theorem Affine.Simplex.scalene_reindex_iff {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {m n : } {s : Simplex R P m} (e : Fin (m + 1) Fin (n + 1)) :
    def Affine.Simplex.Equilateral {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {n : } (s : Simplex R P n) :

    A simplex is equilateral if all the edge lengths are equal.

    Equations
    Instances For
      theorem Affine.Simplex.Equilateral.dist_eq {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {n : } {s : Simplex R P n} (he : s.Equilateral) {i₁ i₂ i₃ i₄ : Fin (n + 1)} (h₁₂ : i₁ i₂) (h₃₄ : i₃ i₄) :
      dist (s.points i₁) (s.points i₂) = dist (s.points i₃) (s.points i₄)
      @[simp]
      theorem Affine.Simplex.equilateral_reindex_iff {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {m n : } {s : Simplex R P m} (e : Fin (m + 1) Fin (n + 1)) :
      def Affine.Simplex.Regular {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {n : } (s : Simplex R P n) :

      A simplex is regular if it is equivalent under an isometry to any reindexing.

      Equations
      Instances For
        @[simp]
        theorem Affine.Simplex.regular_reindex_iff {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {m n : } {s : Simplex R P m} (e : Fin (m + 1) Fin (n + 1)) :
        theorem Affine.Simplex.Regular.equilateral {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {n : } {s : Simplex R P n} (hr : s.Regular) :
        theorem Affine.Triangle.equilateral_iff_dist_eq_and_dist_eq {R : Type u_1} {V : Type u_2} {P : Type u_3} [Ring R] [SeminormedAddCommGroup V] [PseudoMetricSpace P] [Module R V] [NormedAddTorsor V P] {t : Triangle R P} {i₁ i₂ i₃ : Fin 3} (h₁₂ : i₁ i₂) (h₁₃ : i₁ i₃) (h₂₃ : i₂ i₃) :
        Simplex.Equilateral t dist (t.points i₁) (t.points i₂) = dist (t.points i₁) (t.points i₃) dist (t.points i₁) (t.points i₂) = dist (t.points i₂) (t.points i₃)