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