| author | wenzelm | 
| Fri, 05 Dec 2008 00:23:37 +0100 | |
| changeset 28981 | c9cf71e161b9 | 
| parent 27193 | 740159cfbf0e | 
| child 35317 | d57da4abb47d | 
| 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 | ID: $Id$ | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 3 | Author: Markus Wenzel, TU Muenchen | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 6 | header {* Bounds *}
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 7 | |
| 16417 | 8 | theory Bounds imports Orders begin | 
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 9 | |
| 27193 | 10 | hide (open) const inf sup | 
| 22426 | 11 | |
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 12 | subsection {* Infimum and supremum *}
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 13 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 14 | text {*
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 15 | Given a partial order, we define infimum (greatest lower bound) and | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 16 |   supremum (least upper bound) wrt.\ @{text \<sqsubseteq>} for two and for any
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 17 | number of elements. | 
| 
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 | |
| 19736 | 20 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 21 | is_inf :: "'a::partial_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where | 
| 19736 | 22 | "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 | 23 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 24 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 25 | is_sup :: "'a::partial_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where | 
| 19736 | 26 | "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 | 27 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 28 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 29 | is_Inf :: "'a::partial_order set \<Rightarrow> 'a \<Rightarrow> bool" where | 
| 19736 | 30 | "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 | 31 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 32 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19736diff
changeset | 33 | is_Sup :: "'a::partial_order set \<Rightarrow> 'a \<Rightarrow> bool" where | 
| 19736 | 34 | "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 | 35 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 36 | text {*
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 37 | These definitions entail the following basic properties of boundary | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 38 | elements. | 
| 
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 | 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 | 42 | (\<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 | 43 | by (unfold is_inf_def) blast | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 44 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 45 | lemma is_inf_greatest [elim?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 46 | "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 | 47 | by (unfold is_inf_def) blast | 
| 
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 | lemma is_inf_lower [elim?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 50 | "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 | 51 | by (unfold is_inf_def) blast | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 54 | 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 | 55 | (\<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 | 56 | by (unfold is_sup_def) blast | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 57 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 58 | lemma is_sup_least [elim?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 59 | "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 | 60 | by (unfold is_sup_def) blast | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 61 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 62 | lemma is_sup_upper [elim?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 63 | "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 | 64 | by (unfold is_sup_def) blast | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 67 | 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 | 68 | (\<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 | 69 | by (unfold is_Inf_def) blast | 
| 
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 is_Inf_greatest [elim?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 72 | "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 | 73 | by (unfold is_Inf_def) blast | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 74 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 75 | lemma is_Inf_lower [dest?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 76 | "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 | 77 | by (unfold is_Inf_def) blast | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 80 | 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 | 81 | (\<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 | 82 | by (unfold is_Sup_def) blast | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 83 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 84 | lemma is_Sup_least [elim?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 85 | "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 | 86 | by (unfold is_Sup_def) blast | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 87 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 88 | lemma is_Sup_upper [dest?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 89 | "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 | 90 | by (unfold is_Sup_def) blast | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 93 | subsection {* Duality *}
 | 
| 
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 | text {*
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 96 | Infimum and supremum are dual to each other. | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 99 | theorem dual_inf [iff?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 100 | "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 | 101 | 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 | 102 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 103 | theorem dual_sup [iff?]: | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 104 | "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 | 105 | 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 | 106 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 107 | theorem dual_Inf [iff?]: | 
| 10834 | 108 | "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 | 109 | 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 | 110 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 111 | theorem dual_Sup [iff?]: | 
| 10834 | 112 | "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 | 113 | 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 | 114 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 115 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 116 | subsection {* Uniqueness *}
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 117 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 118 | text {*
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 119 | 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 | 120 | to anti-symmetry of the underlying relation. | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 123 | 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 | 124 | proof - | 
| 
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 | assume inf': "is_inf x y inf'" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 127 | show ?thesis | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 128 | proof (rule leq_antisym) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 129 | from inf' show "inf \<sqsubseteq> inf'" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 130 | proof (rule is_inf_greatest) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 131 | from inf show "inf \<sqsubseteq> x" .. | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 132 | from inf show "inf \<sqsubseteq> y" .. | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 133 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 134 | from inf show "inf' \<sqsubseteq> inf" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 135 | proof (rule is_inf_greatest) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 136 | from inf' show "inf' \<sqsubseteq> x" .. | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 137 | from inf' show "inf' \<sqsubseteq> y" .. | 
| 
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 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 141 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 142 | 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 | 143 | proof - | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 144 | 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 | 145 | have "dual sup = dual sup'" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 146 | proof (rule is_inf_uniq) | 
| 
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 | 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 | 149 | qed | 
| 23373 | 150 | then show "sup = sup'" .. | 
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 151 | qed | 
| 
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 | 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 | 154 | proof - | 
| 
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 | assume inf': "is_Inf A inf'" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 157 | show ?thesis | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 158 | proof (rule leq_antisym) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 159 | from inf' show "inf \<sqsubseteq> inf'" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 160 | proof (rule is_Inf_greatest) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 161 | fix x assume "x \<in> A" | 
| 23373 | 162 | with inf show "inf \<sqsubseteq> x" .. | 
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 163 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 164 | from inf show "inf' \<sqsubseteq> inf" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 165 | proof (rule is_Inf_greatest) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 166 | fix x assume "x \<in> A" | 
| 23373 | 167 | with inf' show "inf' \<sqsubseteq> x" .. | 
| 10157 
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 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 171 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 172 | 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 | 173 | proof - | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 174 | 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 | 175 | have "dual sup = dual sup'" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 176 | proof (rule is_Inf_uniq) | 
| 10834 | 177 | from sup show "is_Inf (dual ` A) (dual sup)" .. | 
| 178 | 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 | 179 | qed | 
| 23373 | 180 | then show "sup = sup'" .. | 
| 10157 
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 | |
| 
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 | subsection {* Related elements *}
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 185 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 186 | text {*
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 187 | The binary bound of related elements is either one of the argument. | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 190 | 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 | 191 | proof - | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 192 | assume "x \<sqsubseteq> y" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 193 | show ?thesis | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 194 | proof | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 195 | show "x \<sqsubseteq> x" .. | 
| 23393 | 196 | show "x \<sqsubseteq> y" by fact | 
| 197 | 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 | 198 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 199 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 200 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 201 | 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 | 202 | proof - | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 203 | assume "x \<sqsubseteq> y" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 204 | show ?thesis | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 205 | proof | 
| 23393 | 206 | show "x \<sqsubseteq> y" by fact | 
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 207 | show "y \<sqsubseteq> y" .. | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 208 | fix z assume "x \<sqsubseteq> z" and "y \<sqsubseteq> z" | 
| 23393 | 209 | show "y \<sqsubseteq> z" by fact | 
| 10157 
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 | qed | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 214 | subsection {* General versus binary bounds \label{sec:gen-bin-bounds} *}
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 215 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 216 | text {*
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 217 | General bounds of two-element sets coincide with binary bounds. | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 220 | 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 | 221 | proof - | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 222 |   let ?A = "{x, y}"
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 223 | show ?thesis | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 224 | proof | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 225 | assume is_Inf: "is_Inf ?A inf" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 226 | show "is_inf x y inf" | 
| 
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 | have "x \<in> ?A" by simp | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 229 | with is_Inf show "inf \<sqsubseteq> x" .. | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 230 | have "y \<in> ?A" by simp | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 231 | with is_Inf show "inf \<sqsubseteq> y" .. | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 232 | 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 | 233 | from is_Inf show "z \<sqsubseteq> inf" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 234 | proof (rule is_Inf_greatest) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 235 | fix a assume "a \<in> ?A" | 
| 23373 | 236 | then have "a = x \<or> a = y" by blast | 
| 237 | then show "z \<sqsubseteq> a" | |
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 238 | proof | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 239 | assume "a = x" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 240 | with zx show ?thesis by simp | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 241 | next | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 242 | assume "a = y" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 243 | with zy show ?thesis by simp | 
| 
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 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 247 | next | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 248 | assume is_inf: "is_inf x y inf" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 249 |     show "is_Inf {x, y} inf"
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 250 | proof | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 251 | fix a assume "a \<in> ?A" | 
| 23373 | 252 | then have "a = x \<or> a = y" by blast | 
| 253 | then show "inf \<sqsubseteq> a" | |
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 254 | proof | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 255 | assume "a = x" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 256 | also from is_inf have "inf \<sqsubseteq> x" .. | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 257 | finally show ?thesis . | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 258 | next | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 259 | assume "a = y" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 260 | also from is_inf have "inf \<sqsubseteq> y" .. | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 261 | finally show ?thesis . | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 262 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 263 | next | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 264 | 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 | 265 | from is_inf show "z \<sqsubseteq> inf" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 266 | proof (rule is_inf_greatest) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 267 | from z show "z \<sqsubseteq> x" by blast | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 268 | from z show "z \<sqsubseteq> y" by blast | 
| 
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 | qed | 
| 
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 | 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 | 275 | proof - | 
| 10834 | 276 |   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 | 277 | by (simp only: dual_Inf) | 
| 10834 | 278 |   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 | 279 | by simp | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 280 | 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 | 281 | by (rule is_Inf_binary) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 282 | also have "\<dots> = is_sup x y sup" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 283 | by (simp only: dual_inf) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 284 | finally show ?thesis . | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 287 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 288 | subsection {* Connecting general bounds \label{sec:connect-bounds} *}
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 289 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 290 | text {*
 | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 291 | 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 | 292 | 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 | 293 | 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 | 294 | holds as well; the dual statement holds as well. | 
| 
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 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 297 | 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 | 298 | proof - | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 299 |   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 | 300 | assume is_Inf: "is_Inf ?B sup" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 301 | show "is_Sup A sup" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 302 | proof | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 303 | fix x assume x: "x \<in> A" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 304 | from is_Inf show "x \<sqsubseteq> sup" | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 305 | proof (rule is_Inf_greatest) | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 306 | fix y assume "y \<in> ?B" | 
| 23373 | 307 | 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 | 308 | from this x show "x \<sqsubseteq> y" .. | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 309 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 310 | next | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 311 | fix z assume "\<forall>x \<in> A. x \<sqsubseteq> z" | 
| 23373 | 312 | then have "z \<in> ?B" .. | 
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 313 | with is_Inf show "sup \<sqsubseteq> z" .. | 
| 
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 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 316 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 317 | 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 | 318 | proof - | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 319 |   assume "is_Sup {b. \<forall>a \<in> A. b \<sqsubseteq> a} inf"
 | 
| 23373 | 320 |   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 | 321 | by (simp only: dual_Inf dual_leq) | 
| 10834 | 322 |   also have "dual ` {b. \<forall>a \<in> A. dual a \<sqsubseteq> dual b} = {b'. \<forall>a' \<in> dual ` A. a' \<sqsubseteq> b'}"
 | 
| 11265 | 323 | 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 | 324 | finally have "is_Inf \<dots> (dual inf)" . | 
| 23373 | 325 | then have "is_Sup (dual ` A) (dual inf)" | 
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 326 | by (rule Inf_Sup) | 
| 23373 | 327 | then show ?thesis .. | 
| 10157 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 328 | qed | 
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 329 | |
| 
6d3987f3aad9
* HOL/Lattice: fundamental concepts of lattice theory and order structures;
 wenzelm parents: diff
changeset | 330 | end |