The inequality p⁻¹ + q⁻¹ + r⁻¹ > 1
#
In this file we classify solutions to the inequality
(p⁻¹ + q⁻¹ + r⁻¹ : ℚ) > 1
, for positive natural numbers p
, q
, and r
.
The solutions are exactly of the form.
This inequality shows up in Lie theory, in the classification of Dynkin diagrams, root systems, and semisimple Lie algebras.
Main declarations #
A r := {1,1,r}
is a Multiset ℕ+
that is a solution to the inequality
(p⁻¹ + q⁻¹ + r⁻¹ : ℚ) > 1
.
These solutions are related to the Dynkin diagrams $A_r$.
Equations
- ADEInequality.A r = ADEInequality.A' 1 r
Instances For
sum_inv pqr
for a pqr : Multiset ℕ+
is the sum of the inverses
of the elements of pqr
, as rational number.
The intended argument is a multiset {p,q,r}
of cardinality 3
.
Equations
- ADEInequality.sumInv pqr = (Multiset.map (fun (x : ℕ+) => (↑↑x)⁻¹) pqr).sum
Instances For
theorem
ADEInequality.admissible_of_one_lt_sumInv_aux
{pqr : List ℕ+}
:
List.Sorted (fun (x1 x2 : ℕ+) => x1 ≤ x2) pqr → pqr.length = 3 → 1 < sumInv ↑pqr → Admissible ↑pqr
A multiset {p,q,r}
of positive natural numbers
is a solution to (p⁻¹ + q⁻¹ + r⁻¹ : ℚ) > 1
if and only if
it is admissible
which means it is one of: