author | wenzelm |
Thu, 08 Dec 2022 11:24:43 +0100 | |
changeset 76598 | 9f97eda3fcf1 |
parent 65099 | 30d0b2f1df76 |
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/Galois_Connection.thy |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
2 |
Author: Alasdair Armstrong and Simon Foster |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
3 |
Copyright: Alasdair Armstrong and Simon Foster |
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 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
6 |
theory Galois_Connection |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
7 |
imports Complete_Lattice |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
8 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
9 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
10 |
section \<open>Galois connections\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
11 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
12 |
subsection \<open>Definition and basic properties\<close> |
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 |
record ('a, 'b, 'c, 'd) galcon = |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
15 |
orderA :: "('a, 'c) gorder_scheme" ("\<X>\<index>") |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
16 |
orderB :: "('b, 'd) gorder_scheme" ("\<Y>\<index>") |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
17 |
lower :: "'a \<Rightarrow> 'b" ("\<pi>\<^sup>*\<index>") |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
18 |
upper :: "'b \<Rightarrow> 'a" ("\<pi>\<^sub>*\<index>") |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
19 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
20 |
type_synonym ('a, 'b) galois = "('a, 'b, unit, unit) galcon" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
21 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
22 |
abbreviation "inv_galcon G \<equiv> \<lparr> orderA = inv_gorder \<Y>\<^bsub>G\<^esub>, orderB = inv_gorder \<X>\<^bsub>G\<^esub>, lower = upper G, upper = lower G \<rparr>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
23 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
24 |
definition comp_galcon :: "('b, 'c) galois \<Rightarrow> ('a, 'b) galois \<Rightarrow> ('a, 'c) galois" (infixr "\<circ>\<^sub>g" 85) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
25 |
where "G \<circ>\<^sub>g F = \<lparr> orderA = orderA F, orderB = orderB G, lower = lower G \<circ> lower F, upper = upper F \<circ> upper G \<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 |
definition id_galcon :: "'a gorder \<Rightarrow> ('a, 'a) galois" ("I\<^sub>g") where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
28 |
"I\<^sub>g(A) = \<lparr> orderA = A, orderB = A, lower = id, upper = id \<rparr>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
29 |
|
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 |
subsection \<open>Well-typed connections\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
32 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
33 |
locale connection = |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
34 |
fixes G (structure) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
35 |
assumes is_order_A: "partial_order \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
36 |
and is_order_B: "partial_order \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
37 |
and lower_closure: "\<pi>\<^sup>* \<in> carrier \<X> \<rightarrow> carrier \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
38 |
and upper_closure: "\<pi>\<^sub>* \<in> carrier \<Y> \<rightarrow> carrier \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
39 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
40 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
41 |
lemma lower_closed: "x \<in> carrier \<X> \<Longrightarrow> \<pi>\<^sup>* x \<in> carrier \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
42 |
using lower_closure by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
43 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
44 |
lemma upper_closed: "y \<in> carrier \<Y> \<Longrightarrow> \<pi>\<^sub>* y \<in> carrier \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
45 |
using upper_closure by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
46 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
47 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
48 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
49 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
50 |
subsection \<open>Galois connections\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
51 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
52 |
locale galois_connection = connection + |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
53 |
assumes galois_property: "\<lbrakk>x \<in> carrier \<X>; y \<in> carrier \<Y>\<rbrakk> \<Longrightarrow> \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
54 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
55 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
56 |
lemma is_weak_order_A: "weak_partial_order \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
57 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
58 |
interpret po: partial_order \<X> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
59 |
by (metis is_order_A) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
60 |
show ?thesis .. |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
61 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
62 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
63 |
lemma is_weak_order_B: "weak_partial_order \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
64 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
65 |
interpret po: partial_order \<Y> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
66 |
by (metis is_order_B) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
67 |
show ?thesis .. |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
68 |
qed |
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 right: "\<lbrakk>x \<in> carrier \<X>; y \<in> carrier \<Y>; \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y\<rbrakk> \<Longrightarrow> x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
71 |
by (metis galois_property) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
72 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
73 |
lemma left: "\<lbrakk>x \<in> carrier \<X>; y \<in> carrier \<Y>; x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y\<rbrakk> \<Longrightarrow> \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
74 |
by (metis galois_property) |
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 deflation: "y \<in> carrier \<Y> \<Longrightarrow> \<pi>\<^sup>* (\<pi>\<^sub>* y) \<sqsubseteq>\<^bsub>\<Y>\<^esub> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
77 |
by (metis Pi_iff is_weak_order_A left upper_closure weak_partial_order.le_refl) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
78 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
79 |
lemma inflation: "x \<in> carrier \<X> \<Longrightarrow> x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* (\<pi>\<^sup>* x)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
80 |
by (metis (no_types, lifting) PiE galois_connection.right galois_connection_axioms is_weak_order_B lower_closure weak_partial_order.le_refl) |
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 lower_iso: "isotone \<X> \<Y> \<pi>\<^sup>*" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
83 |
proof (auto simp add:isotone_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
84 |
show "weak_partial_order \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
85 |
by (metis is_weak_order_A) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
86 |
show "weak_partial_order \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
87 |
by (metis is_weak_order_B) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
88 |
fix x y |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
89 |
assume a: "x \<in> carrier \<X>" "y \<in> carrier \<X>" "x \<sqsubseteq>\<^bsub>\<X>\<^esub> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
90 |
have b: "\<pi>\<^sup>* y \<in> carrier \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
91 |
using a(2) lower_closure by blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
92 |
then have "\<pi>\<^sub>* (\<pi>\<^sup>* y) \<in> carrier \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
93 |
using upper_closure by blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
94 |
then have "x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* (\<pi>\<^sup>* y)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
95 |
by (meson a inflation is_weak_order_A weak_partial_order.le_trans) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
96 |
thus "\<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> \<pi>\<^sup>* y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
97 |
by (meson b a(1) Pi_iff galois_property lower_closure upper_closure) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
98 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
99 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
100 |
lemma upper_iso: "isotone \<Y> \<X> \<pi>\<^sub>*" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
101 |
apply (auto simp add:isotone_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
102 |
apply (metis is_weak_order_B) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
103 |
apply (metis is_weak_order_A) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
104 |
apply (metis (no_types, lifting) Pi_mem deflation is_weak_order_B lower_closure right upper_closure weak_partial_order.le_trans) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
105 |
done |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
106 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
107 |
lemma lower_comp: "x \<in> carrier \<X> \<Longrightarrow> \<pi>\<^sup>* (\<pi>\<^sub>* (\<pi>\<^sup>* x)) = \<pi>\<^sup>* x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
108 |
by (meson deflation funcset_mem inflation is_order_B lower_closure lower_iso partial_order.le_antisym upper_closure use_iso2) |
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 lower_comp': "x \<in> carrier \<X> \<Longrightarrow> (\<pi>\<^sup>* \<circ> \<pi>\<^sub>* \<circ> \<pi>\<^sup>*) x = \<pi>\<^sup>* x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
111 |
by (simp add: lower_comp) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
112 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
113 |
lemma upper_comp: "y \<in> carrier \<Y> \<Longrightarrow> \<pi>\<^sub>* (\<pi>\<^sup>* (\<pi>\<^sub>* y)) = \<pi>\<^sub>* y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
114 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
115 |
assume a1: "y \<in> carrier \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
116 |
hence f1: "\<pi>\<^sub>* y \<in> carrier \<X>" using upper_closure by blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
117 |
have f2: "\<pi>\<^sup>* (\<pi>\<^sub>* y) \<sqsubseteq>\<^bsub>\<Y>\<^esub> y" using a1 deflation by blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
118 |
have f3: "\<pi>\<^sub>* (\<pi>\<^sup>* (\<pi>\<^sub>* y)) \<in> carrier \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
119 |
using f1 lower_closure upper_closure by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
120 |
have "\<pi>\<^sup>* (\<pi>\<^sub>* y) \<in> carrier \<Y>" using f1 lower_closure by blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
121 |
thus "\<pi>\<^sub>* (\<pi>\<^sup>* (\<pi>\<^sub>* y)) = \<pi>\<^sub>* y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
122 |
by (meson a1 f1 f2 f3 inflation is_order_A partial_order.le_antisym upper_iso use_iso2) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
123 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
124 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
125 |
lemma upper_comp': "y \<in> carrier \<Y> \<Longrightarrow> (\<pi>\<^sub>* \<circ> \<pi>\<^sup>* \<circ> \<pi>\<^sub>*) y = \<pi>\<^sub>* y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
126 |
by (simp add: upper_comp) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
127 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
128 |
lemma adjoint_idem1: "idempotent \<Y> (\<pi>\<^sup>* \<circ> \<pi>\<^sub>*)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
129 |
by (simp add: idempotent_def is_order_B partial_order.eq_is_equal upper_comp) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
130 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
131 |
lemma adjoint_idem2: "idempotent \<X> (\<pi>\<^sub>* \<circ> \<pi>\<^sup>*)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
132 |
by (simp add: idempotent_def is_order_A partial_order.eq_is_equal lower_comp) |
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 |
lemma fg_iso: "isotone \<Y> \<Y> (\<pi>\<^sup>* \<circ> \<pi>\<^sub>*)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
135 |
by (metis iso_compose lower_closure lower_iso upper_closure upper_iso) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
136 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
137 |
lemma gf_iso: "isotone \<X> \<X> (\<pi>\<^sub>* \<circ> \<pi>\<^sup>*)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
138 |
by (metis iso_compose lower_closure lower_iso upper_closure upper_iso) |
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 |
lemma semi_inverse1: "x \<in> carrier \<X> \<Longrightarrow> \<pi>\<^sup>* x = \<pi>\<^sup>* (\<pi>\<^sub>* (\<pi>\<^sup>* x))" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
141 |
by (metis lower_comp) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
142 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
143 |
lemma semi_inverse2: "x \<in> carrier \<Y> \<Longrightarrow> \<pi>\<^sub>* x = \<pi>\<^sub>* (\<pi>\<^sup>* (\<pi>\<^sub>* x))" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
144 |
by (metis upper_comp) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
145 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
146 |
theorem lower_by_complete_lattice: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
147 |
assumes "complete_lattice \<Y>" "x \<in> carrier \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
148 |
shows "\<pi>\<^sup>*(x) = \<Sqinter>\<^bsub>\<Y>\<^esub> { y \<in> carrier \<Y>. x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>*(y) }" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
149 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
150 |
interpret Y: complete_lattice \<Y> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
151 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
152 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
153 |
show ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
154 |
proof (rule Y.le_antisym) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
155 |
show x: "\<pi>\<^sup>* x \<in> carrier \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
156 |
using assms(2) lower_closure by blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
157 |
show "\<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> \<Sqinter>\<^bsub>\<Y>\<^esub>{y \<in> carrier \<Y>. x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y}" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
158 |
proof (rule Y.weak.inf_greatest) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
159 |
show "{y \<in> carrier \<Y>. x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y} \<subseteq> carrier \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
160 |
by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
161 |
show "\<pi>\<^sup>* x \<in> carrier \<Y>" by (fact x) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
162 |
fix z |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
163 |
assume "z \<in> {y \<in> carrier \<Y>. x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y}" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
164 |
thus "\<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> z" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
165 |
using assms(2) left by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
166 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
167 |
show "\<Sqinter>\<^bsub>\<Y>\<^esub>{y \<in> carrier \<Y>. x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y} \<sqsubseteq>\<^bsub>\<Y>\<^esub> \<pi>\<^sup>* x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
168 |
proof (rule Y.weak.inf_lower) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
169 |
show "{y \<in> carrier \<Y>. x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y} \<subseteq> carrier \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
170 |
by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
171 |
show "\<pi>\<^sup>* x \<in> {y \<in> carrier \<Y>. x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y}" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
172 |
proof (auto) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
173 |
show "\<pi>\<^sup>* x \<in> carrier \<Y>" by (fact x) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
174 |
show "x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* (\<pi>\<^sup>* x)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
175 |
using assms(2) inflation by blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
176 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
177 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
178 |
show "\<Sqinter>\<^bsub>\<Y>\<^esub>{y \<in> carrier \<Y>. x \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y} \<in> carrier \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
179 |
by (auto intro: Y.weak.inf_closed) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
180 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
181 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
182 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
183 |
theorem upper_by_complete_lattice: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
184 |
assumes "complete_lattice \<X>" "y \<in> carrier \<Y>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
185 |
shows "\<pi>\<^sub>*(y) = \<Squnion>\<^bsub>\<X>\<^esub> { x \<in> carrier \<X>. \<pi>\<^sup>*(x) \<sqsubseteq>\<^bsub>\<Y>\<^esub> y }" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
186 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
187 |
interpret X: complete_lattice \<X> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
188 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
189 |
show ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
190 |
proof (rule X.le_antisym) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
191 |
show y: "\<pi>\<^sub>* y \<in> carrier \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
192 |
using assms(2) upper_closure by blast |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
193 |
show "\<pi>\<^sub>* y \<sqsubseteq>\<^bsub>\<X>\<^esub> \<Squnion>\<^bsub>\<X>\<^esub>{x \<in> carrier \<X>. \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y}" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
194 |
proof (rule X.weak.sup_upper) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
195 |
show "{x \<in> carrier \<X>. \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y} \<subseteq> carrier \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
196 |
by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
197 |
show "\<pi>\<^sub>* y \<in> {x \<in> carrier \<X>. \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y}" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
198 |
proof (auto) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
199 |
show "\<pi>\<^sub>* y \<in> carrier \<X>" by (fact y) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
200 |
show "\<pi>\<^sup>* (\<pi>\<^sub>* y) \<sqsubseteq>\<^bsub>\<Y>\<^esub> y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
201 |
by (simp add: assms(2) deflation) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
202 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
203 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
204 |
show "\<Squnion>\<^bsub>\<X>\<^esub>{x \<in> carrier \<X>. \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y} \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
205 |
proof (rule X.weak.sup_least) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
206 |
show "{x \<in> carrier \<X>. \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y} \<subseteq> carrier \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
207 |
by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
208 |
show "\<pi>\<^sub>* y \<in> carrier \<X>" by (fact y) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
209 |
fix z |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
210 |
assume "z \<in> {x \<in> carrier \<X>. \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y}" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
211 |
thus "z \<sqsubseteq>\<^bsub>\<X>\<^esub> \<pi>\<^sub>* y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
212 |
by (simp add: assms(2) right) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
213 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
214 |
show "\<Squnion>\<^bsub>\<X>\<^esub>{x \<in> carrier \<X>. \<pi>\<^sup>* x \<sqsubseteq>\<^bsub>\<Y>\<^esub> y} \<in> carrier \<X>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
215 |
by (auto intro: X.weak.sup_closed) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
216 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
217 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
218 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
219 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
220 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
221 |
lemma dual_galois [simp]: " galois_connection \<lparr> orderA = inv_gorder B, orderB = inv_gorder A, lower = f, upper = g \<rparr> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
222 |
= galois_connection \<lparr> orderA = A, orderB = B, lower = g, upper = f \<rparr>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
223 |
by (auto simp add: galois_connection_def galois_connection_axioms_def connection_def dual_order_iff) |
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 |
definition lower_adjoint :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
226 |
"lower_adjoint A B f \<equiv> \<exists>g. galois_connection \<lparr> orderA = A, orderB = B, lower = f, upper = g \<rparr>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
227 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
228 |
definition upper_adjoint :: "('a, 'c) gorder_scheme \<Rightarrow> ('b, 'd) gorder_scheme \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> bool" where |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
229 |
"upper_adjoint A B g \<equiv> \<exists>f. galois_connection \<lparr> orderA = A, orderB = B, lower = f, upper = g \<rparr>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
230 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
231 |
lemma lower_adjoint_dual [simp]: "lower_adjoint (inv_gorder A) (inv_gorder B) f = upper_adjoint B A f" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
232 |
by (simp add: lower_adjoint_def upper_adjoint_def) |
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 upper_adjoint_dual [simp]: "upper_adjoint (inv_gorder A) (inv_gorder B) f = lower_adjoint B A f" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
235 |
by (simp add: lower_adjoint_def upper_adjoint_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
236 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
237 |
lemma lower_type: "lower_adjoint A B f \<Longrightarrow> f \<in> carrier A \<rightarrow> carrier B" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
238 |
by (auto simp add:lower_adjoint_def galois_connection_def galois_connection_axioms_def connection_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
239 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
240 |
lemma upper_type: "upper_adjoint A B g \<Longrightarrow> g \<in> carrier B \<rightarrow> carrier A" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
241 |
by (auto simp add:upper_adjoint_def galois_connection_def galois_connection_axioms_def connection_def) |
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 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
244 |
subsection \<open>Composition of Galois connections\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
245 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
246 |
lemma id_galois: "partial_order A \<Longrightarrow> galois_connection (I\<^sub>g(A))" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
247 |
by (simp add: id_galcon_def galois_connection_def galois_connection_axioms_def connection_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
248 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
249 |
lemma comp_galcon_closed: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
250 |
assumes "galois_connection G" "galois_connection F" "\<Y>\<^bsub>F\<^esub> = \<X>\<^bsub>G\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
251 |
shows "galois_connection (G \<circ>\<^sub>g F)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
252 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
253 |
interpret F: galois_connection F |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
254 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
255 |
interpret G: galois_connection G |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
256 |
by (simp add: assms) |
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 |
have "partial_order \<X>\<^bsub>G \<circ>\<^sub>g F\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
259 |
by (simp add: F.is_order_A comp_galcon_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
260 |
moreover have "partial_order \<Y>\<^bsub>G \<circ>\<^sub>g F\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
261 |
by (simp add: G.is_order_B comp_galcon_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
262 |
moreover have "\<pi>\<^sup>*\<^bsub>G\<^esub> \<circ> \<pi>\<^sup>*\<^bsub>F\<^esub> \<in> carrier \<X>\<^bsub>F\<^esub> \<rightarrow> carrier \<Y>\<^bsub>G\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
263 |
using F.lower_closure G.lower_closure assms(3) by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
264 |
moreover have "\<pi>\<^sub>*\<^bsub>F\<^esub> \<circ> \<pi>\<^sub>*\<^bsub>G\<^esub> \<in> carrier \<Y>\<^bsub>G\<^esub> \<rightarrow> carrier \<X>\<^bsub>F\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
265 |
using F.upper_closure G.upper_closure assms(3) by auto |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
266 |
moreover |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
267 |
have "\<And> x y. \<lbrakk>x \<in> carrier \<X>\<^bsub>F\<^esub>; y \<in> carrier \<Y>\<^bsub>G\<^esub> \<rbrakk> \<Longrightarrow> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
268 |
(\<pi>\<^sup>*\<^bsub>G\<^esub> (\<pi>\<^sup>*\<^bsub>F\<^esub> x) \<sqsubseteq>\<^bsub>\<Y>\<^bsub>G\<^esub>\<^esub> y) = (x \<sqsubseteq>\<^bsub>\<X>\<^bsub>F\<^esub>\<^esub> \<pi>\<^sub>*\<^bsub>F\<^esub> (\<pi>\<^sub>*\<^bsub>G\<^esub> y))" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
269 |
by (metis F.galois_property F.lower_closure G.galois_property G.upper_closure assms(3) Pi_iff) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
270 |
ultimately show ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
271 |
by (simp add: comp_galcon_def galois_connection_def galois_connection_axioms_def connection_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
272 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
273 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
274 |
lemma comp_galcon_right_unit [simp]: "F \<circ>\<^sub>g I\<^sub>g(\<X>\<^bsub>F\<^esub>) = F" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
275 |
by (simp add: comp_galcon_def id_galcon_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
276 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
277 |
lemma comp_galcon_left_unit [simp]: "I\<^sub>g(\<Y>\<^bsub>F\<^esub>) \<circ>\<^sub>g F = F" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
278 |
by (simp add: comp_galcon_def id_galcon_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
279 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
280 |
lemma galois_connectionI: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
281 |
assumes |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
282 |
"partial_order A" "partial_order B" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
283 |
"L \<in> carrier A \<rightarrow> carrier B" "R \<in> carrier B \<rightarrow> carrier A" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
284 |
"isotone A B L" "isotone B A R" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
285 |
"\<And> x y. \<lbrakk> x \<in> carrier A; y \<in> carrier B \<rbrakk> \<Longrightarrow> L x \<sqsubseteq>\<^bsub>B\<^esub> y \<longleftrightarrow> x \<sqsubseteq>\<^bsub>A\<^esub> R y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
286 |
shows "galois_connection \<lparr> orderA = A, orderB = B, lower = L, upper = R \<rparr>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
287 |
using assms by (simp add: galois_connection_def connection_def galois_connection_axioms_def) |
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 galois_connectionI': |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
290 |
assumes |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
291 |
"partial_order A" "partial_order B" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
292 |
"L \<in> carrier A \<rightarrow> carrier B" "R \<in> carrier B \<rightarrow> carrier A" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
293 |
"isotone A B L" "isotone B A R" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
294 |
"\<And> X. X \<in> carrier(B) \<Longrightarrow> L(R(X)) \<sqsubseteq>\<^bsub>B\<^esub> X" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
295 |
"\<And> X. X \<in> carrier(A) \<Longrightarrow> X \<sqsubseteq>\<^bsub>A\<^esub> R(L(X))" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
296 |
shows "galois_connection \<lparr> orderA = A, orderB = B, lower = L, upper = R \<rparr>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
297 |
using assms |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
298 |
by (auto simp add: galois_connection_def connection_def galois_connection_axioms_def, (meson PiE isotone_def weak_partial_order.le_trans)+) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
299 |
|
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 |
subsection \<open>Retracts\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
302 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
303 |
locale retract = galois_connection + |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
304 |
assumes retract_property: "x \<in> carrier \<X> \<Longrightarrow> \<pi>\<^sub>* (\<pi>\<^sup>* x) \<sqsubseteq>\<^bsub>\<X>\<^esub> x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
305 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
306 |
lemma retract_inverse: "x \<in> carrier \<X> \<Longrightarrow> \<pi>\<^sub>* (\<pi>\<^sup>* x) = x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
307 |
by (meson funcset_mem inflation is_order_A lower_closure partial_order.le_antisym retract_axioms retract_axioms_def retract_def upper_closure) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
308 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
309 |
lemma retract_injective: "inj_on \<pi>\<^sup>* (carrier \<X>)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
310 |
by (metis inj_onI retract_inverse) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
311 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
312 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
313 |
theorem comp_retract_closed: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
314 |
assumes "retract G" "retract F" "\<Y>\<^bsub>F\<^esub> = \<X>\<^bsub>G\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
315 |
shows "retract (G \<circ>\<^sub>g F)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
316 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
317 |
interpret f: retract F |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
318 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
319 |
interpret g: retract G |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
320 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
321 |
interpret gf: galois_connection "(G \<circ>\<^sub>g F)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
322 |
by (simp add: assms(1) assms(2) assms(3) comp_galcon_closed retract.axioms(1)) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
323 |
show ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
324 |
proof |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
325 |
fix x |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
326 |
assume "x \<in> carrier \<X>\<^bsub>G \<circ>\<^sub>g F\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
327 |
thus "le \<X>\<^bsub>G \<circ>\<^sub>g F\<^esub> (\<pi>\<^sub>*\<^bsub>G \<circ>\<^sub>g F\<^esub> (\<pi>\<^sup>*\<^bsub>G \<circ>\<^sub>g F\<^esub> x)) x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
328 |
using assms(3) f.inflation f.lower_closed f.retract_inverse g.retract_inverse by (auto simp add: comp_galcon_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
329 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
330 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
331 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
332 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
333 |
subsection \<open>Coretracts\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
334 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
335 |
locale coretract = galois_connection + |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
336 |
assumes coretract_property: "y \<in> carrier \<Y> \<Longrightarrow> y \<sqsubseteq>\<^bsub>\<Y>\<^esub> \<pi>\<^sup>* (\<pi>\<^sub>* y)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
337 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
338 |
lemma coretract_inverse: "y \<in> carrier \<Y> \<Longrightarrow> \<pi>\<^sup>* (\<pi>\<^sub>* y) = y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
339 |
by (meson coretract_axioms coretract_axioms_def coretract_def deflation funcset_mem is_order_B lower_closure partial_order.le_antisym upper_closure) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
340 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
341 |
lemma retract_injective: "inj_on \<pi>\<^sub>* (carrier \<Y>)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
342 |
by (metis coretract_inverse inj_onI) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
343 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
344 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
345 |
theorem comp_coretract_closed: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
346 |
assumes "coretract G" "coretract F" "\<Y>\<^bsub>F\<^esub> = \<X>\<^bsub>G\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
347 |
shows "coretract (G \<circ>\<^sub>g F)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
348 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
349 |
interpret f: coretract F |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
350 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
351 |
interpret g: coretract G |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
352 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
353 |
interpret gf: galois_connection "(G \<circ>\<^sub>g F)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
354 |
by (simp add: assms(1) assms(2) assms(3) comp_galcon_closed coretract.axioms(1)) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
355 |
show ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
356 |
proof |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
357 |
fix y |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
358 |
assume "y \<in> carrier \<Y>\<^bsub>G \<circ>\<^sub>g F\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
359 |
thus "le \<Y>\<^bsub>G \<circ>\<^sub>g F\<^esub> y (\<pi>\<^sup>*\<^bsub>G \<circ>\<^sub>g F\<^esub> (\<pi>\<^sub>*\<^bsub>G \<circ>\<^sub>g F\<^esub> y))" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
360 |
by (simp add: comp_galcon_def assms(3) f.coretract_inverse g.coretract_property g.upper_closed) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
361 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
362 |
qed |
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 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
365 |
subsection \<open>Galois Bijections\<close> |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
366 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
367 |
locale galois_bijection = connection + |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
368 |
assumes lower_iso: "isotone \<X> \<Y> \<pi>\<^sup>*" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
369 |
and upper_iso: "isotone \<Y> \<X> \<pi>\<^sub>*" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
370 |
and lower_inv_eq: "x \<in> carrier \<X> \<Longrightarrow> \<pi>\<^sub>* (\<pi>\<^sup>* x) = x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
371 |
and upper_inv_eq: "y \<in> carrier \<Y> \<Longrightarrow> \<pi>\<^sup>* (\<pi>\<^sub>* y) = y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
372 |
begin |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
373 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
374 |
lemma lower_bij: "bij_betw \<pi>\<^sup>* (carrier \<X>) (carrier \<Y>)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
375 |
by (rule bij_betwI[where g="\<pi>\<^sub>*"], auto intro: upper_inv_eq lower_inv_eq upper_closed lower_closed) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
376 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
377 |
lemma upper_bij: "bij_betw \<pi>\<^sub>* (carrier \<Y>) (carrier \<X>)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
378 |
by (rule bij_betwI[where g="\<pi>\<^sup>*"], auto intro: upper_inv_eq lower_inv_eq upper_closed lower_closed) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
379 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
380 |
sublocale gal_bij_conn: galois_connection |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
381 |
apply (unfold_locales, auto) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
382 |
using lower_closed lower_inv_eq upper_iso use_iso2 apply fastforce |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
383 |
using lower_iso upper_closed upper_inv_eq use_iso2 apply fastforce |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
384 |
done |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
385 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
386 |
sublocale gal_bij_ret: retract |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
387 |
by (unfold_locales, simp add: gal_bij_conn.is_weak_order_A lower_inv_eq weak_partial_order.le_refl) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
388 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
389 |
sublocale gal_bij_coret: coretract |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
390 |
by (unfold_locales, simp add: gal_bij_conn.is_weak_order_B upper_inv_eq weak_partial_order.le_refl) |
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 |
end |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
393 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
394 |
theorem comp_galois_bijection_closed: |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
395 |
assumes "galois_bijection G" "galois_bijection F" "\<Y>\<^bsub>F\<^esub> = \<X>\<^bsub>G\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
396 |
shows "galois_bijection (G \<circ>\<^sub>g F)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
397 |
proof - |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
398 |
interpret f: galois_bijection F |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
399 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
400 |
interpret g: galois_bijection G |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
401 |
by (simp add: assms) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
402 |
interpret gf: galois_connection "(G \<circ>\<^sub>g F)" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
403 |
by (simp add: assms(3) comp_galcon_closed f.gal_bij_conn.galois_connection_axioms g.gal_bij_conn.galois_connection_axioms galois_connection.axioms(1)) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
404 |
show ?thesis |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
405 |
proof |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
406 |
show "isotone \<X>\<^bsub>G \<circ>\<^sub>g F\<^esub> \<Y>\<^bsub>G \<circ>\<^sub>g F\<^esub> \<pi>\<^sup>*\<^bsub>G \<circ>\<^sub>g F\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
407 |
by (simp add: comp_galcon_def, metis comp_galcon_def galcon.select_convs(1) galcon.select_convs(2) galcon.select_convs(3) gf.lower_iso) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
408 |
show "isotone \<Y>\<^bsub>G \<circ>\<^sub>g F\<^esub> \<X>\<^bsub>G \<circ>\<^sub>g F\<^esub> \<pi>\<^sub>*\<^bsub>G \<circ>\<^sub>g F\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
409 |
by (simp add: gf.upper_iso) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
410 |
fix x |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
411 |
assume "x \<in> carrier \<X>\<^bsub>G \<circ>\<^sub>g F\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
412 |
thus "\<pi>\<^sub>*\<^bsub>G \<circ>\<^sub>g F\<^esub> (\<pi>\<^sup>*\<^bsub>G \<circ>\<^sub>g F\<^esub> x) = x" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
413 |
using assms(3) f.lower_closed f.lower_inv_eq g.lower_inv_eq by (auto simp add: comp_galcon_def) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
414 |
next |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
415 |
fix y |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
416 |
assume "y \<in> carrier \<Y>\<^bsub>G \<circ>\<^sub>g F\<^esub>" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
417 |
thus "\<pi>\<^sup>*\<^bsub>G \<circ>\<^sub>g F\<^esub> (\<pi>\<^sub>*\<^bsub>G \<circ>\<^sub>g F\<^esub> y) = y" |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
418 |
by (simp add: comp_galcon_def assms(3) f.upper_inv_eq g.upper_closed g.upper_inv_eq) |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
419 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
420 |
qed |
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
421 |
|
30d0b2f1df76
Knaster-Tarski fixed point theorem and Galois Connections.
ballarin
parents:
diff
changeset
|
422 |
end |