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