| 8020 |      1 | 
 | 
| 10007 |      2 | header {* An old chestnut *}
 | 
| 8020 |      3 | 
 | 
| 16417 |      4 | theory Puzzle imports Main begin
 | 
| 8020 |      5 | 
 | 
|  |      6 | text_raw {*
 | 
| 18193 |      7 |   \footnote{A question from ``Bundeswettbewerb Mathematik''.  Original
 | 
|  |      8 |   pen-and-paper proof due to Herbert Ehler; Isabelle tactic script by
 | 
|  |      9 |   Tobias Nipkow.}
 | 
|  |     10 | *}
 | 
| 8020 |     11 | 
 | 
| 18193 |     12 | text {*
 | 
|  |     13 |   \textbf{Problem.}  Given some function $f\colon \Nat \to \Nat$ such
 | 
|  |     14 |   that $f \ap (f \ap n) < f \ap (\idt{Suc} \ap n)$ for all $n$.
 | 
|  |     15 |   Demonstrate that $f$ is the identity.
 | 
| 10007 |     16 | *}
 | 
| 8020 |     17 | 
 | 
| 18191 |     18 | theorem
 | 
|  |     19 |   assumes f_ax: "\<And>n. f (f n) < f (Suc n)"
 | 
|  |     20 |   shows "f n = n"
 | 
| 10436 |     21 | proof (rule order_antisym)
 | 
| 18191 |     22 |   {
 | 
|  |     23 |     fix n show "n \<le> f n"
 | 
| 20503 |     24 |     proof (induct k \<equiv> "f n" arbitrary: n rule: less_induct)
 | 
| 18191 |     25 |       case (less k n)
 | 
|  |     26 |       then have hyp: "\<And>m. f m < f n \<Longrightarrow> m \<le> f m" by (simp only:)
 | 
|  |     27 |       show "n \<le> f n"
 | 
| 10436 |     28 |       proof (cases n)
 | 
| 18191 |     29 | 	case (Suc m)
 | 
|  |     30 | 	from f_ax have "f (f m) < f n" by (simp only: Suc)
 | 
|  |     31 | 	with hyp have "f m \<le> f (f m)" .
 | 
|  |     32 | 	also from f_ax have "\<dots> < f n" by (simp only: Suc)
 | 
|  |     33 | 	finally have "f m < f n" .
 | 
|  |     34 | 	with hyp have "m \<le> f m" .
 | 
|  |     35 | 	also note `\<dots> < f n`
 | 
|  |     36 | 	finally have "m < f n" .
 | 
|  |     37 | 	then have "n \<le> f n" by (simp only: Suc)
 | 
|  |     38 | 	then show ?thesis .
 | 
| 10007 |     39 |       next
 | 
| 18191 |     40 | 	case 0
 | 
|  |     41 | 	then show ?thesis by simp
 | 
| 10007 |     42 |       qed
 | 
|  |     43 |     qed
 | 
| 18191 |     44 |   } note ge = this
 | 
| 8020 |     45 | 
 | 
| 10436 |     46 |   {
 | 
| 18191 |     47 |     fix m n :: nat
 | 
|  |     48 |     assume "m \<le> n"
 | 
|  |     49 |     then have "f m \<le> f n"
 | 
| 10007 |     50 |     proof (induct n)
 | 
| 18191 |     51 |       case 0
 | 
|  |     52 |       then have "m = 0" by simp
 | 
|  |     53 |       then show ?case by simp
 | 
| 10436 |     54 |     next
 | 
| 18191 |     55 |       case (Suc n)
 | 
|  |     56 |       from Suc.prems show "f m \<le> f (Suc n)"
 | 
| 10436 |     57 |       proof (rule le_SucE)
 | 
| 18191 |     58 |         assume "m \<le> n"
 | 
|  |     59 |         with Suc.hyps have "f m \<le> f n" .
 | 
|  |     60 |         also from ge f_ax have "\<dots> < f (Suc n)"
 | 
| 10436 |     61 |           by (rule le_less_trans)
 | 
|  |     62 |         finally show ?thesis by simp
 | 
|  |     63 |       next
 | 
|  |     64 |         assume "m = Suc n"
 | 
| 18191 |     65 |         then show ?thesis by simp
 | 
| 10007 |     66 |       qed
 | 
|  |     67 |     qed
 | 
| 10436 |     68 |   } note mono = this
 | 
| 8020 |     69 | 
 | 
| 18191 |     70 |   show "f n \<le> n"
 | 
| 10436 |     71 |   proof -
 | 
| 18191 |     72 |     have "\<not> n < f n"
 | 
| 10007 |     73 |     proof
 | 
|  |     74 |       assume "n < f n"
 | 
| 18191 |     75 |       then have "Suc n \<le> f n" by simp
 | 
|  |     76 |       then have "f (Suc n) \<le> f (f n)" by (rule mono)
 | 
|  |     77 |       also have "\<dots> < f (Suc n)" by (rule f_ax)
 | 
|  |     78 |       finally have "\<dots> < \<dots>" . then show False ..
 | 
| 10007 |     79 |     qed
 | 
| 18191 |     80 |     then show ?thesis by simp
 | 
| 10007 |     81 |   qed
 | 
|  |     82 | qed
 | 
| 8020 |     83 | 
 | 
| 10007 |     84 | end
 |