mathlib3 documentation

mathlib-archive / imo.imo1962_q1

IMO 1962 Q1 #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

Find the smallest natural number $n$ which has the following properties:

(a) Its decimal representation has 6 as the last digit.

(b) If the last digit 6 is erased and placed in front of the remaining digits, the resulting number is four times as large as the original number $n$.

Since Lean does not explicitly express problems of the form "find the smallest number satisfying X", we define the problem as a predicate, and then prove a particular number is the smallest member of a set satisfying it.

def imo1962_q1.problem_predicate (n : ) :
Prop
Equations

First, it's inconvenient to work with digits, so let's simplify them out of the problem.

@[reducible]
def imo1962_q1.problem_predicate' (c n : ) :
Prop

Now we can eliminate possibilities for (digits 10 c).length until we get to the one that works.

theorem imo1962_q1.helper_5_digit {c : } (h : 6 * 10 ^ 5 + c = 4 * (10 * c + 6)) :
c = 15384

Putting this inline causes a deep recursion error, so we separate it out.

theorem imo1962_q1.case_5_digit {c n : } (h1 : (10.digits c).length = 5) (h2 : imo1962_q1.problem_predicate' c n) :
c = 15384
theorem imo1962_q1.case_more_digits {c n : } (h1 : (10.digits c).length 6) (h2 : imo1962_q1.problem_predicate' c n) :
n 153846

linarith fails on numbers this large, so this lemma spells out some of the arithmetic that normally would be automated.

Now we combine these cases to show that 153846 is the smallest solution.