author | wenzelm |
Mon, 25 May 2020 22:37:14 +0200 | |
changeset 71892 | dff81ce866d4 |
parent 61986 | 2461779da2b8 |
permissions | -rw-r--r-- |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
1 |
(* Title: HOL/Lattice/Bounds.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>Bounds\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
6 |
|
16417 | 7 |
theory Bounds imports Orders begin |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
8 |
|
36176
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
wenzelm
parents:
35317
diff
changeset
|
9 |
hide_const (open) inf sup |
22426 | 10 |
|
61986 | 11 |
subsection \<open>Infimum and supremum\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
12 |
|
61986 | 13 |
text \<open> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
14 |
Given a partial order, we define infimum (greatest lower bound) and |
61986 | 15 |
supremum (least upper bound) wrt.\ \<open>\<sqsubseteq>\<close> for two and for any |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
16 |
number of elements. |
61986 | 17 |
\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
18 |
|
19736 | 19 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
20 |
is_inf :: "'a::partial_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where |
19736 | 21 |
"is_inf x y inf = (inf \<sqsubseteq> x \<and> inf \<sqsubseteq> y \<and> (\<forall>z. z \<sqsubseteq> x \<and> z \<sqsubseteq> y \<longrightarrow> z \<sqsubseteq> inf))" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
22 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
23 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
24 |
is_sup :: "'a::partial_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where |
19736 | 25 |
"is_sup x y sup = (x \<sqsubseteq> sup \<and> y \<sqsubseteq> sup \<and> (\<forall>z. x \<sqsubseteq> z \<and> y \<sqsubseteq> z \<longrightarrow> sup \<sqsubseteq> z))" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
26 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
27 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
28 |
is_Inf :: "'a::partial_order set \<Rightarrow> 'a \<Rightarrow> bool" where |
19736 | 29 |
"is_Inf A inf = ((\<forall>x \<in> A. inf \<sqsubseteq> x) \<and> (\<forall>z. (\<forall>x \<in> A. z \<sqsubseteq> x) \<longrightarrow> z \<sqsubseteq> inf))" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
30 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
31 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
32 |
is_Sup :: "'a::partial_order set \<Rightarrow> 'a \<Rightarrow> bool" where |
19736 | 33 |
"is_Sup A sup = ((\<forall>x \<in> A. x \<sqsubseteq> sup) \<and> (\<forall>z. (\<forall>x \<in> A. x \<sqsubseteq> z) \<longrightarrow> sup \<sqsubseteq> z))" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
34 |
|
61986 | 35 |
text \<open> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
36 |
These definitions entail the following basic properties of boundary |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
37 |
elements. |
61986 | 38 |
\<close> |
10157
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 |
lemma is_infI [intro?]: "inf \<sqsubseteq> x \<Longrightarrow> inf \<sqsubseteq> y \<Longrightarrow> |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
41 |
(\<And>z. z \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> y \<Longrightarrow> z \<sqsubseteq> inf) \<Longrightarrow> is_inf x y inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
42 |
by (unfold is_inf_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
43 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
44 |
lemma is_inf_greatest [elim?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
45 |
"is_inf x y inf \<Longrightarrow> z \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> y \<Longrightarrow> z \<sqsubseteq> inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
46 |
by (unfold is_inf_def) blast |
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 |
lemma is_inf_lower [elim?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
49 |
"is_inf x y inf \<Longrightarrow> (inf \<sqsubseteq> x \<Longrightarrow> inf \<sqsubseteq> y \<Longrightarrow> C) \<Longrightarrow> C" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
50 |
by (unfold is_inf_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
51 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
52 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
53 |
lemma is_supI [intro?]: "x \<sqsubseteq> sup \<Longrightarrow> y \<sqsubseteq> sup \<Longrightarrow> |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
54 |
(\<And>z. x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> sup \<sqsubseteq> z) \<Longrightarrow> is_sup x y sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
55 |
by (unfold is_sup_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
56 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
57 |
lemma is_sup_least [elim?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
58 |
"is_sup x y sup \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> sup \<sqsubseteq> z" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
59 |
by (unfold is_sup_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
60 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
61 |
lemma is_sup_upper [elim?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
62 |
"is_sup x y sup \<Longrightarrow> (x \<sqsubseteq> sup \<Longrightarrow> y \<sqsubseteq> sup \<Longrightarrow> C) \<Longrightarrow> C" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
63 |
by (unfold is_sup_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
64 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
65 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
66 |
lemma is_InfI [intro?]: "(\<And>x. x \<in> A \<Longrightarrow> inf \<sqsubseteq> x) \<Longrightarrow> |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
67 |
(\<And>z. (\<forall>x \<in> A. z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> inf) \<Longrightarrow> is_Inf A inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
68 |
by (unfold is_Inf_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
69 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
70 |
lemma is_Inf_greatest [elim?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
71 |
"is_Inf A inf \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> z \<sqsubseteq> x) \<Longrightarrow> z \<sqsubseteq> inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
72 |
by (unfold is_Inf_def) blast |
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 is_Inf_lower [dest?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
75 |
"is_Inf A inf \<Longrightarrow> x \<in> A \<Longrightarrow> inf \<sqsubseteq> x" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
76 |
by (unfold is_Inf_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
77 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
78 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
79 |
lemma is_SupI [intro?]: "(\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> sup) \<Longrightarrow> |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
80 |
(\<And>z. (\<forall>x \<in> A. x \<sqsubseteq> z) \<Longrightarrow> sup \<sqsubseteq> z) \<Longrightarrow> is_Sup A sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
81 |
by (unfold is_Sup_def) blast |
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 is_Sup_least [elim?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
84 |
"is_Sup A sup \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> x \<sqsubseteq> z) \<Longrightarrow> sup \<sqsubseteq> z" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
85 |
by (unfold is_Sup_def) blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
86 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
87 |
lemma is_Sup_upper [dest?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
88 |
"is_Sup A sup \<Longrightarrow> x \<in> A \<Longrightarrow> x \<sqsubseteq> sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
89 |
by (unfold is_Sup_def) blast |
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 |
|
61986 | 92 |
subsection \<open>Duality\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
93 |
|
61986 | 94 |
text \<open> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
95 |
Infimum and supremum are dual to each other. |
61986 | 96 |
\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
97 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
98 |
theorem dual_inf [iff?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
99 |
"is_inf (dual x) (dual y) (dual sup) = is_sup x y sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
100 |
by (simp add: is_inf_def is_sup_def dual_all [symmetric] dual_leq) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
101 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
102 |
theorem dual_sup [iff?]: |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
103 |
"is_sup (dual x) (dual y) (dual inf) = is_inf x y inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
104 |
by (simp add: is_inf_def is_sup_def dual_all [symmetric] dual_leq) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
105 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
106 |
theorem dual_Inf [iff?]: |
10834 | 107 |
"is_Inf (dual ` A) (dual sup) = is_Sup A sup" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
108 |
by (simp add: is_Inf_def is_Sup_def dual_all [symmetric] dual_leq) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
109 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
110 |
theorem dual_Sup [iff?]: |
10834 | 111 |
"is_Sup (dual ` A) (dual inf) = is_Inf A inf" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
112 |
by (simp add: is_Inf_def is_Sup_def dual_all [symmetric] dual_leq) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
113 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
114 |
|
61986 | 115 |
subsection \<open>Uniqueness\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
116 |
|
61986 | 117 |
text \<open> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
118 |
Infima and suprema on partial orders are unique; this is mainly due |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
119 |
to anti-symmetry of the underlying relation. |
61986 | 120 |
\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
121 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
122 |
theorem is_inf_uniq: "is_inf x y inf \<Longrightarrow> is_inf x y inf' \<Longrightarrow> inf = inf'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
123 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
124 |
assume inf: "is_inf x y inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
125 |
assume inf': "is_inf x y inf'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
126 |
show ?thesis |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
127 |
proof (rule leq_antisym) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
128 |
from inf' show "inf \<sqsubseteq> inf'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
129 |
proof (rule is_inf_greatest) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
130 |
from inf show "inf \<sqsubseteq> x" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
131 |
from inf show "inf \<sqsubseteq> y" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
132 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
133 |
from inf show "inf' \<sqsubseteq> inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
134 |
proof (rule is_inf_greatest) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
135 |
from inf' show "inf' \<sqsubseteq> x" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
136 |
from inf' show "inf' \<sqsubseteq> y" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
137 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
138 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
139 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
140 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
141 |
theorem is_sup_uniq: "is_sup x y sup \<Longrightarrow> is_sup x y sup' \<Longrightarrow> sup = sup'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
142 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
143 |
assume sup: "is_sup x y sup" and sup': "is_sup x y sup'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
144 |
have "dual sup = dual sup'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
145 |
proof (rule is_inf_uniq) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
146 |
from sup show "is_inf (dual x) (dual y) (dual sup)" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
147 |
from sup' show "is_inf (dual x) (dual y) (dual sup')" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
148 |
qed |
23373 | 149 |
then show "sup = sup'" .. |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
150 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
151 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
152 |
theorem is_Inf_uniq: "is_Inf A inf \<Longrightarrow> is_Inf A inf' \<Longrightarrow> inf = inf'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
153 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
154 |
assume inf: "is_Inf A inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
155 |
assume inf': "is_Inf A inf'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
156 |
show ?thesis |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
157 |
proof (rule leq_antisym) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
158 |
from inf' show "inf \<sqsubseteq> inf'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
159 |
proof (rule is_Inf_greatest) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
160 |
fix x assume "x \<in> A" |
23373 | 161 |
with inf show "inf \<sqsubseteq> x" .. |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
162 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
163 |
from inf show "inf' \<sqsubseteq> inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
164 |
proof (rule is_Inf_greatest) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
165 |
fix x assume "x \<in> A" |
23373 | 166 |
with inf' show "inf' \<sqsubseteq> x" .. |
10157
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 |
qed |
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 |
theorem is_Sup_uniq: "is_Sup A sup \<Longrightarrow> is_Sup A sup' \<Longrightarrow> sup = sup'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
172 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
173 |
assume sup: "is_Sup A sup" and sup': "is_Sup A sup'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
174 |
have "dual sup = dual sup'" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
175 |
proof (rule is_Inf_uniq) |
10834 | 176 |
from sup show "is_Inf (dual ` A) (dual sup)" .. |
177 |
from sup' show "is_Inf (dual ` A) (dual sup')" .. |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
178 |
qed |
23373 | 179 |
then show "sup = sup'" .. |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
180 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
181 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
182 |
|
61986 | 183 |
subsection \<open>Related elements\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
184 |
|
61986 | 185 |
text \<open> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
186 |
The binary bound of related elements is either one of the argument. |
61986 | 187 |
\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
188 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
189 |
theorem is_inf_related [elim?]: "x \<sqsubseteq> y \<Longrightarrow> is_inf x y x" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
190 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
191 |
assume "x \<sqsubseteq> y" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
192 |
show ?thesis |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
193 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
194 |
show "x \<sqsubseteq> x" .. |
23393 | 195 |
show "x \<sqsubseteq> y" by fact |
196 |
fix z assume "z \<sqsubseteq> x" and "z \<sqsubseteq> y" show "z \<sqsubseteq> x" by fact |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
197 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
198 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
199 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
200 |
theorem is_sup_related [elim?]: "x \<sqsubseteq> y \<Longrightarrow> is_sup x y y" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
201 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
202 |
assume "x \<sqsubseteq> y" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
203 |
show ?thesis |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
204 |
proof |
23393 | 205 |
show "x \<sqsubseteq> y" by fact |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
206 |
show "y \<sqsubseteq> y" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
207 |
fix z assume "x \<sqsubseteq> z" and "y \<sqsubseteq> z" |
23393 | 208 |
show "y \<sqsubseteq> z" by fact |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
209 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
210 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
211 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
212 |
|
61986 | 213 |
subsection \<open>General versus binary bounds \label{sec:gen-bin-bounds}\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
214 |
|
61986 | 215 |
text \<open> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
216 |
General bounds of two-element sets coincide with binary bounds. |
61986 | 217 |
\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
218 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
219 |
theorem is_Inf_binary: "is_Inf {x, y} inf = is_inf x y inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
220 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
221 |
let ?A = "{x, y}" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
222 |
show ?thesis |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
223 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
224 |
assume is_Inf: "is_Inf ?A inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
225 |
show "is_inf x y inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
226 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
227 |
have "x \<in> ?A" by simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
228 |
with is_Inf show "inf \<sqsubseteq> x" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
229 |
have "y \<in> ?A" by simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
230 |
with is_Inf show "inf \<sqsubseteq> y" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
231 |
fix z assume zx: "z \<sqsubseteq> x" and zy: "z \<sqsubseteq> y" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
232 |
from is_Inf show "z \<sqsubseteq> inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
233 |
proof (rule is_Inf_greatest) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
234 |
fix a assume "a \<in> ?A" |
23373 | 235 |
then have "a = x \<or> a = y" by blast |
236 |
then show "z \<sqsubseteq> a" |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
237 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
238 |
assume "a = x" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
239 |
with zx show ?thesis by simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
240 |
next |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
241 |
assume "a = y" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
242 |
with zy show ?thesis by simp |
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 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
246 |
next |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
247 |
assume is_inf: "is_inf x y inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
248 |
show "is_Inf {x, y} inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
249 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
250 |
fix a assume "a \<in> ?A" |
23373 | 251 |
then have "a = x \<or> a = y" by blast |
252 |
then show "inf \<sqsubseteq> a" |
|
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
253 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
254 |
assume "a = x" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
255 |
also from is_inf have "inf \<sqsubseteq> x" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
256 |
finally show ?thesis . |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
257 |
next |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
258 |
assume "a = y" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
259 |
also from is_inf have "inf \<sqsubseteq> y" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
260 |
finally show ?thesis . |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
261 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
262 |
next |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
263 |
fix z assume z: "\<forall>a \<in> ?A. z \<sqsubseteq> a" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
264 |
from is_inf show "z \<sqsubseteq> inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
265 |
proof (rule is_inf_greatest) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
266 |
from z show "z \<sqsubseteq> x" by blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
267 |
from z show "z \<sqsubseteq> y" by blast |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
268 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
269 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
270 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
271 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
272 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
273 |
theorem is_Sup_binary: "is_Sup {x, y} sup = is_sup x y sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
274 |
proof - |
10834 | 275 |
have "is_Sup {x, y} sup = is_Inf (dual ` {x, y}) (dual sup)" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
276 |
by (simp only: dual_Inf) |
10834 | 277 |
also have "dual ` {x, y} = {dual x, dual y}" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
278 |
by simp |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
279 |
also have "is_Inf \<dots> (dual sup) = is_inf (dual x) (dual y) (dual sup)" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
280 |
by (rule is_Inf_binary) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
281 |
also have "\<dots> = is_sup x y sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
282 |
by (simp only: dual_inf) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
283 |
finally show ?thesis . |
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 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
286 |
|
61986 | 287 |
subsection \<open>Connecting general bounds \label{sec:connect-bounds}\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
288 |
|
61986 | 289 |
text \<open> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
290 |
Either kind of general bounds is sufficient to express the other. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
291 |
The least upper bound (supremum) is the same as the the greatest |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
292 |
lower bound of the set of all upper bounds; the dual statements |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
293 |
holds as well; the dual statement holds as well. |
61986 | 294 |
\<close> |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
295 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
296 |
theorem Inf_Sup: "is_Inf {b. \<forall>a \<in> A. a \<sqsubseteq> b} sup \<Longrightarrow> is_Sup A sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
297 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
298 |
let ?B = "{b. \<forall>a \<in> A. a \<sqsubseteq> b}" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
299 |
assume is_Inf: "is_Inf ?B sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
300 |
show "is_Sup A sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
301 |
proof |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
302 |
fix x assume x: "x \<in> A" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
303 |
from is_Inf show "x \<sqsubseteq> sup" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
304 |
proof (rule is_Inf_greatest) |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
305 |
fix y assume "y \<in> ?B" |
23373 | 306 |
then have "\<forall>a \<in> A. a \<sqsubseteq> y" .. |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
307 |
from this x show "x \<sqsubseteq> y" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
308 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
309 |
next |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
310 |
fix z assume "\<forall>x \<in> A. x \<sqsubseteq> z" |
23373 | 311 |
then have "z \<in> ?B" .. |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
312 |
with is_Inf show "sup \<sqsubseteq> z" .. |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
313 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
314 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
315 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
316 |
theorem Sup_Inf: "is_Sup {b. \<forall>a \<in> A. b \<sqsubseteq> a} inf \<Longrightarrow> is_Inf A inf" |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
317 |
proof - |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
318 |
assume "is_Sup {b. \<forall>a \<in> A. b \<sqsubseteq> a} inf" |
23373 | 319 |
then have "is_Inf (dual ` {b. \<forall>a \<in> A. dual a \<sqsubseteq> dual b}) (dual inf)" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
320 |
by (simp only: dual_Inf dual_leq) |
10834 | 321 |
also have "dual ` {b. \<forall>a \<in> A. dual a \<sqsubseteq> dual b} = {b'. \<forall>a' \<in> dual ` A. a' \<sqsubseteq> b'}" |
11265 | 322 |
by (auto iff: dual_ball dual_Collect simp add: image_Collect) (* FIXME !? *) |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
323 |
finally have "is_Inf \<dots> (dual inf)" . |
23373 | 324 |
then have "is_Sup (dual ` A) (dual inf)" |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
325 |
by (rule Inf_Sup) |
23373 | 326 |
then show ?thesis .. |
10157
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
327 |
qed |
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
328 |
|
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
wenzelm
parents:
diff
changeset
|
329 |
end |