author | wenzelm |
Sat, 27 Nov 2021 14:03:44 +0100 | |
changeset 74853 | 7420a7ac1a4c |
parent 69597 | ff784d5a5bfb |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
1 |
(* Title: HOL/Algebra/Order.thy |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
2 |
Author: Clemens Ballarin, started 7 November 2003 |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
3 |
Copyright: Clemens Ballarin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
4 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
5 |
Most congruence rules by Stephan Hohe. |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
6 |
With additional contributions from Alasdair Armstrong and Simon Foster. |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
7 |
*) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
8 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
9 |
theory Order |
68072
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67613
diff
changeset
|
10 |
imports |
493b818e8e10
added Johannes' generalizations Modules.thy and Vector_Spaces.thy; adapted HOL and HOL-Analysis accordingly
immler
parents:
67613
diff
changeset
|
11 |
Congruence |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
12 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
13 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
14 |
section \<open>Orders\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
15 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
16 |
subsection \<open>Partial Orders\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
17 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
18 |
record 'a gorder = "'a eq_object" + |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
19 |
le :: "['a, 'a] => bool" (infixl "\<sqsubseteq>\<index>" 50) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
20 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
21 |
abbreviation inv_gorder :: "_ \<Rightarrow> 'a gorder" where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
22 |
"inv_gorder L \<equiv> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
23 |
\<lparr> carrier = carrier L, |
67399 | 24 |
eq = (.=\<^bsub>L\<^esub>), |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
25 |
le = (\<lambda> x y. y \<sqsubseteq>\<^bsub>L \<^esub>x) \<rparr>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
26 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
27 |
lemma inv_gorder_inv: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
28 |
"inv_gorder (inv_gorder L) = L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
29 |
by simp |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
30 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
31 |
locale weak_partial_order = equivalence L for L (structure) + |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
32 |
assumes le_refl [intro, simp]: |
68004 | 33 |
"x \<in> carrier L \<Longrightarrow> x \<sqsubseteq> x" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
34 |
and weak_le_antisym [intro]: |
68004 | 35 |
"\<lbrakk>x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x .= y" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
36 |
and le_trans [trans]: |
68004 | 37 |
"\<lbrakk>x \<sqsubseteq> y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> z" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
38 |
and le_cong: |
68004 | 39 |
"\<lbrakk>x .= y; z .= w; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L; w \<in> carrier L\<rbrakk> \<Longrightarrow> |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
40 |
x \<sqsubseteq> z \<longleftrightarrow> y \<sqsubseteq> w" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
41 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
42 |
definition |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
43 |
lless :: "[_, 'a, 'a] => bool" (infixl "\<sqsubset>\<index>" 50) |
67091 | 44 |
where "x \<sqsubset>\<^bsub>L\<^esub> y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> y \<and> x .\<noteq>\<^bsub>L\<^esub> y" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
45 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
46 |
subsubsection \<open>The order relation\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
47 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
48 |
context weak_partial_order |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
49 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
50 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
51 |
lemma le_cong_l [intro, trans]: |
68004 | 52 |
"\<lbrakk>x .= y; y \<sqsubseteq> z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> z" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
53 |
by (auto intro: le_cong [THEN iffD2]) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
54 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
55 |
lemma le_cong_r [intro, trans]: |
68004 | 56 |
"\<lbrakk>x \<sqsubseteq> y; y .= z; x \<in> carrier L; y \<in> carrier L; z \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> z" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
57 |
by (auto intro: le_cong [THEN iffD1]) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
58 |
|
68004 | 59 |
lemma weak_refl [intro, simp]: "\<lbrakk>x .= y; x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
60 |
by (simp add: le_cong_l) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
61 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
62 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
63 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
64 |
lemma weak_llessI: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
65 |
fixes R (structure) |
67091 | 66 |
assumes "x \<sqsubseteq> y" and "\<not>(x .= y)" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
67 |
shows "x \<sqsubset> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
68 |
using assms unfolding lless_def by simp |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
69 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
70 |
lemma lless_imp_le: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
71 |
fixes R (structure) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
72 |
assumes "x \<sqsubset> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
73 |
shows "x \<sqsubseteq> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
74 |
using assms unfolding lless_def by simp |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
75 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
76 |
lemma weak_lless_imp_not_eq: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
77 |
fixes R (structure) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
78 |
assumes "x \<sqsubset> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
79 |
shows "\<not> (x .= y)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
80 |
using assms unfolding lless_def by simp |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
81 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
82 |
lemma weak_llessE: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
83 |
fixes R (structure) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
84 |
assumes p: "x \<sqsubset> y" and e: "\<lbrakk>x \<sqsubseteq> y; \<not> (x .= y)\<rbrakk> \<Longrightarrow> P" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
85 |
shows "P" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
86 |
using p by (blast dest: lless_imp_le weak_lless_imp_not_eq e) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
87 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
88 |
lemma (in weak_partial_order) lless_cong_l [trans]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
89 |
assumes xx': "x .= x'" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
90 |
and xy: "x' \<sqsubset> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
91 |
and carr: "x \<in> carrier L" "x' \<in> carrier L" "y \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
92 |
shows "x \<sqsubset> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
93 |
using assms unfolding lless_def by (auto intro: trans sym) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
94 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
95 |
lemma (in weak_partial_order) lless_cong_r [trans]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
96 |
assumes xy: "x \<sqsubset> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
97 |
and yy': "y .= y'" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
98 |
and carr: "x \<in> carrier L" "y \<in> carrier L" "y' \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
99 |
shows "x \<sqsubset> y'" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
100 |
using assms unfolding lless_def by (auto intro: trans sym) (*slow*) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
101 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
102 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
103 |
lemma (in weak_partial_order) lless_antisym: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
104 |
assumes "a \<in> carrier L" "b \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
105 |
and "a \<sqsubset> b" "b \<sqsubset> a" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
106 |
shows "P" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
107 |
using assms |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
108 |
by (elim weak_llessE) auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
109 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
110 |
lemma (in weak_partial_order) lless_trans [trans]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
111 |
assumes "a \<sqsubset> b" "b \<sqsubset> c" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
112 |
and carr[simp]: "a \<in> carrier L" "b \<in> carrier L" "c \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
113 |
shows "a \<sqsubset> c" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
114 |
using assms unfolding lless_def by (blast dest: le_trans intro: sym) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
115 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
116 |
lemma weak_partial_order_subset: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
117 |
assumes "weak_partial_order L" "A \<subseteq> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
118 |
shows "weak_partial_order (L\<lparr> carrier := A \<rparr>)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
119 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
120 |
interpret L: weak_partial_order L |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
121 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
122 |
interpret equivalence "(L\<lparr> carrier := A \<rparr>)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
123 |
by (simp add: L.equivalence_axioms assms(2) equivalence_subset) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
124 |
show ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
125 |
apply (unfold_locales, simp_all) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
126 |
using assms(2) apply auto[1] |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
127 |
using assms(2) apply auto[1] |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
128 |
apply (meson L.le_trans assms(2) contra_subsetD) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
129 |
apply (meson L.le_cong assms(2) subsetCE) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
130 |
done |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
131 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
132 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
133 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
134 |
subsubsection \<open>Upper and lower bounds of a set\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
135 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
136 |
definition |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
137 |
Upper :: "[_, 'a set] => 'a set" |
67091 | 138 |
where "Upper L A = {u. (\<forall>x. x \<in> A \<inter> carrier L \<longrightarrow> x \<sqsubseteq>\<^bsub>L\<^esub> u)} \<inter> carrier L" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
139 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
140 |
definition |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
141 |
Lower :: "[_, 'a set] => 'a set" |
67091 | 142 |
where "Lower L A = {l. (\<forall>x. x \<in> A \<inter> carrier L \<longrightarrow> l \<sqsubseteq>\<^bsub>L\<^esub> x)} \<inter> carrier L" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
143 |
|
68004 | 144 |
lemma Lower_dual [simp]: |
145 |
"Lower (inv_gorder L) A = Upper L A" |
|
146 |
by (simp add:Upper_def Lower_def) |
|
147 |
||
148 |
lemma Upper_dual [simp]: |
|
149 |
"Upper (inv_gorder L) A = Lower L A" |
|
150 |
by (simp add:Upper_def Lower_def) |
|
151 |
||
152 |
lemma (in weak_partial_order) equivalence_dual: "equivalence (inv_gorder L)" |
|
153 |
by (rule equivalence.intro) (auto simp: intro: sym trans) |
|
154 |
||
155 |
lemma (in weak_partial_order) dual_weak_order: "weak_partial_order (inv_gorder L)" |
|
156 |
by intro_locales (auto simp add: weak_partial_order_axioms_def le_cong intro: equivalence_dual le_trans) |
|
157 |
||
158 |
lemma (in weak_partial_order) dual_eq_iff [simp]: "A {.=}\<^bsub>inv_gorder L\<^esub> A' \<longleftrightarrow> A {.=} A'" |
|
159 |
by (auto simp: set_eq_def elem_def) |
|
160 |
||
161 |
lemma dual_weak_order_iff: |
|
162 |
"weak_partial_order (inv_gorder A) \<longleftrightarrow> weak_partial_order A" |
|
163 |
proof |
|
164 |
assume "weak_partial_order (inv_gorder A)" |
|
165 |
then interpret dpo: weak_partial_order "inv_gorder A" |
|
166 |
rewrites "carrier (inv_gorder A) = carrier A" |
|
167 |
and "le (inv_gorder A) = (\<lambda> x y. le A y x)" |
|
168 |
and "eq (inv_gorder A) = eq A" |
|
169 |
by (simp_all) |
|
170 |
show "weak_partial_order A" |
|
171 |
by (unfold_locales, auto intro: dpo.sym dpo.trans dpo.le_trans) |
|
172 |
next |
|
173 |
assume "weak_partial_order A" |
|
174 |
thus "weak_partial_order (inv_gorder A)" |
|
175 |
by (metis weak_partial_order.dual_weak_order) |
|
176 |
qed |
|
177 |
||
178 |
lemma Upper_closed [iff]: |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
179 |
"Upper L A \<subseteq> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
180 |
by (unfold Upper_def) clarify |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
181 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
182 |
lemma Upper_memD [dest]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
183 |
fixes L (structure) |
68004 | 184 |
shows "\<lbrakk>u \<in> Upper L A; x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u \<and> u \<in> carrier L" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
185 |
by (unfold Upper_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
186 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
187 |
lemma (in weak_partial_order) Upper_elemD [dest]: |
68004 | 188 |
"\<lbrakk>u .\<in> Upper L A; u \<in> carrier L; x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> u" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
189 |
unfolding Upper_def elem_def |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
190 |
by (blast dest: sym) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
191 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
192 |
lemma Upper_memI: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
193 |
fixes L (structure) |
68004 | 194 |
shows "\<lbrakk>!! y. y \<in> A \<Longrightarrow> y \<sqsubseteq> x; x \<in> carrier L\<rbrakk> \<Longrightarrow> x \<in> Upper L A" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
195 |
by (unfold Upper_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
196 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
197 |
lemma (in weak_partial_order) Upper_elemI: |
68004 | 198 |
"\<lbrakk>!! y. y \<in> A \<Longrightarrow> y \<sqsubseteq> x; x \<in> carrier L\<rbrakk> \<Longrightarrow> x .\<in> Upper L A" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
199 |
unfolding Upper_def by blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
200 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
201 |
lemma Upper_antimono: |
68004 | 202 |
"A \<subseteq> B \<Longrightarrow> Upper L B \<subseteq> Upper L A" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
203 |
by (unfold Upper_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
204 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
205 |
lemma (in weak_partial_order) Upper_is_closed [simp]: |
68004 | 206 |
"A \<subseteq> carrier L \<Longrightarrow> is_closed (Upper L A)" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
207 |
by (rule is_closedI) (blast intro: Upper_memI)+ |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
208 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
209 |
lemma (in weak_partial_order) Upper_mem_cong: |
68004 | 210 |
assumes "a' \<in> carrier L" "A \<subseteq> carrier L" "a .= a'" "a \<in> Upper L A" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
211 |
shows "a' \<in> Upper L A" |
68004 | 212 |
by (metis assms Upper_closed Upper_is_closed closure_of_eq complete_classes) |
213 |
||
214 |
lemma (in weak_partial_order) Upper_semi_cong: |
|
215 |
assumes "A \<subseteq> carrier L" "A {.=} A'" |
|
216 |
shows "Upper L A \<subseteq> Upper L A'" |
|
217 |
unfolding Upper_def |
|
218 |
by clarsimp (meson assms equivalence.refl equivalence_axioms le_cong set_eqD2 subset_eq) |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
219 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
220 |
lemma (in weak_partial_order) Upper_cong: |
68004 | 221 |
assumes "A \<subseteq> carrier L" "A' \<subseteq> carrier L" "A {.=} A'" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
222 |
shows "Upper L A = Upper L A'" |
68004 | 223 |
using assms by (simp add: Upper_semi_cong set_eq_sym subset_antisym) |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
224 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
225 |
lemma Lower_closed [intro!, simp]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
226 |
"Lower L A \<subseteq> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
227 |
by (unfold Lower_def) clarify |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
228 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
229 |
lemma Lower_memD [dest]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
230 |
fixes L (structure) |
68004 | 231 |
shows "\<lbrakk>l \<in> Lower L A; x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> l \<sqsubseteq> x \<and> l \<in> carrier L" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
232 |
by (unfold Lower_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
233 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
234 |
lemma Lower_memI: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
235 |
fixes L (structure) |
68004 | 236 |
shows "\<lbrakk>!! y. y \<in> A \<Longrightarrow> x \<sqsubseteq> y; x \<in> carrier L\<rbrakk> \<Longrightarrow> x \<in> Lower L A" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
237 |
by (unfold Lower_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
238 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
239 |
lemma Lower_antimono: |
68004 | 240 |
"A \<subseteq> B \<Longrightarrow> Lower L B \<subseteq> Lower L A" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
241 |
by (unfold Lower_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
242 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
243 |
lemma (in weak_partial_order) Lower_is_closed [simp]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
244 |
"A \<subseteq> carrier L \<Longrightarrow> is_closed (Lower L A)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
245 |
by (rule is_closedI) (blast intro: Lower_memI dest: sym)+ |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
246 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
247 |
lemma (in weak_partial_order) Lower_mem_cong: |
68004 | 248 |
assumes "a' \<in> carrier L" "A \<subseteq> carrier L" "a .= a'" "a \<in> Lower L A" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
249 |
shows "a' \<in> Lower L A" |
68004 | 250 |
by (meson assms Lower_closed Lower_is_closed is_closed_eq subsetCE) |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
251 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
252 |
lemma (in weak_partial_order) Lower_cong: |
68004 | 253 |
assumes "A \<subseteq> carrier L" "A' \<subseteq> carrier L" "A {.=} A'" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
254 |
shows "Lower L A = Lower L A'" |
68004 | 255 |
unfolding Upper_dual [symmetric] |
256 |
by (rule weak_partial_order.Upper_cong [OF dual_weak_order]) (simp_all add: assms) |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
257 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
258 |
text \<open>Jacobson: Theorem 8.1\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
259 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
260 |
lemma Lower_empty [simp]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
261 |
"Lower L {} = carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
262 |
by (unfold Lower_def) simp |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
263 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
264 |
lemma Upper_empty [simp]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
265 |
"Upper L {} = carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
266 |
by (unfold Upper_def) simp |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
267 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
268 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
269 |
subsubsection \<open>Least and greatest, as predicate\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
270 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
271 |
definition |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
272 |
least :: "[_, 'a, 'a set] => bool" |
67091 | 273 |
where "least L l A \<longleftrightarrow> A \<subseteq> carrier L \<and> l \<in> A \<and> (\<forall>x\<in>A. l \<sqsubseteq>\<^bsub>L\<^esub> x)" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
274 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
275 |
definition |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
276 |
greatest :: "[_, 'a, 'a set] => bool" |
67091 | 277 |
where "greatest L g A \<longleftrightarrow> A \<subseteq> carrier L \<and> g \<in> A \<and> (\<forall>x\<in>A. x \<sqsubseteq>\<^bsub>L\<^esub> g)" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
278 |
|
69597 | 279 |
text (in weak_partial_order) \<open>Could weaken these to \<^term>\<open>l \<in> carrier L \<and> l .\<in> A\<close> and \<^term>\<open>g \<in> carrier L \<and> g .\<in> A\<close>.\<close> |
68004 | 280 |
|
281 |
lemma least_dual [simp]: |
|
282 |
"least (inv_gorder L) x A = greatest L x A" |
|
283 |
by (simp add:least_def greatest_def) |
|
284 |
||
285 |
lemma greatest_dual [simp]: |
|
286 |
"greatest (inv_gorder L) x A = least L x A" |
|
287 |
by (simp add:least_def greatest_def) |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
288 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
289 |
lemma least_closed [intro, simp]: |
68004 | 290 |
"least L l A \<Longrightarrow> l \<in> carrier L" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
291 |
by (unfold least_def) fast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
292 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
293 |
lemma least_mem: |
68004 | 294 |
"least L l A \<Longrightarrow> l \<in> A" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
295 |
by (unfold least_def) fast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
296 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
297 |
lemma (in weak_partial_order) weak_least_unique: |
68004 | 298 |
"\<lbrakk>least L x A; least L y A\<rbrakk> \<Longrightarrow> x .= y" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
299 |
by (unfold least_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
300 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
301 |
lemma least_le: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
302 |
fixes L (structure) |
68004 | 303 |
shows "\<lbrakk>least L x A; a \<in> A\<rbrakk> \<Longrightarrow> x \<sqsubseteq> a" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
304 |
by (unfold least_def) fast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
305 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
306 |
lemma (in weak_partial_order) least_cong: |
68004 | 307 |
"\<lbrakk>x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A\<rbrakk> \<Longrightarrow> least L x A = least L x' A" |
308 |
unfolding least_def |
|
309 |
by (meson is_closed_eq is_closed_eq_rev le_cong local.refl subset_iff) |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
310 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
311 |
abbreviation is_lub :: "[_, 'a, 'a set] => bool" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
312 |
where "is_lub L x A \<equiv> least L x (Upper L A)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
313 |
|
69597 | 314 |
text (in weak_partial_order) \<open>\<^const>\<open>least\<close> is not congruent in the second parameter for |
315 |
\<^term>\<open>A {.=} A'\<close>\<close> |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
316 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
317 |
lemma (in weak_partial_order) least_Upper_cong_l: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
318 |
assumes "x .= x'" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
319 |
and "x \<in> carrier L" "x' \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
320 |
and "A \<subseteq> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
321 |
shows "least L x (Upper L A) = least L x' (Upper L A)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
322 |
apply (rule least_cong) using assms by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
323 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
324 |
lemma (in weak_partial_order) least_Upper_cong_r: |
68004 | 325 |
assumes "A \<subseteq> carrier L" "A' \<subseteq> carrier L" "A {.=} A'" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
326 |
shows "least L x (Upper L A) = least L x (Upper L A')" |
68004 | 327 |
using Upper_cong assms by auto |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
328 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
329 |
lemma least_UpperI: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
330 |
fixes L (structure) |
68004 | 331 |
assumes above: "!! x. x \<in> A \<Longrightarrow> x \<sqsubseteq> s" |
332 |
and below: "!! y. y \<in> Upper L A \<Longrightarrow> s \<sqsubseteq> y" |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
333 |
and L: "A \<subseteq> carrier L" "s \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
334 |
shows "least L s (Upper L A)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
335 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
336 |
have "Upper L A \<subseteq> carrier L" by simp |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
337 |
moreover from above L have "s \<in> Upper L A" by (simp add: Upper_def) |
67613 | 338 |
moreover from below have "\<forall>x \<in> Upper L A. s \<sqsubseteq> x" by fast |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
339 |
ultimately show ?thesis by (simp add: least_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
340 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
341 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
342 |
lemma least_Upper_above: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
343 |
fixes L (structure) |
68004 | 344 |
shows "\<lbrakk>least L s (Upper L A); x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> s" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
345 |
by (unfold least_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
346 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
347 |
lemma greatest_closed [intro, simp]: |
68004 | 348 |
"greatest L l A \<Longrightarrow> l \<in> carrier L" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
349 |
by (unfold greatest_def) fast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
350 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
351 |
lemma greatest_mem: |
68004 | 352 |
"greatest L l A \<Longrightarrow> l \<in> A" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
353 |
by (unfold greatest_def) fast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
354 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
355 |
lemma (in weak_partial_order) weak_greatest_unique: |
68004 | 356 |
"\<lbrakk>greatest L x A; greatest L y A\<rbrakk> \<Longrightarrow> x .= y" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
357 |
by (unfold greatest_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
358 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
359 |
lemma greatest_le: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
360 |
fixes L (structure) |
68004 | 361 |
shows "\<lbrakk>greatest L x A; a \<in> A\<rbrakk> \<Longrightarrow> a \<sqsubseteq> x" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
362 |
by (unfold greatest_def) fast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
363 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
364 |
lemma (in weak_partial_order) greatest_cong: |
68004 | 365 |
"\<lbrakk>x .= x'; x \<in> carrier L; x' \<in> carrier L; is_closed A\<rbrakk> \<Longrightarrow> |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
366 |
greatest L x A = greatest L x' A" |
68004 | 367 |
unfolding greatest_def |
368 |
by (meson is_closed_eq_rev le_cong_r local.sym subset_eq) |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
369 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
370 |
abbreviation is_glb :: "[_, 'a, 'a set] => bool" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
371 |
where "is_glb L x A \<equiv> greatest L x (Lower L A)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
372 |
|
69597 | 373 |
text (in weak_partial_order) \<open>\<^const>\<open>greatest\<close> is not congruent in the second parameter for |
374 |
\<^term>\<open>A {.=} A'\<close> \<close> |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
375 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
376 |
lemma (in weak_partial_order) greatest_Lower_cong_l: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
377 |
assumes "x .= x'" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
378 |
and "x \<in> carrier L" "x' \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
379 |
shows "greatest L x (Lower L A) = greatest L x' (Lower L A)" |
68004 | 380 |
proof - |
381 |
have "\<forall>A. is_closed (Lower L (A \<inter> carrier L))" |
|
382 |
by simp |
|
383 |
then show ?thesis |
|
384 |
by (simp add: Lower_def assms greatest_cong) |
|
385 |
qed |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
386 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
387 |
lemma (in weak_partial_order) greatest_Lower_cong_r: |
68004 | 388 |
assumes "A \<subseteq> carrier L" "A' \<subseteq> carrier L" "A {.=} A'" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
389 |
shows "greatest L x (Lower L A) = greatest L x (Lower L A')" |
68004 | 390 |
using Lower_cong assms by auto |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
391 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
392 |
lemma greatest_LowerI: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
393 |
fixes L (structure) |
68004 | 394 |
assumes below: "!! x. x \<in> A \<Longrightarrow> i \<sqsubseteq> x" |
395 |
and above: "!! y. y \<in> Lower L A \<Longrightarrow> y \<sqsubseteq> i" |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
396 |
and L: "A \<subseteq> carrier L" "i \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
397 |
shows "greatest L i (Lower L A)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
398 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
399 |
have "Lower L A \<subseteq> carrier L" by simp |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
400 |
moreover from below L have "i \<in> Lower L A" by (simp add: Lower_def) |
67613 | 401 |
moreover from above have "\<forall>x \<in> Lower L A. x \<sqsubseteq> i" by fast |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
402 |
ultimately show ?thesis by (simp add: greatest_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
403 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
404 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
405 |
lemma greatest_Lower_below: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
406 |
fixes L (structure) |
68004 | 407 |
shows "\<lbrakk>greatest L i (Lower L A); x \<in> A; A \<subseteq> carrier L\<rbrakk> \<Longrightarrow> i \<sqsubseteq> x" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
408 |
by (unfold greatest_def) blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
409 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
410 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
411 |
subsubsection \<open>Intervals\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
412 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
413 |
definition |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
414 |
at_least_at_most :: "('a, 'c) gorder_scheme \<Rightarrow> 'a => 'a => 'a set" ("(1\<lbrace>_.._\<rbrace>\<index>)") |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
415 |
where "\<lbrace>l..u\<rbrace>\<^bsub>A\<^esub> = {x \<in> carrier A. l \<sqsubseteq>\<^bsub>A\<^esub> x \<and> x \<sqsubseteq>\<^bsub>A\<^esub> u}" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
416 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
417 |
context weak_partial_order |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
418 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
419 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
420 |
lemma at_least_at_most_upper [dest]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
421 |
"x \<in> \<lbrace>a..b\<rbrace> \<Longrightarrow> x \<sqsubseteq> b" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
422 |
by (simp add: at_least_at_most_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
423 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
424 |
lemma at_least_at_most_lower [dest]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
425 |
"x \<in> \<lbrace>a..b\<rbrace> \<Longrightarrow> a \<sqsubseteq> x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
426 |
by (simp add: at_least_at_most_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
427 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
428 |
lemma at_least_at_most_closed: "\<lbrace>a..b\<rbrace> \<subseteq> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
429 |
by (auto simp add: at_least_at_most_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
430 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
431 |
lemma at_least_at_most_member [intro]: |
68004 | 432 |
"\<lbrakk>x \<in> carrier L; a \<sqsubseteq> x; x \<sqsubseteq> b\<rbrakk> \<Longrightarrow> x \<in> \<lbrace>a..b\<rbrace>" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
433 |
by (simp add: at_least_at_most_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
434 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
435 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
436 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
437 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
438 |
subsubsection \<open>Isotone functions\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
439 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
440 |
definition isotone :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
441 |
where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
442 |
"isotone A B f \<equiv> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
443 |
weak_partial_order A \<and> weak_partial_order B \<and> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
444 |
(\<forall>x\<in>carrier A. \<forall>y\<in>carrier A. x \<sqsubseteq>\<^bsub>A\<^esub> y \<longrightarrow> f x \<sqsubseteq>\<^bsub>B\<^esub> f y)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
445 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
446 |
lemma isotoneI [intro?]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
447 |
fixes f :: "'a \<Rightarrow> 'b" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
448 |
assumes "weak_partial_order L1" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
449 |
"weak_partial_order L2" |
68004 | 450 |
"(\<And>x y. \<lbrakk>x \<in> carrier L1; y \<in> carrier L1; x \<sqsubseteq>\<^bsub>L1\<^esub> y\<rbrakk> |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
451 |
\<Longrightarrow> f x \<sqsubseteq>\<^bsub>L2\<^esub> f y)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
452 |
shows "isotone L1 L2 f" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
453 |
using assms by (auto simp add:isotone_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
454 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
455 |
abbreviation Monotone :: "('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" ("Mono\<index>") |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
456 |
where "Monotone L f \<equiv> isotone L L f" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
457 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
458 |
lemma use_iso1: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
459 |
"\<lbrakk>isotone A A f; x \<in> carrier A; y \<in> carrier A; x \<sqsubseteq>\<^bsub>A\<^esub> y\<rbrakk> \<Longrightarrow> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
460 |
f x \<sqsubseteq>\<^bsub>A\<^esub> f y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
461 |
by (simp add: isotone_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
462 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
463 |
lemma use_iso2: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
464 |
"\<lbrakk>isotone A B f; x \<in> carrier A; y \<in> carrier A; x \<sqsubseteq>\<^bsub>A\<^esub> y\<rbrakk> \<Longrightarrow> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
465 |
f x \<sqsubseteq>\<^bsub>B\<^esub> f y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
466 |
by (simp add: isotone_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
467 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
468 |
lemma iso_compose: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
469 |
"\<lbrakk>f \<in> carrier A \<rightarrow> carrier B; isotone A B f; g \<in> carrier B \<rightarrow> carrier C; isotone B C g\<rbrakk> \<Longrightarrow> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
470 |
isotone A C (g \<circ> f)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
471 |
by (simp add: isotone_def, safe, metis Pi_iff) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
472 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
473 |
lemma (in weak_partial_order) inv_isotone [simp]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
474 |
"isotone (inv_gorder A) (inv_gorder B) f = isotone A B f" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
475 |
by (auto simp add:isotone_def dual_weak_order dual_weak_order_iff) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
476 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
477 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
478 |
subsubsection \<open>Idempotent functions\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
479 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
480 |
definition idempotent :: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
481 |
"('a, 'b) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" ("Idem\<index>") where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
482 |
"idempotent L f \<equiv> \<forall>x\<in>carrier L. f (f x) .=\<^bsub>L\<^esub> f x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
483 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
484 |
lemma (in weak_partial_order) idempotent: |
68004 | 485 |
"\<lbrakk>Idem f; x \<in> carrier L\<rbrakk> \<Longrightarrow> f (f x) .= f x" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
486 |
by (auto simp add: idempotent_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
487 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
488 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
489 |
subsubsection \<open>Order embeddings\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
490 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
491 |
definition order_emb :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
492 |
where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
493 |
"order_emb A B f \<equiv> weak_partial_order A |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
494 |
\<and> weak_partial_order B |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
495 |
\<and> (\<forall>x\<in>carrier A. \<forall>y\<in>carrier A. f x \<sqsubseteq>\<^bsub>B\<^esub> f y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>A\<^esub> y )" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
496 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
497 |
lemma order_emb_isotone: "order_emb A B f \<Longrightarrow> isotone A B f" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
498 |
by (auto simp add: isotone_def order_emb_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
499 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
500 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
501 |
subsubsection \<open>Commuting functions\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
502 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
503 |
definition commuting :: "('a, 'c) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
504 |
"commuting A f g = (\<forall>x\<in>carrier A. (f \<circ> g) x .=\<^bsub>A\<^esub> (g \<circ> f) x)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
505 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
506 |
subsection \<open>Partial orders where \<open>eq\<close> is the Equality\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
507 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
508 |
locale partial_order = weak_partial_order + |
67399 | 509 |
assumes eq_is_equal: "(.=) = (=)" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
510 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
511 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
512 |
declare weak_le_antisym [rule del] |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
513 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
514 |
lemma le_antisym [intro]: |
68004 | 515 |
"\<lbrakk>x \<sqsubseteq> y; y \<sqsubseteq> x; x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x = y" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
516 |
using weak_le_antisym unfolding eq_is_equal . |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
517 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
518 |
lemma lless_eq: |
67091 | 519 |
"x \<sqsubset> y \<longleftrightarrow> x \<sqsubseteq> y \<and> x \<noteq> y" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
520 |
unfolding lless_def by (simp add: eq_is_equal) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
521 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
522 |
lemma set_eq_is_eq: "A {.=} B \<longleftrightarrow> A = B" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
523 |
by (auto simp add: set_eq_def elem_def eq_is_equal) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
524 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
525 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
526 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
527 |
lemma (in partial_order) dual_order: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
528 |
"partial_order (inv_gorder L)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
529 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
530 |
interpret dwo: weak_partial_order "inv_gorder L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
531 |
by (metis dual_weak_order) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
532 |
show ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
533 |
by (unfold_locales, simp add:eq_is_equal) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
534 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
535 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
536 |
lemma dual_order_iff: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
537 |
"partial_order (inv_gorder A) \<longleftrightarrow> partial_order A" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
538 |
proof |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
539 |
assume assm:"partial_order (inv_gorder A)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
540 |
then interpret po: partial_order "inv_gorder A" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
541 |
rewrites "carrier (inv_gorder A) = carrier A" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
542 |
and "le (inv_gorder A) = (\<lambda> x y. le A y x)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
543 |
and "eq (inv_gorder A) = eq A" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
544 |
by (simp_all) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
545 |
show "partial_order A" |
68004 | 546 |
apply (unfold_locales, simp_all add: po.sym) |
547 |
apply (metis po.trans) |
|
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
548 |
apply (metis po.weak_le_antisym, metis po.le_trans) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
549 |
apply (metis (full_types) po.eq_is_equal, metis po.eq_is_equal) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
550 |
done |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
551 |
next |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
552 |
assume "partial_order A" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
553 |
thus "partial_order (inv_gorder A)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
554 |
by (metis partial_order.dual_order) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
555 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
556 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
557 |
text \<open>Least and greatest, as predicate\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
558 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
559 |
lemma (in partial_order) least_unique: |
68004 | 560 |
"\<lbrakk>least L x A; least L y A\<rbrakk> \<Longrightarrow> x = y" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
561 |
using weak_least_unique unfolding eq_is_equal . |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
562 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
563 |
lemma (in partial_order) greatest_unique: |
68004 | 564 |
"\<lbrakk>greatest L x A; greatest L y A\<rbrakk> \<Longrightarrow> x = y" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
565 |
using weak_greatest_unique unfolding eq_is_equal . |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
566 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
567 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
568 |
subsection \<open>Bounded Orders\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
569 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
570 |
definition |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
571 |
top :: "_ => 'a" ("\<top>\<index>") where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
572 |
"\<top>\<^bsub>L\<^esub> = (SOME x. greatest L x (carrier L))" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
573 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
574 |
definition |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
575 |
bottom :: "_ => 'a" ("\<bottom>\<index>") where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
576 |
"\<bottom>\<^bsub>L\<^esub> = (SOME x. least L x (carrier L))" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
577 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
578 |
locale weak_partial_order_bottom = weak_partial_order L for L (structure) + |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
579 |
assumes bottom_exists: "\<exists> x. least L x (carrier L)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
580 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
581 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
582 |
lemma bottom_least: "least L \<bottom> (carrier L)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
583 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
584 |
obtain x where "least L x (carrier L)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
585 |
by (metis bottom_exists) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
586 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
587 |
thus ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
588 |
by (auto intro:someI2 simp add: bottom_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
589 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
590 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
591 |
lemma bottom_closed [simp, intro]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
592 |
"\<bottom> \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
593 |
by (metis bottom_least least_mem) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
594 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
595 |
lemma bottom_lower [simp, intro]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
596 |
"x \<in> carrier L \<Longrightarrow> \<bottom> \<sqsubseteq> x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
597 |
by (metis bottom_least least_le) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
598 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
599 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
600 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
601 |
locale weak_partial_order_top = weak_partial_order L for L (structure) + |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
602 |
assumes top_exists: "\<exists> x. greatest L x (carrier L)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
603 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
604 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
605 |
lemma top_greatest: "greatest L \<top> (carrier L)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
606 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
607 |
obtain x where "greatest L x (carrier L)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
608 |
by (metis top_exists) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
609 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
610 |
thus ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
611 |
by (auto intro:someI2 simp add: top_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
612 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
613 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
614 |
lemma top_closed [simp, intro]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
615 |
"\<top> \<in> carrier L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
616 |
by (metis greatest_mem top_greatest) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
617 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
618 |
lemma top_higher [simp, intro]: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
619 |
"x \<in> carrier L \<Longrightarrow> x \<sqsubseteq> \<top>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
620 |
by (metis greatest_le top_greatest) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
621 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
622 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
623 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
624 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
625 |
subsection \<open>Total Orders\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
626 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
627 |
locale weak_total_order = weak_partial_order + |
68004 | 628 |
assumes total: "\<lbrakk>x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
629 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
630 |
text \<open>Introduction rule: the usual definition of total order\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
631 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
632 |
lemma (in weak_partial_order) weak_total_orderI: |
68004 | 633 |
assumes total: "!!x y. \<lbrakk>x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
634 |
shows "weak_total_order L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
635 |
by unfold_locales (rule total) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
636 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
637 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
638 |
subsection \<open>Total orders where \<open>eq\<close> is the Equality\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
639 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
640 |
locale total_order = partial_order + |
68004 | 641 |
assumes total_order_total: "\<lbrakk>x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
642 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
643 |
sublocale total_order < weak?: weak_total_order |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
644 |
by unfold_locales (rule total_order_total) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
645 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
646 |
text \<open>Introduction rule: the usual definition of total order\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
647 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
648 |
lemma (in partial_order) total_orderI: |
68004 | 649 |
assumes total: "!!x y. \<lbrakk>x \<in> carrier L; y \<in> carrier L\<rbrakk> \<Longrightarrow> x \<sqsubseteq> y \<or> y \<sqsubseteq> x" |
65099
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
650 |
shows "total_order L" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
651 |
by unfold_locales (rule total) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
652 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
653 |
end |