author | wenzelm |
Tue, 10 Nov 2015 19:56:51 +0100 | |
changeset 61616 | abbecf4e6601 |
parent 61541 | 846c72206207 |
child 61932 | 2e48182cc82c |
permissions | -rw-r--r-- |
33026 | 1 |
(* Title: HOL/Isar_Examples/Knaster_Tarski.thy |
6882 | 2 |
Author: Markus Wenzel, TU Muenchen |
3 |
||
4 |
Typical textbook proof example. |
|
5 |
*) |
|
6 |
||
58882 | 7 |
section \<open>Textbook-style reasoning: the Knaster-Tarski Theorem\<close> |
6882 | 8 |
|
31758 | 9 |
theory Knaster_Tarski |
41413
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
wenzelm
parents:
37671
diff
changeset
|
10 |
imports Main "~~/src/HOL/Library/Lattice_Syntax" |
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
11 |
begin |
6882 | 12 |
|
7761 | 13 |
|
58614 | 14 |
subsection \<open>Prose version\<close> |
7761 | 15 |
|
58614 | 16 |
text \<open>According to the textbook @{cite \<open>pages 93--94\<close> "davey-priestley"}, |
61541 | 17 |
the Knaster-Tarski fixpoint theorem is as follows.\footnote{We have |
18 |
dualized the argument, and tuned the notation a little bit.} |
|
7153 | 19 |
|
61541 | 20 |
\<^bold>\<open>The Knaster-Tarski Fixpoint Theorem.\<close> Let \<open>L\<close> be a complete lattice and |
21 |
\<open>f: L \<rightarrow> L\<close> an order-preserving map. Then \<open>\<Sqinter>{x \<in> L | f(x) \<le> x}\<close> is a |
|
22 |
fixpoint of \<open>f\<close>. |
|
23 |
||
24 |
\<^bold>\<open>Proof.\<close> Let \<open>H = {x \<in> L | f(x) \<le> x}\<close> and \<open>a = \<Sqinter>H\<close>. For all \<open>x \<in> H\<close> we |
|
25 |
have \<open>a \<le> x\<close>, so \<open>f(a) \<le> f(x) \<le> x\<close>. Thus \<open>f(a)\<close> is a lower bound of \<open>H\<close>, |
|
26 |
whence \<open>f(a) \<le> a\<close>. We now use this inequality to prove the reverse one (!) |
|
27 |
and thereby complete the proof that \<open>a\<close> is a fixpoint. Since \<open>f\<close> is |
|
28 |
order-preserving, \<open>f(f(a)) \<le> f(a)\<close>. This says \<open>f(a) \<in> H\<close>, so \<open>a \<le> f(a)\<close>.\<close> |
|
6883 | 29 |
|
7761 | 30 |
|
58614 | 31 |
subsection \<open>Formal versions\<close> |
7761 | 32 |
|
61541 | 33 |
text \<open>The Isar proof below closely follows the original presentation. |
34 |
Virtually all of the prose narration has been rephrased in terms of formal |
|
35 |
Isar language elements. Just as many textbook-style proofs, there is a |
|
36 |
strong bias towards forward proof, and several bends in the course of |
|
37 |
reasoning.\<close> |
|
6882 | 38 |
|
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
39 |
theorem Knaster_Tarski: |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
40 |
fixes f :: "'a::complete_lattice \<Rightarrow> 'a" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
41 |
assumes "mono f" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
42 |
shows "\<exists>a. f a = a" |
10007 | 43 |
proof |
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
44 |
let ?H = "{u. f u \<le> u}" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
45 |
let ?a = "\<Sqinter>?H" |
10007 | 46 |
show "f ?a = ?a" |
47 |
proof - |
|
48 |
{ |
|
49 |
fix x |
|
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
50 |
assume "x \<in> ?H" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
51 |
then have "?a \<le> x" by (rule Inf_lower) |
58614 | 52 |
with \<open>mono f\<close> have "f ?a \<le> f x" .. |
53 |
also from \<open>x \<in> ?H\<close> have "\<dots> \<le> x" .. |
|
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
54 |
finally have "f ?a \<le> x" . |
10007 | 55 |
} |
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
56 |
then have "f ?a \<le> ?a" by (rule Inf_greatest) |
10007 | 57 |
{ |
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
58 |
also presume "\<dots> \<le> f ?a" |
10007 | 59 |
finally (order_antisym) show ?thesis . |
60 |
} |
|
58614 | 61 |
from \<open>mono f\<close> and \<open>f ?a \<le> ?a\<close> have "f (f ?a) \<le> f ?a" .. |
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
62 |
then have "f ?a \<in> ?H" .. |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
63 |
then show "?a \<le> f ?a" by (rule Inf_lower) |
10007 | 64 |
qed |
65 |
qed |
|
6898 | 66 |
|
61541 | 67 |
text \<open>Above we have used several advanced Isar language elements, such as |
68 |
explicit block structure and weak assumptions. Thus we have mimicked the |
|
69 |
particular way of reasoning of the original text. |
|
7818 | 70 |
|
61541 | 71 |
In the subsequent version the order of reasoning is changed to achieve |
72 |
structured top-down decomposition of the problem at the outer level, while |
|
73 |
only the inner steps of reasoning are done in a forward manner. We are |
|
74 |
certainly more at ease here, requiring only the most basic features of the |
|
75 |
Isar language.\<close> |
|
7818 | 76 |
|
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
77 |
theorem Knaster_Tarski': |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
78 |
fixes f :: "'a::complete_lattice \<Rightarrow> 'a" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
79 |
assumes "mono f" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
80 |
shows "\<exists>a. f a = a" |
10007 | 81 |
proof |
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
82 |
let ?H = "{u. f u \<le> u}" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
83 |
let ?a = "\<Sqinter>?H" |
10007 | 84 |
show "f ?a = ?a" |
85 |
proof (rule order_antisym) |
|
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
86 |
show "f ?a \<le> ?a" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
87 |
proof (rule Inf_greatest) |
10007 | 88 |
fix x |
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
89 |
assume "x \<in> ?H" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
90 |
then have "?a \<le> x" by (rule Inf_lower) |
58614 | 91 |
with \<open>mono f\<close> have "f ?a \<le> f x" .. |
92 |
also from \<open>x \<in> ?H\<close> have "\<dots> \<le> x" .. |
|
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
93 |
finally show "f ?a \<le> x" . |
10007 | 94 |
qed |
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
95 |
show "?a \<le> f ?a" |
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
96 |
proof (rule Inf_lower) |
58614 | 97 |
from \<open>mono f\<close> and \<open>f ?a \<le> ?a\<close> have "f (f ?a) \<le> f ?a" .. |
30816
4de62c902f9a
replaced 'a set by abstract 'a::complete_lattice, and recover plain reasoning instead of adhoc automation (by simp);
wenzelm
parents:
26812
diff
changeset
|
98 |
then show "f ?a \<in> ?H" .. |
10007 | 99 |
qed |
100 |
qed |
|
101 |
qed |
|
7818 | 102 |
|
10007 | 103 |
end |