author | wenzelm |
Sat, 05 Jan 2019 17:24:33 +0100 | |
changeset 69597 | ff784d5a5bfb |
parent 66453 | cc19f7ca2ed6 |
permissions | -rw-r--r-- |
33026 | 1 |
(* Title: HOL/Isar_Examples/Knaster_Tarski.thy |
61932 | 2 |
Author: Makarius |
6882 | 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 |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
63583
diff
changeset
|
10 |
imports Main "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 |
|
61932 | 16 |
text \<open> |
17 |
According to the textbook @{cite \<open>pages 93--94\<close> "davey-priestley"}, the |
|
18 |
Knaster-Tarski fixpoint theorem is as follows.\<^footnote>\<open>We have dualized the |
|
19 |
argument, and tuned the notation a little bit.\<close> |
|
7153 | 20 |
|
61541 | 21 |
\<^bold>\<open>The Knaster-Tarski Fixpoint Theorem.\<close> Let \<open>L\<close> be a complete lattice and |
61932 | 22 |
\<open>f: L \<rightarrow> L\<close> an order-preserving map. Then \<open>\<Sqinter>{x \<in> L | f(x) \<le> x}\<close> is a fixpoint |
23 |
of \<open>f\<close>. |
|
61541 | 24 |
|
61932 | 25 |
\<^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 have |
26 |
\<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>, whence |
|
27 |
\<open>f(a) \<le> a\<close>. We now use this inequality to prove the reverse one (!) and |
|
28 |
thereby complete the proof that \<open>a\<close> is a fixpoint. Since \<open>f\<close> is |
|
61541 | 29 |
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 | 30 |
|
7761 | 31 |
|
58614 | 32 |
subsection \<open>Formal versions\<close> |
7761 | 33 |
|
61932 | 34 |
text \<open> |
35 |
The Isar proof below closely follows the original presentation. Virtually |
|
36 |
all of the prose narration has been rephrased in terms of formal Isar |
|
37 |
language elements. Just as many textbook-style proofs, there is a strong |
|
38 |
bias towards forward proof, and several bends in the course of reasoning. |
|
39 |
\<close> |
|
6882 | 40 |
|
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
|
41 |
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
|
42 |
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
|
43 |
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
|
44 |
shows "\<exists>a. f a = a" |
10007 | 45 |
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
|
46 |
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
|
47 |
let ?a = "\<Sqinter>?H" |
10007 | 48 |
show "f ?a = ?a" |
49 |
proof - |
|
50 |
{ |
|
51 |
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
|
52 |
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
|
53 |
then have "?a \<le> x" by (rule Inf_lower) |
58614 | 54 |
with \<open>mono f\<close> have "f ?a \<le> f x" .. |
55 |
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
|
56 |
finally have "f ?a \<le> x" . |
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 |
then have "f ?a \<le> ?a" by (rule Inf_greatest) |
10007 | 59 |
{ |
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
|
60 |
also presume "\<dots> \<le> f ?a" |
10007 | 61 |
finally (order_antisym) show ?thesis . |
62 |
} |
|
58614 | 63 |
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
|
64 |
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
|
65 |
then show "?a \<le> f ?a" by (rule Inf_lower) |
10007 | 66 |
qed |
67 |
qed |
|
6898 | 68 |
|
61932 | 69 |
text \<open> |
70 |
Above we have used several advanced Isar language elements, such as explicit |
|
71 |
block structure and weak assumptions. Thus we have mimicked the particular |
|
72 |
way of reasoning of the original text. |
|
7818 | 73 |
|
61541 | 74 |
In the subsequent version the order of reasoning is changed to achieve |
75 |
structured top-down decomposition of the problem at the outer level, while |
|
76 |
only the inner steps of reasoning are done in a forward manner. We are |
|
77 |
certainly more at ease here, requiring only the most basic features of the |
|
61932 | 78 |
Isar language. |
79 |
\<close> |
|
7818 | 80 |
|
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
|
81 |
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
|
82 |
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
|
83 |
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
|
84 |
shows "\<exists>a. f a = a" |
10007 | 85 |
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
|
86 |
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
|
87 |
let ?a = "\<Sqinter>?H" |
10007 | 88 |
show "f ?a = ?a" |
89 |
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
|
90 |
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
|
91 |
proof (rule Inf_greatest) |
10007 | 92 |
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
|
93 |
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
|
94 |
then have "?a \<le> x" by (rule Inf_lower) |
58614 | 95 |
with \<open>mono f\<close> have "f ?a \<le> f x" .. |
96 |
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
|
97 |
finally show "f ?a \<le> x" . |
10007 | 98 |
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
|
99 |
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
|
100 |
proof (rule Inf_lower) |
58614 | 101 |
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
|
102 |
then show "f ?a \<in> ?H" .. |
10007 | 103 |
qed |
104 |
qed |
|
105 |
qed |
|
7818 | 106 |
|
10007 | 107 |
end |