author | haftmann |
Tue, 04 May 2010 08:55:43 +0200 | |
changeset 36635 | 080b755377c0 |
parent 35317 | d57da4abb47d |
child 37678 | 0040bafffdef |
permissions | -rw-r--r-- |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
1 |
(* Title: HOL/Lattice/Orders.thy |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
2 |
Author: Markus Wenzel, TU Muenchen |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
3 |
*) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
4 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
5 |
header {* Orders *} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
6 |
|
16417 | 7 |
theory Orders imports Main begin |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
8 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
9 |
subsection {* Ordered structures *} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
10 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
11 |
text {* |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
12 |
We define several classes of ordered structures over some type @{typ |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
13 |
'a} with relation @{text "\<sqsubseteq> \<Colon> 'a \<Rightarrow> 'a \<Rightarrow> bool"}. For a |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
14 |
\emph{quasi-order} that relation is required to be reflexive and |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
15 |
transitive, for a \emph{partial order} it also has to be |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
16 |
anti-symmetric, while for a \emph{linear order} all elements are |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
17 |
required to be related (in either direction). |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
18 |
*} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
19 |
|
35317 | 20 |
class leq = |
21 |
fixes leq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "[=" 50) |
|
22 |
||
21210 | 23 |
notation (xsymbols) |
19736 | 24 |
leq (infixl "\<sqsubseteq>" 50) |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
25 |
|
35317 | 26 |
class quasi_order = leq + |
27 |
assumes leq_refl [intro?]: "x \<sqsubseteq> x" |
|
28 |
assumes leq_trans [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z" |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
29 |
|
35317 | 30 |
class partial_order = quasi_order + |
31 |
assumes leq_antisym [trans]: "x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> x \<Longrightarrow> x = y" |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
32 |
|
35317 | 33 |
class linear_order = partial_order + |
34 |
assumes leq_linear: "x \<sqsubseteq> y \<or> y \<sqsubseteq> x" |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
35 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
36 |
lemma linear_order_cases: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
37 |
"((x::'a::linear_order) \<sqsubseteq> y \<Longrightarrow> C) \<Longrightarrow> (y \<sqsubseteq> x \<Longrightarrow> C) \<Longrightarrow> C" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
38 |
by (insert leq_linear) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
39 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
40 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
41 |
subsection {* Duality *} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
42 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
43 |
text {* |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
44 |
The \emph{dual} of an ordered structure is an isomorphic copy of the |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
45 |
underlying type, with the @{text \<sqsubseteq>} relation defined as the inverse |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
46 |
of the original one. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
47 |
*} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
48 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
49 |
datatype 'a dual = dual 'a |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
50 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
51 |
consts |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
52 |
undual :: "'a dual \<Rightarrow> 'a" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
53 |
primrec |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
54 |
undual_dual: "undual (dual x) = x" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
55 |
|
35317 | 56 |
instantiation dual :: (leq) leq |
57 |
begin |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
58 |
|
35317 | 59 |
definition |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
60 |
leq_dual_def: "x' \<sqsubseteq> y' \<equiv> undual y' \<sqsubseteq> undual x'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
61 |
|
35317 | 62 |
instance .. |
63 |
||
64 |
end |
|
65 |
||
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
66 |
lemma undual_leq [iff?]: "(undual x' \<sqsubseteq> undual y') = (y' \<sqsubseteq> x')" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
67 |
by (simp add: leq_dual_def) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
68 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
69 |
lemma dual_leq [iff?]: "(dual x \<sqsubseteq> dual y) = (y \<sqsubseteq> x)" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
70 |
by (simp add: leq_dual_def) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
71 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
72 |
text {* |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
73 |
\medskip Functions @{term dual} and @{term undual} are inverse to |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
74 |
each other; this entails the following fundamental properties. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
75 |
*} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
76 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
77 |
lemma dual_undual [simp]: "dual (undual x') = x'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
78 |
by (cases x') simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
79 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
80 |
lemma undual_dual_id [simp]: "undual o dual = id" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
81 |
by (rule ext) simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
82 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
83 |
lemma dual_undual_id [simp]: "dual o undual = id" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
84 |
by (rule ext) simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
85 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
86 |
text {* |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
87 |
\medskip Since @{term dual} (and @{term undual}) are both injective |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
88 |
and surjective, the basic logical connectives (equality, |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
89 |
quantification etc.) are transferred as follows. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
90 |
*} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
91 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
92 |
lemma undual_equality [iff?]: "(undual x' = undual y') = (x' = y')" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
93 |
by (cases x', cases y') simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
94 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
95 |
lemma dual_equality [iff?]: "(dual x = dual y) = (x = y)" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
96 |
by simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
97 |
|
10834 | 98 |
lemma dual_ball [iff?]: "(\<forall>x \<in> A. P (dual x)) = (\<forall>x' \<in> dual ` A. P x')" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
99 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
100 |
assume a: "\<forall>x \<in> A. P (dual x)" |
10834 | 101 |
show "\<forall>x' \<in> dual ` A. P x'" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
102 |
proof |
10834 | 103 |
fix x' assume x': "x' \<in> dual ` A" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
104 |
have "undual x' \<in> A" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
105 |
proof - |
10834 | 106 |
from x' have "undual x' \<in> undual ` dual ` A" by simp |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
107 |
thus "undual x' \<in> A" by (simp add: image_compose [symmetric]) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
108 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
109 |
with a have "P (dual (undual x'))" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
110 |
also have "\<dots> = x'" by simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
111 |
finally show "P x'" . |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
112 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
113 |
next |
10834 | 114 |
assume a: "\<forall>x' \<in> dual ` A. P x'" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
115 |
show "\<forall>x \<in> A. P (dual x)" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
116 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
117 |
fix x assume "x \<in> A" |
10834 | 118 |
hence "dual x \<in> dual ` A" by simp |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
119 |
with a show "P (dual x)" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
120 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
121 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
122 |
|
10834 | 123 |
lemma range_dual [simp]: "dual ` UNIV = UNIV" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
124 |
proof (rule surj_range) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
125 |
have "\<And>x'. dual (undual x') = x'" by simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
126 |
thus "surj dual" by (rule surjI) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
127 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
128 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
129 |
lemma dual_all [iff?]: "(\<forall>x. P (dual x)) = (\<forall>x'. P x')" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
130 |
proof - |
10834 | 131 |
have "(\<forall>x \<in> UNIV. P (dual x)) = (\<forall>x' \<in> dual ` UNIV. P x')" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
132 |
by (rule dual_ball) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
133 |
thus ?thesis by simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
134 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
135 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
136 |
lemma dual_ex: "(\<exists>x. P (dual x)) = (\<exists>x'. P x')" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
137 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
138 |
have "(\<forall>x. \<not> P (dual x)) = (\<forall>x'. \<not> P x')" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
139 |
by (rule dual_all) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
140 |
thus ?thesis by blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
141 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
142 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
143 |
lemma dual_Collect: "{dual x| x. P (dual x)} = {x'. P x'}" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
144 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
145 |
have "{dual x| x. P (dual x)} = {x'. \<exists>x''. x' = x'' \<and> P x''}" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
146 |
by (simp only: dual_ex [symmetric]) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
147 |
thus ?thesis by blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
148 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
149 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
150 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
151 |
subsection {* Transforming orders *} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
152 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
153 |
subsubsection {* Duals *} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
154 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
155 |
text {* |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
156 |
The classes of quasi, partial, and linear orders are all closed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
157 |
under formation of dual structures. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
158 |
*} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
159 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
160 |
instance dual :: (quasi_order) quasi_order |
10309 | 161 |
proof |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
162 |
fix x' y' z' :: "'a::quasi_order dual" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
163 |
have "undual x' \<sqsubseteq> undual x'" .. thus "x' \<sqsubseteq> x'" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
164 |
assume "y' \<sqsubseteq> z'" hence "undual z' \<sqsubseteq> undual y'" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
165 |
also assume "x' \<sqsubseteq> y'" hence "undual y' \<sqsubseteq> undual x'" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
166 |
finally show "x' \<sqsubseteq> z'" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
167 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
168 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
169 |
instance dual :: (partial_order) partial_order |
10309 | 170 |
proof |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
171 |
fix x' y' :: "'a::partial_order dual" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
172 |
assume "y' \<sqsubseteq> x'" hence "undual x' \<sqsubseteq> undual y'" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
173 |
also assume "x' \<sqsubseteq> y'" hence "undual y' \<sqsubseteq> undual x'" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
174 |
finally show "x' = y'" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
175 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
176 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
177 |
instance dual :: (linear_order) linear_order |
10309 | 178 |
proof |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
179 |
fix x' y' :: "'a::linear_order dual" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
180 |
show "x' \<sqsubseteq> y' \<or> y' \<sqsubseteq> x'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
181 |
proof (rule linear_order_cases) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
182 |
assume "undual y' \<sqsubseteq> undual x'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
183 |
hence "x' \<sqsubseteq> y'" .. thus ?thesis .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
184 |
next |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
185 |
assume "undual x' \<sqsubseteq> undual y'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
186 |
hence "y' \<sqsubseteq> x'" .. thus ?thesis .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
187 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
188 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
189 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
190 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
191 |
subsubsection {* Binary products \label{sec:prod-order} *} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
192 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
193 |
text {* |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
194 |
The classes of quasi and partial orders are closed under binary |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
195 |
products. Note that the direct product of linear orders need |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
196 |
\emph{not} be linear in general. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
197 |
*} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
198 |
|
35317 | 199 |
instantiation * :: (leq, leq) leq |
200 |
begin |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
201 |
|
35317 | 202 |
definition |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
203 |
leq_prod_def: "p \<sqsubseteq> q \<equiv> fst p \<sqsubseteq> fst q \<and> snd p \<sqsubseteq> snd q" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
204 |
|
35317 | 205 |
instance .. |
206 |
||
207 |
end |
|
208 |
||
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
209 |
lemma leq_prodI [intro?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
210 |
"fst p \<sqsubseteq> fst q \<Longrightarrow> snd p \<sqsubseteq> snd q \<Longrightarrow> p \<sqsubseteq> q" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
211 |
by (unfold leq_prod_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
212 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
213 |
lemma leq_prodE [elim?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
214 |
"p \<sqsubseteq> q \<Longrightarrow> (fst p \<sqsubseteq> fst q \<Longrightarrow> snd p \<sqsubseteq> snd q \<Longrightarrow> C) \<Longrightarrow> C" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
215 |
by (unfold leq_prod_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
216 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
217 |
instance * :: (quasi_order, quasi_order) quasi_order |
10309 | 218 |
proof |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
219 |
fix p q r :: "'a::quasi_order \<times> 'b::quasi_order" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
220 |
show "p \<sqsubseteq> p" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
221 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
222 |
show "fst p \<sqsubseteq> fst p" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
223 |
show "snd p \<sqsubseteq> snd p" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
224 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
225 |
assume pq: "p \<sqsubseteq> q" and qr: "q \<sqsubseteq> r" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
226 |
show "p \<sqsubseteq> r" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
227 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
228 |
from pq have "fst p \<sqsubseteq> fst q" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
229 |
also from qr have "\<dots> \<sqsubseteq> fst r" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
230 |
finally show "fst p \<sqsubseteq> fst r" . |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
231 |
from pq have "snd p \<sqsubseteq> snd q" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
232 |
also from qr have "\<dots> \<sqsubseteq> snd r" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
233 |
finally show "snd p \<sqsubseteq> snd r" . |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
234 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
235 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
236 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
237 |
instance * :: (partial_order, partial_order) partial_order |
10309 | 238 |
proof |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
239 |
fix p q :: "'a::partial_order \<times> 'b::partial_order" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
240 |
assume pq: "p \<sqsubseteq> q" and qp: "q \<sqsubseteq> p" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
241 |
show "p = q" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
242 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
243 |
from pq have "fst p \<sqsubseteq> fst q" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
244 |
also from qp have "\<dots> \<sqsubseteq> fst p" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
245 |
finally show "fst p = fst q" . |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
246 |
from pq have "snd p \<sqsubseteq> snd q" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
247 |
also from qp have "\<dots> \<sqsubseteq> snd p" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
248 |
finally show "snd p = snd q" . |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
249 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
250 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
251 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
252 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
253 |
subsubsection {* General products \label{sec:fun-order} *} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
254 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
255 |
text {* |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
256 |
The classes of quasi and partial orders are closed under general |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
257 |
products (function spaces). Note that the direct product of linear |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
258 |
orders need \emph{not} be linear in general. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
259 |
*} |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
260 |
|
35317 | 261 |
instantiation "fun" :: (type, leq) leq |
262 |
begin |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
263 |
|
35317 | 264 |
definition |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
265 |
leq_fun_def: "f \<sqsubseteq> g \<equiv> \<forall>x. f x \<sqsubseteq> g x" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
266 |
|
35317 | 267 |
instance .. |
268 |
||
269 |
end |
|
270 |
||
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
271 |
lemma leq_funI [intro?]: "(\<And>x. f x \<sqsubseteq> g x) \<Longrightarrow> f \<sqsubseteq> g" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
272 |
by (unfold leq_fun_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
273 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
274 |
lemma leq_funD [dest?]: "f \<sqsubseteq> g \<Longrightarrow> f x \<sqsubseteq> g x" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
275 |
by (unfold leq_fun_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
276 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
277 |
instance "fun" :: (type, quasi_order) quasi_order |
10309 | 278 |
proof |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
279 |
fix f g h :: "'a \<Rightarrow> 'b::quasi_order" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
280 |
show "f \<sqsubseteq> f" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
281 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
282 |
fix x show "f x \<sqsubseteq> f x" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
283 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
284 |
assume fg: "f \<sqsubseteq> g" and gh: "g \<sqsubseteq> h" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
285 |
show "f \<sqsubseteq> h" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
286 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
287 |
fix x from fg have "f x \<sqsubseteq> g x" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
288 |
also from gh have "\<dots> \<sqsubseteq> h x" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
289 |
finally show "f x \<sqsubseteq> h x" . |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
290 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
291 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
292 |
|
20523
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
krauss
parents:
19736
diff
changeset
|
293 |
instance "fun" :: (type, partial_order) partial_order |
10309 | 294 |
proof |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
295 |
fix f g :: "'a \<Rightarrow> 'b::partial_order" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
296 |
assume fg: "f \<sqsubseteq> g" and gf: "g \<sqsubseteq> f" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
297 |
show "f = g" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
298 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
299 |
fix x from fg have "f x \<sqsubseteq> g x" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
300 |
also from gf have "\<dots> \<sqsubseteq> f x" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
301 |
finally show "f x = g x" . |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
302 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
303 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
304 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
305 |
end |