Documentation

Archive.Imo.Imo2021Q1

IMO 2021 Q1 #

Let n≥100 be an integer. Ivan writes the numbers n, n+1,..., 2n each on different cards. He then shuffles these n+1 cards, and divides them into two piles. Prove that at least one of the piles contains two cards such that the sum of their numbers is a perfect square.

Solution #

We show there exists a triplet a, b, c ∈ [n , 2n] with a < b < c and each of the sums (a + b), (b + c), (a + c) being a perfect square. Specifically, we consider the linear system of equations

a + b = (2 * l - 1) ^ 2
a + c = (2 * l) ^ 2
b + c = (2 * l + 1) ^ 2

which can be solved to give

a = 2 * l ^ 2 - 4 * l
b = 2 * l ^ 2 + 1
c = 2 * l ^ 2 + 4 * l

Therefore, it is enough to show that there exists a natural number l such that n ≤ 2 * l ^ 2 - 4 * l and 2 * l ^ 2 + 4 * l ≤ 2 * n for n ≥ 100.

Then, by the Pigeonhole principle, at least two numbers in the triplet must lie in the same pile, which finishes the proof.

theorem Imo2021Q1.exists_numbers_in_interval (n : ) (hn : 100 n) :
l, n + 4 * l 2 * l ^ 2 2 * l ^ 2 + 4 * l 2 * n
theorem Imo2021Q1.exists_triplet_summing_to_squares (n : ) (hn : 100 n) :
a b c, n a a < b b < c c 2 * n (k, a + b = k ^ 2) (l, c + a = l ^ 2) m, b + c = m ^ 2
theorem Imo2021Q1.exists_finset_3_le_card_with_pairs_summing_to_squares (n : ) (hn : 100 n) :
B, 2 * 1 + 1 Finset.card B (∀ (a : ), a B∀ (b : ), b Ba bk, a + b = k ^ 2) ∀ (c : ), c Bn c c 2 * n
theorem imo2021_q1 (n : ) :
100 n∀ (A : Finset ), A Finset.Icc n (2 * n)(a x b x, a b k, a + b = k ^ 2) a x b x, a b k, a + b = k ^ 2