Documentation

Archive.Imo.Imo1975Q1

IMO 1975 Q1 #

Let x₁, x₂, ... , xₙ and y₁, y₂, ... , yₙ be two sequences of real numbers, such that x₁ ≥ x₂ ≥ ... ≥ xₙ and y₁ ≥ y₂ ≥ ... ≥ yₙ. Prove that if z₁, z₂, ... , zₙ is any permutation of y₁, y₂, ... , yₙ, then ∑ (xᵢ - yᵢ)^2 ≤ ∑ (xᵢ - zᵢ)^2

Solution #

Firstly, we expand the squares within both sums and distribute into separate finite sums. Then, noting that ∑ yᵢ ^ 2 = ∑ zᵢ ^ 2, it remains to prove that ∑ xᵢ * zᵢ ≤ ∑ xᵢ * yᵢ, which is true by the Rearrangement Inequality

theorem imo1975_q1 (n : ) (σ : Equiv.Perm ) (hσ : {x : | σ x x} (Finset.Icc 1 n)) (x : ) (y : ) (hx : AntitoneOn x (Finset.Icc 1 n)) (hy : AntitoneOn y (Finset.Icc 1 n)) :
((Finset.Icc 1 n).sum fun (i : ) => (x i - y i) ^ 2) (Finset.Icc 1 n).sum fun (i : ) => (x i - y (σ i)) ^ 2