| author | wenzelm | 
| Sun, 22 Mar 2015 12:45:34 +0100 | |
| changeset 59774 | d1b83f7ff6fe | 
| parent 59545 | 12a6088ed195 | 
| child 60758 | d8d85a8172b5 | 
| permissions | -rw-r--r-- | 
| 21249 | 1 | (* Title: HOL/Lattices.thy | 
| 2 | Author: Tobias Nipkow | |
| 3 | *) | |
| 4 | ||
| 58889 | 5 | section {* Abstract lattices *}
 | 
| 21249 | 6 | |
| 7 | theory Lattices | |
| 54555 | 8 | imports Groups | 
| 21249 | 9 | begin | 
| 10 | ||
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 11 | subsection {* Abstract semilattice *}
 | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 12 | |
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 13 | text {*
 | 
| 51487 | 14 | These locales provide a basic structure for interpretation into | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 15 | bigger structures; extensions require careful thinking, otherwise | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 16 | undesired effects may occur due to interpretation. | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 17 | *} | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 18 | |
| 51487 | 19 | no_notation times (infixl "*" 70) | 
| 20 | no_notation Groups.one ("1")
 | |
| 21 | ||
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 22 | locale semilattice = abel_semigroup + | 
| 51487 | 23 | assumes idem [simp]: "a * a = a" | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 24 | begin | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 25 | |
| 51487 | 26 | lemma left_idem [simp]: "a * (a * b) = a * b" | 
| 50615 | 27 | by (simp add: assoc [symmetric]) | 
| 28 | ||
| 51487 | 29 | lemma right_idem [simp]: "(a * b) * b = a * b" | 
| 50615 | 30 | by (simp add: assoc) | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 31 | |
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 32 | end | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 33 | |
| 51487 | 34 | locale semilattice_neutr = semilattice + comm_monoid | 
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 35 | |
| 51487 | 36 | locale semilattice_order = semilattice + | 
| 37 | fixes less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<preceq>" 50) | |
| 38 | and less :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<prec>" 50) | |
| 39 | assumes order_iff: "a \<preceq> b \<longleftrightarrow> a = a * b" | |
| 59545 
12a6088ed195
explicit equivalence for strict order on lattices
 haftmann parents: 
58889diff
changeset | 40 | and strict_order_iff: "a \<prec> b \<longleftrightarrow> a = a * b \<and> a \<noteq> b" | 
| 51487 | 41 | begin | 
| 42 | ||
| 43 | lemma orderI: | |
| 44 | "a = a * b \<Longrightarrow> a \<preceq> b" | |
| 45 | by (simp add: order_iff) | |
| 46 | ||
| 47 | lemma orderE: | |
| 48 | assumes "a \<preceq> b" | |
| 49 | obtains "a = a * b" | |
| 50 | using assms by (unfold order_iff) | |
| 51 | ||
| 52152 | 52 | sublocale ordering less_eq less | 
| 51487 | 53 | proof | 
| 54 | fix a b | |
| 55 | show "a \<prec> b \<longleftrightarrow> a \<preceq> b \<and> a \<noteq> b" | |
| 59545 
12a6088ed195
explicit equivalence for strict order on lattices
 haftmann parents: 
58889diff
changeset | 56 | by (simp add: order_iff strict_order_iff) | 
| 51487 | 57 | next | 
| 58 | fix a | |
| 59 | show "a \<preceq> a" | |
| 60 | by (simp add: order_iff) | |
| 61 | next | |
| 62 | fix a b | |
| 63 | assume "a \<preceq> b" "b \<preceq> a" | |
| 64 | then have "a = a * b" "a * b = b" | |
| 65 | by (simp_all add: order_iff commute) | |
| 66 | then show "a = b" by simp | |
| 67 | next | |
| 68 | fix a b c | |
| 69 | assume "a \<preceq> b" "b \<preceq> c" | |
| 70 | then have "a = a * b" "b = b * c" | |
| 71 | by (simp_all add: order_iff commute) | |
| 72 | then have "a = a * (b * c)" | |
| 73 | by simp | |
| 74 | then have "a = (a * b) * c" | |
| 75 | by (simp add: assoc) | |
| 76 | with `a = a * b` [symmetric] have "a = a * c" by simp | |
| 77 | then show "a \<preceq> c" by (rule orderI) | |
| 78 | qed | |
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 79 | |
| 51487 | 80 | lemma cobounded1 [simp]: | 
| 81 | "a * b \<preceq> a" | |
| 82 | by (simp add: order_iff commute) | |
| 83 | ||
| 84 | lemma cobounded2 [simp]: | |
| 85 | "a * b \<preceq> b" | |
| 86 | by (simp add: order_iff) | |
| 87 | ||
| 88 | lemma boundedI: | |
| 89 | assumes "a \<preceq> b" and "a \<preceq> c" | |
| 90 | shows "a \<preceq> b * c" | |
| 91 | proof (rule orderI) | |
| 92 | from assms obtain "a * b = a" and "a * c = a" by (auto elim!: orderE) | |
| 93 | then show "a = a * (b * c)" by (simp add: assoc [symmetric]) | |
| 94 | qed | |
| 95 | ||
| 96 | lemma boundedE: | |
| 97 | assumes "a \<preceq> b * c" | |
| 98 | obtains "a \<preceq> b" and "a \<preceq> c" | |
| 99 | using assms by (blast intro: trans cobounded1 cobounded2) | |
| 100 | ||
| 54859 | 101 | lemma bounded_iff [simp]: | 
| 51487 | 102 | "a \<preceq> b * c \<longleftrightarrow> a \<preceq> b \<and> a \<preceq> c" | 
| 103 | by (blast intro: boundedI elim: boundedE) | |
| 104 | ||
| 105 | lemma strict_boundedE: | |
| 106 | assumes "a \<prec> b * c" | |
| 107 | obtains "a \<prec> b" and "a \<prec> c" | |
| 54859 | 108 | using assms by (auto simp add: commute strict_iff_order elim: orderE intro!: that)+ | 
| 51487 | 109 | |
| 110 | lemma coboundedI1: | |
| 111 | "a \<preceq> c \<Longrightarrow> a * b \<preceq> c" | |
| 112 | by (rule trans) auto | |
| 113 | ||
| 114 | lemma coboundedI2: | |
| 115 | "b \<preceq> c \<Longrightarrow> a * b \<preceq> c" | |
| 116 | by (rule trans) auto | |
| 117 | ||
| 54858 | 118 | lemma strict_coboundedI1: | 
| 119 | "a \<prec> c \<Longrightarrow> a * b \<prec> c" | |
| 120 | using irrefl | |
| 121 | by (auto intro: not_eq_order_implies_strict coboundedI1 strict_implies_order elim: strict_boundedE) | |
| 122 | ||
| 123 | lemma strict_coboundedI2: | |
| 124 | "b \<prec> c \<Longrightarrow> a * b \<prec> c" | |
| 125 | using strict_coboundedI1 [of b c a] by (simp add: commute) | |
| 126 | ||
| 51487 | 127 | lemma mono: "a \<preceq> c \<Longrightarrow> b \<preceq> d \<Longrightarrow> a * b \<preceq> c * d" | 
| 128 | by (blast intro: boundedI coboundedI1 coboundedI2) | |
| 129 | ||
| 130 | lemma absorb1: "a \<preceq> b \<Longrightarrow> a * b = a" | |
| 54859 | 131 | by (rule antisym) (auto simp add: refl) | 
| 51487 | 132 | |
| 133 | lemma absorb2: "b \<preceq> a \<Longrightarrow> a * b = b" | |
| 54859 | 134 | by (rule antisym) (auto simp add: refl) | 
| 54858 | 135 | |
| 136 | lemma absorb_iff1: "a \<preceq> b \<longleftrightarrow> a * b = a" | |
| 137 | using order_iff by auto | |
| 138 | ||
| 139 | lemma absorb_iff2: "b \<preceq> a \<longleftrightarrow> a * b = b" | |
| 140 | using order_iff by (auto simp add: commute) | |
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 141 | |
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 142 | end | 
| 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 143 | |
| 51487 | 144 | locale semilattice_neutr_order = semilattice_neutr + semilattice_order | 
| 52152 | 145 | begin | 
| 51487 | 146 | |
| 52152 | 147 | sublocale ordering_top less_eq less 1 | 
| 51487 | 148 | by default (simp add: order_iff) | 
| 149 | ||
| 52152 | 150 | end | 
| 151 | ||
| 51487 | 152 | notation times (infixl "*" 70) | 
| 153 | notation Groups.one ("1")
 | |
| 154 | ||
| 35301 
90e42f9ba4d1
distributed theory Algebras to theories Groups and Lattices
 haftmann parents: 
35121diff
changeset | 155 | |
| 46691 | 156 | subsection {* Syntactic infimum and supremum operations *}
 | 
| 41082 | 157 | |
| 44845 | 158 | class inf = | 
| 159 | fixes inf :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<sqinter>" 70) | |
| 25206 | 160 | |
| 44845 | 161 | class sup = | 
| 162 | fixes sup :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "\<squnion>" 65) | |
| 163 | ||
| 46691 | 164 | |
| 165 | subsection {* Concrete lattices *}
 | |
| 166 | ||
| 167 | notation | |
| 168 | less_eq (infix "\<sqsubseteq>" 50) and | |
| 169 | less (infix "\<sqsubset>" 50) | |
| 170 | ||
| 44845 | 171 | class semilattice_inf = order + inf + | 
| 22737 | 172 | assumes inf_le1 [simp]: "x \<sqinter> y \<sqsubseteq> x" | 
| 173 | and inf_le2 [simp]: "x \<sqinter> y \<sqsubseteq> y" | |
| 21733 | 174 | and inf_greatest: "x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<sqinter> z" | 
| 21249 | 175 | |
| 44845 | 176 | class semilattice_sup = order + sup + | 
| 22737 | 177 | assumes sup_ge1 [simp]: "x \<sqsubseteq> x \<squnion> y" | 
| 178 | and sup_ge2 [simp]: "y \<sqsubseteq> x \<squnion> y" | |
| 21733 | 179 | and sup_least: "y \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> x \<Longrightarrow> y \<squnion> z \<sqsubseteq> x" | 
| 26014 | 180 | begin | 
| 181 | ||
| 182 | text {* Dual lattice *}
 | |
| 183 | ||
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 184 | lemma dual_semilattice: | 
| 44845 | 185 | "class.semilattice_inf sup greater_eq greater" | 
| 36635 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 haftmann parents: 
36352diff
changeset | 186 | by (rule class.semilattice_inf.intro, rule dual_order) | 
| 27682 | 187 | (unfold_locales, simp_all add: sup_least) | 
| 26014 | 188 | |
| 189 | end | |
| 21249 | 190 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 191 | class lattice = semilattice_inf + semilattice_sup | 
| 21249 | 192 | |
| 25382 | 193 | |
| 28562 | 194 | subsubsection {* Intro and elim rules*}
 | 
| 21733 | 195 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 196 | context semilattice_inf | 
| 21733 | 197 | begin | 
| 21249 | 198 | |
| 32064 | 199 | lemma le_infI1: | 
| 200 | "a \<sqsubseteq> x \<Longrightarrow> a \<sqinter> b \<sqsubseteq> x" | |
| 201 | by (rule order_trans) auto | |
| 21249 | 202 | |
| 32064 | 203 | lemma le_infI2: | 
| 204 | "b \<sqsubseteq> x \<Longrightarrow> a \<sqinter> b \<sqsubseteq> x" | |
| 205 | by (rule order_trans) auto | |
| 21733 | 206 | |
| 32064 | 207 | lemma le_infI: "x \<sqsubseteq> a \<Longrightarrow> x \<sqsubseteq> b \<Longrightarrow> x \<sqsubseteq> a \<sqinter> b" | 
| 54857 | 208 | by (fact inf_greatest) (* FIXME: duplicate lemma *) | 
| 21249 | 209 | |
| 32064 | 210 | lemma le_infE: "x \<sqsubseteq> a \<sqinter> b \<Longrightarrow> (x \<sqsubseteq> a \<Longrightarrow> x \<sqsubseteq> b \<Longrightarrow> P) \<Longrightarrow> P" | 
| 36008 | 211 | by (blast intro: order_trans inf_le1 inf_le2) | 
| 21249 | 212 | |
| 54859 | 213 | lemma le_inf_iff: | 
| 32064 | 214 | "x \<sqsubseteq> y \<sqinter> z \<longleftrightarrow> x \<sqsubseteq> y \<and> x \<sqsubseteq> z" | 
| 215 | by (blast intro: le_infI elim: le_infE) | |
| 21733 | 216 | |
| 32064 | 217 | lemma le_iff_inf: | 
| 218 | "x \<sqsubseteq> y \<longleftrightarrow> x \<sqinter> y = x" | |
| 54859 | 219 | by (auto intro: le_infI1 antisym dest: eq_iff [THEN iffD1] simp add: le_inf_iff) | 
| 21249 | 220 | |
| 43753 | 221 | lemma inf_mono: "a \<sqsubseteq> c \<Longrightarrow> b \<sqsubseteq> d \<Longrightarrow> a \<sqinter> b \<sqsubseteq> c \<sqinter> d" | 
| 36008 | 222 | by (fast intro: inf_greatest le_infI1 le_infI2) | 
| 223 | ||
| 25206 | 224 | lemma mono_inf: | 
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 225 | fixes f :: "'a \<Rightarrow> 'b\<Colon>semilattice_inf" | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 226 | shows "mono f \<Longrightarrow> f (A \<sqinter> B) \<sqsubseteq> f A \<sqinter> f B" | 
| 25206 | 227 | by (auto simp add: mono_def intro: Lattices.inf_greatest) | 
| 21733 | 228 | |
| 25206 | 229 | end | 
| 21733 | 230 | |
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 231 | context semilattice_sup | 
| 21733 | 232 | begin | 
| 21249 | 233 | |
| 32064 | 234 | lemma le_supI1: | 
| 235 | "x \<sqsubseteq> a \<Longrightarrow> x \<sqsubseteq> a \<squnion> b" | |
| 25062 | 236 | by (rule order_trans) auto | 
| 21249 | 237 | |
| 32064 | 238 | lemma le_supI2: | 
| 239 | "x \<sqsubseteq> b \<Longrightarrow> x \<sqsubseteq> a \<squnion> b" | |
| 25062 | 240 | by (rule order_trans) auto | 
| 21733 | 241 | |
| 32064 | 242 | lemma le_supI: | 
| 243 | "a \<sqsubseteq> x \<Longrightarrow> b \<sqsubseteq> x \<Longrightarrow> a \<squnion> b \<sqsubseteq> x" | |
| 54857 | 244 | by (fact sup_least) (* FIXME: duplicate lemma *) | 
| 21249 | 245 | |
| 32064 | 246 | lemma le_supE: | 
| 247 | "a \<squnion> b \<sqsubseteq> x \<Longrightarrow> (a \<sqsubseteq> x \<Longrightarrow> b \<sqsubseteq> x \<Longrightarrow> P) \<Longrightarrow> P" | |
| 36008 | 248 | by (blast intro: order_trans sup_ge1 sup_ge2) | 
| 22422 
ee19cdb07528
stepping towards uniform lattice theory development in HOL
 haftmann parents: 
22384diff
changeset | 249 | |
| 54859 | 250 | lemma le_sup_iff: | 
| 32064 | 251 | "x \<squnion> y \<sqsubseteq> z \<longleftrightarrow> x \<sqsubseteq> z \<and> y \<sqsubseteq> z" | 
| 252 | by (blast intro: le_supI elim: le_supE) | |
| 21733 | 253 | |
| 32064 | 254 | lemma le_iff_sup: | 
| 255 | "x \<sqsubseteq> y \<longleftrightarrow> x \<squnion> y = y" | |
| 54859 | 256 | by (auto intro: le_supI2 antisym dest: eq_iff [THEN iffD1] simp add: le_sup_iff) | 
| 21734 | 257 | |
| 43753 | 258 | lemma sup_mono: "a \<sqsubseteq> c \<Longrightarrow> b \<sqsubseteq> d \<Longrightarrow> a \<squnion> b \<sqsubseteq> c \<squnion> d" | 
| 36008 | 259 | by (fast intro: sup_least le_supI1 le_supI2) | 
| 260 | ||
| 25206 | 261 | lemma mono_sup: | 
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 262 | fixes f :: "'a \<Rightarrow> 'b\<Colon>semilattice_sup" | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 263 | shows "mono f \<Longrightarrow> f A \<squnion> f B \<sqsubseteq> f (A \<squnion> B)" | 
| 25206 | 264 | by (auto simp add: mono_def intro: Lattices.sup_least) | 
| 21733 | 265 | |
| 25206 | 266 | end | 
| 23878 | 267 | |
| 21733 | 268 | |
| 32064 | 269 | subsubsection {* Equational laws *}
 | 
| 21249 | 270 | |
| 52152 | 271 | context semilattice_inf | 
| 272 | begin | |
| 273 | ||
| 274 | sublocale inf!: semilattice inf | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 275 | proof | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 276 | fix a b c | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 277 | show "(a \<sqinter> b) \<sqinter> c = a \<sqinter> (b \<sqinter> c)" | 
| 54859 | 278 | by (rule antisym) (auto intro: le_infI1 le_infI2 simp add: le_inf_iff) | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 279 | show "a \<sqinter> b = b \<sqinter> a" | 
| 54859 | 280 | by (rule antisym) (auto simp add: le_inf_iff) | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 281 | show "a \<sqinter> a = a" | 
| 54859 | 282 | by (rule antisym) (auto simp add: le_inf_iff) | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 283 | qed | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 284 | |
| 52152 | 285 | sublocale inf!: semilattice_order inf less_eq less | 
| 51487 | 286 | by default (auto simp add: le_iff_inf less_le) | 
| 287 | ||
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 288 | lemma inf_assoc: "(x \<sqinter> y) \<sqinter> z = x \<sqinter> (y \<sqinter> z)" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 289 | by (fact inf.assoc) | 
| 21733 | 290 | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 291 | lemma inf_commute: "(x \<sqinter> y) = (y \<sqinter> x)" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 292 | by (fact inf.commute) | 
| 21733 | 293 | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 294 | lemma inf_left_commute: "x \<sqinter> (y \<sqinter> z) = y \<sqinter> (x \<sqinter> z)" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 295 | by (fact inf.left_commute) | 
| 21733 | 296 | |
| 44921 | 297 | lemma inf_idem: "x \<sqinter> x = x" | 
| 298 | by (fact inf.idem) (* already simp *) | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 299 | |
| 50615 | 300 | lemma inf_left_idem: "x \<sqinter> (x \<sqinter> y) = x \<sqinter> y" | 
| 301 | by (fact inf.left_idem) (* already simp *) | |
| 302 | ||
| 303 | lemma inf_right_idem: "(x \<sqinter> y) \<sqinter> y = x \<sqinter> y" | |
| 304 | by (fact inf.right_idem) (* already simp *) | |
| 21733 | 305 | |
| 32642 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 haftmann parents: 
32568diff
changeset | 306 | lemma inf_absorb1: "x \<sqsubseteq> y \<Longrightarrow> x \<sqinter> y = x" | 
| 32064 | 307 | by (rule antisym) auto | 
| 21733 | 308 | |
| 32642 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 haftmann parents: 
32568diff
changeset | 309 | lemma inf_absorb2: "y \<sqsubseteq> x \<Longrightarrow> x \<sqinter> y = y" | 
| 32064 | 310 | by (rule antisym) auto | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 311 | |
| 32064 | 312 | lemmas inf_aci = inf_commute inf_assoc inf_left_commute inf_left_idem | 
| 21733 | 313 | |
| 314 | end | |
| 315 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 316 | context semilattice_sup | 
| 21733 | 317 | begin | 
| 21249 | 318 | |
| 52152 | 319 | sublocale sup!: semilattice sup | 
| 320 | proof | |
| 321 | fix a b c | |
| 322 | show "(a \<squnion> b) \<squnion> c = a \<squnion> (b \<squnion> c)" | |
| 54859 | 323 | by (rule antisym) (auto intro: le_supI1 le_supI2 simp add: le_sup_iff) | 
| 52152 | 324 | show "a \<squnion> b = b \<squnion> a" | 
| 54859 | 325 | by (rule antisym) (auto simp add: le_sup_iff) | 
| 52152 | 326 | show "a \<squnion> a = a" | 
| 54859 | 327 | by (rule antisym) (auto simp add: le_sup_iff) | 
| 52152 | 328 | qed | 
| 329 | ||
| 330 | sublocale sup!: semilattice_order sup greater_eq greater | |
| 331 | by default (auto simp add: le_iff_sup sup.commute less_le) | |
| 332 | ||
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 333 | lemma sup_assoc: "(x \<squnion> y) \<squnion> z = x \<squnion> (y \<squnion> z)" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 334 | by (fact sup.assoc) | 
| 21733 | 335 | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 336 | lemma sup_commute: "(x \<squnion> y) = (y \<squnion> x)" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 337 | by (fact sup.commute) | 
| 21733 | 338 | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 339 | lemma sup_left_commute: "x \<squnion> (y \<squnion> z) = y \<squnion> (x \<squnion> z)" | 
| 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 340 | by (fact sup.left_commute) | 
| 21733 | 341 | |
| 44921 | 342 | lemma sup_idem: "x \<squnion> x = x" | 
| 343 | by (fact sup.idem) (* already simp *) | |
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 344 | |
| 44918 | 345 | lemma sup_left_idem [simp]: "x \<squnion> (x \<squnion> y) = x \<squnion> y" | 
| 34973 
ae634fad947e
dropped mk_left_commute; use interpretation of locale abel_semigroup instead
 haftmann parents: 
34209diff
changeset | 346 | by (fact sup.left_idem) | 
| 21733 | 347 | |
| 32642 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 haftmann parents: 
32568diff
changeset | 348 | lemma sup_absorb1: "y \<sqsubseteq> x \<Longrightarrow> x \<squnion> y = x" | 
| 32064 | 349 | by (rule antisym) auto | 
| 21733 | 350 | |
| 32642 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 haftmann parents: 
32568diff
changeset | 351 | lemma sup_absorb2: "x \<sqsubseteq> y \<Longrightarrow> x \<squnion> y = y" | 
| 32064 | 352 | by (rule antisym) auto | 
| 21249 | 353 | |
| 32064 | 354 | lemmas sup_aci = sup_commute sup_assoc sup_left_commute sup_left_idem | 
| 21733 | 355 | |
| 356 | end | |
| 21249 | 357 | |
| 21733 | 358 | context lattice | 
| 359 | begin | |
| 360 | ||
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 361 | lemma dual_lattice: | 
| 44845 | 362 | "class.lattice sup (op \<ge>) (op >) inf" | 
| 36635 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 haftmann parents: 
36352diff
changeset | 363 | by (rule class.lattice.intro, rule dual_semilattice, rule class.semilattice_sup.intro, rule dual_order) | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 364 | (unfold_locales, auto) | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 365 | |
| 44918 | 366 | lemma inf_sup_absorb [simp]: "x \<sqinter> (x \<squnion> y) = x" | 
| 25102 
db3e412c4cb1
antisymmetry not a default intro rule any longer
 haftmann parents: 
25062diff
changeset | 367 | by (blast intro: antisym inf_le1 inf_greatest sup_ge1) | 
| 21733 | 368 | |
| 44918 | 369 | lemma sup_inf_absorb [simp]: "x \<squnion> (x \<sqinter> y) = x" | 
| 25102 
db3e412c4cb1
antisymmetry not a default intro rule any longer
 haftmann parents: 
25062diff
changeset | 370 | by (blast intro: antisym sup_ge1 sup_least inf_le1) | 
| 21733 | 371 | |
| 32064 | 372 | lemmas inf_sup_aci = inf_aci sup_aci | 
| 21734 | 373 | |
| 22454 | 374 | lemmas inf_sup_ord = inf_le1 inf_le2 sup_ge1 sup_ge2 | 
| 375 | ||
| 21734 | 376 | text{* Towards distributivity *}
 | 
| 21249 | 377 | |
| 21734 | 378 | lemma distrib_sup_le: "x \<squnion> (y \<sqinter> z) \<sqsubseteq> (x \<squnion> y) \<sqinter> (x \<squnion> z)" | 
| 32064 | 379 | by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) | 
| 21734 | 380 | |
| 381 | lemma distrib_inf_le: "(x \<sqinter> y) \<squnion> (x \<sqinter> z) \<sqsubseteq> x \<sqinter> (y \<squnion> z)" | |
| 32064 | 382 | by (auto intro: le_infI1 le_infI2 le_supI1 le_supI2) | 
| 21734 | 383 | |
| 384 | text{* If you have one of them, you have them all. *}
 | |
| 21249 | 385 | |
| 21733 | 386 | lemma distrib_imp1: | 
| 21249 | 387 | assumes D: "!!x y z. x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" | 
| 388 | shows "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" | |
| 389 | proof- | |
| 44918 | 390 | have "x \<squnion> (y \<sqinter> z) = (x \<squnion> (x \<sqinter> z)) \<squnion> (y \<sqinter> z)" by simp | 
| 391 | also have "\<dots> = x \<squnion> (z \<sqinter> (x \<squnion> y))" | |
| 392 | by (simp add: D inf_commute sup_assoc del: sup_inf_absorb) | |
| 21249 | 393 | also have "\<dots> = ((x \<squnion> y) \<sqinter> x) \<squnion> ((x \<squnion> y) \<sqinter> z)" | 
| 44919 | 394 | by(simp add: inf_commute) | 
| 21249 | 395 | also have "\<dots> = (x \<squnion> y) \<sqinter> (x \<squnion> z)" by(simp add:D) | 
| 396 | finally show ?thesis . | |
| 397 | qed | |
| 398 | ||
| 21733 | 399 | lemma distrib_imp2: | 
| 21249 | 400 | assumes D: "!!x y z. x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" | 
| 401 | shows "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" | |
| 402 | proof- | |
| 44918 | 403 | have "x \<sqinter> (y \<squnion> z) = (x \<sqinter> (x \<squnion> z)) \<sqinter> (y \<squnion> z)" by simp | 
| 404 | also have "\<dots> = x \<sqinter> (z \<squnion> (x \<sqinter> y))" | |
| 405 | by (simp add: D sup_commute inf_assoc del: inf_sup_absorb) | |
| 21249 | 406 | also have "\<dots> = ((x \<sqinter> y) \<squnion> x) \<sqinter> ((x \<sqinter> y) \<squnion> z)" | 
| 44919 | 407 | by(simp add: sup_commute) | 
| 21249 | 408 | also have "\<dots> = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" by(simp add:D) | 
| 409 | finally show ?thesis . | |
| 410 | qed | |
| 411 | ||
| 21733 | 412 | end | 
| 21249 | 413 | |
| 32568 | 414 | subsubsection {* Strict order *}
 | 
| 415 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 416 | context semilattice_inf | 
| 32568 | 417 | begin | 
| 418 | ||
| 419 | lemma less_infI1: | |
| 420 | "a \<sqsubset> x \<Longrightarrow> a \<sqinter> b \<sqsubset> x" | |
| 32642 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 haftmann parents: 
32568diff
changeset | 421 | by (auto simp add: less_le inf_absorb1 intro: le_infI1) | 
| 32568 | 422 | |
| 423 | lemma less_infI2: | |
| 424 | "b \<sqsubset> x \<Longrightarrow> a \<sqinter> b \<sqsubset> x" | |
| 32642 
026e7c6a6d08
be more cautious wrt. simp rules: inf_absorb1, inf_absorb2, sup_absorb1, sup_absorb2 are no simp rules by default any longer
 haftmann parents: 
32568diff
changeset | 425 | by (auto simp add: less_le inf_absorb2 intro: le_infI2) | 
| 32568 | 426 | |
| 427 | end | |
| 428 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 429 | context semilattice_sup | 
| 32568 | 430 | begin | 
| 431 | ||
| 432 | lemma less_supI1: | |
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 433 | "x \<sqsubset> a \<Longrightarrow> x \<sqsubset> a \<squnion> b" | 
| 44921 | 434 | using dual_semilattice | 
| 435 | by (rule semilattice_inf.less_infI1) | |
| 32568 | 436 | |
| 437 | lemma less_supI2: | |
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 438 | "x \<sqsubset> b \<Longrightarrow> x \<sqsubset> a \<squnion> b" | 
| 44921 | 439 | using dual_semilattice | 
| 440 | by (rule semilattice_inf.less_infI2) | |
| 32568 | 441 | |
| 442 | end | |
| 443 | ||
| 21249 | 444 | |
| 24164 | 445 | subsection {* Distributive lattices *}
 | 
| 21249 | 446 | |
| 22454 | 447 | class distrib_lattice = lattice + | 
| 21249 | 448 | assumes sup_inf_distrib1: "x \<squnion> (y \<sqinter> z) = (x \<squnion> y) \<sqinter> (x \<squnion> z)" | 
| 449 | ||
| 21733 | 450 | context distrib_lattice | 
| 451 | begin | |
| 452 | ||
| 453 | lemma sup_inf_distrib2: | |
| 44921 | 454 | "(y \<sqinter> z) \<squnion> x = (y \<squnion> x) \<sqinter> (z \<squnion> x)" | 
| 455 | by (simp add: sup_commute sup_inf_distrib1) | |
| 21249 | 456 | |
| 21733 | 457 | lemma inf_sup_distrib1: | 
| 44921 | 458 | "x \<sqinter> (y \<squnion> z) = (x \<sqinter> y) \<squnion> (x \<sqinter> z)" | 
| 459 | by (rule distrib_imp2 [OF sup_inf_distrib1]) | |
| 21249 | 460 | |
| 21733 | 461 | lemma inf_sup_distrib2: | 
| 44921 | 462 | "(y \<squnion> z) \<sqinter> x = (y \<sqinter> x) \<squnion> (z \<sqinter> x)" | 
| 463 | by (simp add: inf_commute inf_sup_distrib1) | |
| 21249 | 464 | |
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 465 | lemma dual_distrib_lattice: | 
| 44845 | 466 | "class.distrib_lattice sup (op \<ge>) (op >) inf" | 
| 36635 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 haftmann parents: 
36352diff
changeset | 467 | by (rule class.distrib_lattice.intro, rule dual_lattice) | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 468 | (unfold_locales, fact inf_sup_distrib1) | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 469 | |
| 36008 | 470 | lemmas sup_inf_distrib = | 
| 471 | sup_inf_distrib1 sup_inf_distrib2 | |
| 472 | ||
| 473 | lemmas inf_sup_distrib = | |
| 474 | inf_sup_distrib1 inf_sup_distrib2 | |
| 475 | ||
| 21733 | 476 | lemmas distrib = | 
| 21249 | 477 | sup_inf_distrib1 sup_inf_distrib2 inf_sup_distrib1 inf_sup_distrib2 | 
| 478 | ||
| 21733 | 479 | end | 
| 480 | ||
| 21249 | 481 | |
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 482 | subsection {* Bounded lattices and boolean algebras *}
 | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 483 | |
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
52152diff
changeset | 484 | class bounded_semilattice_inf_top = semilattice_inf + order_top | 
| 52152 | 485 | begin | 
| 51487 | 486 | |
| 52152 | 487 | sublocale inf_top!: semilattice_neutr inf top | 
| 51546 
2e26df807dc7
more uniform style for interpretation and sublocale declarations
 haftmann parents: 
51540diff
changeset | 488 | + inf_top!: semilattice_neutr_order inf top less_eq less | 
| 51487 | 489 | proof | 
| 490 | fix x | |
| 491 | show "x \<sqinter> \<top> = x" | |
| 492 | by (rule inf_absorb1) simp | |
| 493 | qed | |
| 494 | ||
| 52152 | 495 | end | 
| 51487 | 496 | |
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
52152diff
changeset | 497 | class bounded_semilattice_sup_bot = semilattice_sup + order_bot | 
| 52152 | 498 | begin | 
| 499 | ||
| 500 | sublocale sup_bot!: semilattice_neutr sup bot | |
| 51546 
2e26df807dc7
more uniform style for interpretation and sublocale declarations
 haftmann parents: 
51540diff
changeset | 501 | + sup_bot!: semilattice_neutr_order sup bot greater_eq greater | 
| 51487 | 502 | proof | 
| 503 | fix x | |
| 504 | show "x \<squnion> \<bottom> = x" | |
| 505 | by (rule sup_absorb1) simp | |
| 506 | qed | |
| 507 | ||
| 52152 | 508 | end | 
| 509 | ||
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
52152diff
changeset | 510 | class bounded_lattice_bot = lattice + order_bot | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 511 | begin | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 512 | |
| 51487 | 513 | subclass bounded_semilattice_sup_bot .. | 
| 514 | ||
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 515 | lemma inf_bot_left [simp]: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 516 | "\<bottom> \<sqinter> x = \<bottom>" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 517 | by (rule inf_absorb1) simp | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 518 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 519 | lemma inf_bot_right [simp]: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 520 | "x \<sqinter> \<bottom> = \<bottom>" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 521 | by (rule inf_absorb2) simp | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 522 | |
| 51487 | 523 | lemma sup_bot_left: | 
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 524 | "\<bottom> \<squnion> x = x" | 
| 51487 | 525 | by (fact sup_bot.left_neutral) | 
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 526 | |
| 51487 | 527 | lemma sup_bot_right: | 
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 528 | "x \<squnion> \<bottom> = x" | 
| 51487 | 529 | by (fact sup_bot.right_neutral) | 
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 530 | |
| 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 531 | lemma sup_eq_bot_iff [simp]: | 
| 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 532 | "x \<squnion> y = \<bottom> \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>" | 
| 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 533 | by (simp add: eq_iff) | 
| 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 534 | |
| 51593 | 535 | lemma bot_eq_sup_iff [simp]: | 
| 536 | "\<bottom> = x \<squnion> y \<longleftrightarrow> x = \<bottom> \<and> y = \<bottom>" | |
| 537 | by (simp add: eq_iff) | |
| 538 | ||
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 539 | end | 
| 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 540 | |
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
52152diff
changeset | 541 | class bounded_lattice_top = lattice + order_top | 
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 542 | begin | 
| 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 543 | |
| 51487 | 544 | subclass bounded_semilattice_inf_top .. | 
| 545 | ||
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 546 | lemma sup_top_left [simp]: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 547 | "\<top> \<squnion> x = \<top>" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 548 | by (rule sup_absorb1) simp | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 549 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 550 | lemma sup_top_right [simp]: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 551 | "x \<squnion> \<top> = \<top>" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 552 | by (rule sup_absorb2) simp | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 553 | |
| 51487 | 554 | lemma inf_top_left: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 555 | "\<top> \<sqinter> x = x" | 
| 51487 | 556 | by (fact inf_top.left_neutral) | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 557 | |
| 51487 | 558 | lemma inf_top_right: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 559 | "x \<sqinter> \<top> = x" | 
| 51487 | 560 | by (fact inf_top.right_neutral) | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 561 | |
| 36008 | 562 | lemma inf_eq_top_iff [simp]: | 
| 563 | "x \<sqinter> y = \<top> \<longleftrightarrow> x = \<top> \<and> y = \<top>" | |
| 564 | by (simp add: eq_iff) | |
| 32568 | 565 | |
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 566 | end | 
| 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 567 | |
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
52152diff
changeset | 568 | class bounded_lattice = lattice + order_bot + order_top | 
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 569 | begin | 
| 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 570 | |
| 51487 | 571 | subclass bounded_lattice_bot .. | 
| 572 | subclass bounded_lattice_top .. | |
| 573 | ||
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 574 | lemma dual_bounded_lattice: | 
| 44845 | 575 | "class.bounded_lattice sup greater_eq greater inf \<top> \<bottom>" | 
| 36352 
f71978e47cd5
add bounded_lattice_bot and bounded_lattice_top type classes
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
36096diff
changeset | 576 | by unfold_locales (auto simp add: less_le_not_le) | 
| 32568 | 577 | |
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 578 | end | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 579 | |
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 580 | class boolean_algebra = distrib_lattice + bounded_lattice + minus + uminus + | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 581 | assumes inf_compl_bot: "x \<sqinter> - x = \<bottom>" | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 582 | and sup_compl_top: "x \<squnion> - x = \<top>" | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 583 | assumes diff_eq: "x - y = x \<sqinter> - y" | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 584 | begin | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 585 | |
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 586 | lemma dual_boolean_algebra: | 
| 44845 | 587 | "class.boolean_algebra (\<lambda>x y. x \<squnion> - y) uminus sup greater_eq greater inf \<top> \<bottom>" | 
| 36635 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 haftmann parents: 
36352diff
changeset | 588 | by (rule class.boolean_algebra.intro, rule dual_bounded_lattice, rule dual_distrib_lattice) | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 589 | (unfold_locales, auto simp add: inf_compl_bot sup_compl_top diff_eq) | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 590 | |
| 44918 | 591 | lemma compl_inf_bot [simp]: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 592 | "- x \<sqinter> x = \<bottom>" | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 593 | by (simp add: inf_commute inf_compl_bot) | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 594 | |
| 44918 | 595 | lemma compl_sup_top [simp]: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 596 | "- x \<squnion> x = \<top>" | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 597 | by (simp add: sup_commute sup_compl_top) | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 598 | |
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 599 | lemma compl_unique: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 600 | assumes "x \<sqinter> y = \<bottom>" | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 601 | and "x \<squnion> y = \<top>" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 602 | shows "- x = y" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 603 | proof - | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 604 | have "(x \<sqinter> - x) \<squnion> (- x \<sqinter> y) = (x \<sqinter> y) \<squnion> (- x \<sqinter> y)" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 605 | using inf_compl_bot assms(1) by simp | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 606 | then have "(- x \<sqinter> x) \<squnion> (- x \<sqinter> y) = (y \<sqinter> x) \<squnion> (y \<sqinter> - x)" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 607 | by (simp add: inf_commute) | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 608 | then have "- x \<sqinter> (x \<squnion> y) = y \<sqinter> (x \<squnion> - x)" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 609 | by (simp add: inf_sup_distrib1) | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 610 | then have "- x \<sqinter> \<top> = y \<sqinter> \<top>" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 611 | using sup_compl_top assms(2) by simp | 
| 34209 | 612 | then show "- x = y" by simp | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 613 | qed | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 614 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 615 | lemma double_compl [simp]: | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 616 | "- (- x) = x" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 617 | using compl_inf_bot compl_sup_top by (rule compl_unique) | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 618 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 619 | lemma compl_eq_compl_iff [simp]: | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 620 | "- x = - y \<longleftrightarrow> x = y" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 621 | proof | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 622 | assume "- x = - y" | 
| 36008 | 623 | then have "- (- x) = - (- y)" by (rule arg_cong) | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 624 | then show "x = y" by simp | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 625 | next | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 626 | assume "x = y" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 627 | then show "- x = - y" by simp | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 628 | qed | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 629 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 630 | lemma compl_bot_eq [simp]: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 631 | "- \<bottom> = \<top>" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 632 | proof - | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 633 | from sup_compl_top have "\<bottom> \<squnion> - \<bottom> = \<top>" . | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 634 | then show ?thesis by simp | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 635 | qed | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 636 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 637 | lemma compl_top_eq [simp]: | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 638 | "- \<top> = \<bottom>" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 639 | proof - | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 640 | from inf_compl_bot have "\<top> \<sqinter> - \<top> = \<bottom>" . | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 641 | then show ?thesis by simp | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 642 | qed | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 643 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 644 | lemma compl_inf [simp]: | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 645 | "- (x \<sqinter> y) = - x \<squnion> - y" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 646 | proof (rule compl_unique) | 
| 36008 | 647 | have "(x \<sqinter> y) \<sqinter> (- x \<squnion> - y) = (y \<sqinter> (x \<sqinter> - x)) \<squnion> (x \<sqinter> (y \<sqinter> - y))" | 
| 648 | by (simp only: inf_sup_distrib inf_aci) | |
| 649 | then show "(x \<sqinter> y) \<sqinter> (- x \<squnion> - y) = \<bottom>" | |
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 650 | by (simp add: inf_compl_bot) | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 651 | next | 
| 36008 | 652 | have "(x \<sqinter> y) \<squnion> (- x \<squnion> - y) = (- y \<squnion> (x \<squnion> - x)) \<sqinter> (- x \<squnion> (y \<squnion> - y))" | 
| 653 | by (simp only: sup_inf_distrib sup_aci) | |
| 654 | then show "(x \<sqinter> y) \<squnion> (- x \<squnion> - y) = \<top>" | |
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 655 | by (simp add: sup_compl_top) | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 656 | qed | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 657 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 658 | lemma compl_sup [simp]: | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 659 | "- (x \<squnion> y) = - x \<sqinter> - y" | 
| 44921 | 660 | using dual_boolean_algebra | 
| 661 | by (rule boolean_algebra.compl_inf) | |
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 662 | |
| 36008 | 663 | lemma compl_mono: | 
| 664 | "x \<sqsubseteq> y \<Longrightarrow> - y \<sqsubseteq> - x" | |
| 665 | proof - | |
| 666 | assume "x \<sqsubseteq> y" | |
| 667 | then have "x \<squnion> y = y" by (simp only: le_iff_sup) | |
| 668 | then have "- (x \<squnion> y) = - y" by simp | |
| 669 | then have "- x \<sqinter> - y = - y" by simp | |
| 670 | then have "- y \<sqinter> - x = - y" by (simp only: inf_commute) | |
| 671 | then show "- y \<sqsubseteq> - x" by (simp only: le_iff_inf) | |
| 672 | qed | |
| 673 | ||
| 44918 | 674 | lemma compl_le_compl_iff [simp]: | 
| 43753 | 675 | "- x \<sqsubseteq> - y \<longleftrightarrow> y \<sqsubseteq> x" | 
| 43873 | 676 | by (auto dest: compl_mono) | 
| 677 | ||
| 678 | lemma compl_le_swap1: | |
| 679 | assumes "y \<sqsubseteq> - x" shows "x \<sqsubseteq> -y" | |
| 680 | proof - | |
| 681 | from assms have "- (- x) \<sqsubseteq> - y" by (simp only: compl_le_compl_iff) | |
| 682 | then show ?thesis by simp | |
| 683 | qed | |
| 684 | ||
| 685 | lemma compl_le_swap2: | |
| 686 | assumes "- y \<sqsubseteq> x" shows "- x \<sqsubseteq> y" | |
| 687 | proof - | |
| 688 | from assms have "- x \<sqsubseteq> - (- y)" by (simp only: compl_le_compl_iff) | |
| 689 | then show ?thesis by simp | |
| 690 | qed | |
| 691 | ||
| 692 | lemma compl_less_compl_iff: (* TODO: declare [simp] ? *) | |
| 693 | "- x \<sqsubset> - y \<longleftrightarrow> y \<sqsubset> x" | |
| 44919 | 694 | by (auto simp add: less_le) | 
| 43873 | 695 | |
| 696 | lemma compl_less_swap1: | |
| 697 | assumes "y \<sqsubset> - x" shows "x \<sqsubset> - y" | |
| 698 | proof - | |
| 699 | from assms have "- (- x) \<sqsubset> - y" by (simp only: compl_less_compl_iff) | |
| 700 | then show ?thesis by simp | |
| 701 | qed | |
| 702 | ||
| 703 | lemma compl_less_swap2: | |
| 704 | assumes "- y \<sqsubset> x" shows "- x \<sqsubset> y" | |
| 705 | proof - | |
| 706 | from assms have "- x \<sqsubset> - (- y)" by (simp only: compl_less_compl_iff) | |
| 707 | then show ?thesis by simp | |
| 708 | qed | |
| 36008 | 709 | |
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 710 | end | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 711 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 712 | |
| 51540 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 713 | subsection {* @{text "min/max"} as special case of lattice *}
 | 
| 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 714 | |
| 54861 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 715 | context linorder | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 716 | begin | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 717 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 718 | sublocale min!: semilattice_order min less_eq less | 
| 51540 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 719 | + max!: semilattice_order max greater_eq greater | 
| 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 720 | by default (auto simp add: min_def max_def) | 
| 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 721 | |
| 54861 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 722 | lemma min_le_iff_disj: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 723 | "min x y \<le> z \<longleftrightarrow> x \<le> z \<or> y \<le> z" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 724 | unfolding min_def using linear by (auto intro: order_trans) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 725 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 726 | lemma le_max_iff_disj: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 727 | "z \<le> max x y \<longleftrightarrow> z \<le> x \<or> z \<le> y" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 728 | unfolding max_def using linear by (auto intro: order_trans) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 729 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 730 | lemma min_less_iff_disj: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 731 | "min x y < z \<longleftrightarrow> x < z \<or> y < z" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 732 | unfolding min_def le_less using less_linear by (auto intro: less_trans) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 733 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 734 | lemma less_max_iff_disj: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 735 | "z < max x y \<longleftrightarrow> z < x \<or> z < y" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 736 | unfolding max_def le_less using less_linear by (auto intro: less_trans) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 737 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 738 | lemma min_less_iff_conj [simp]: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 739 | "z < min x y \<longleftrightarrow> z < x \<and> z < y" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 740 | unfolding min_def le_less using less_linear by (auto intro: less_trans) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 741 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 742 | lemma max_less_iff_conj [simp]: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 743 | "max x y < z \<longleftrightarrow> x < z \<and> y < z" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 744 | unfolding max_def le_less using less_linear by (auto intro: less_trans) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 745 | |
| 54862 | 746 | lemma min_max_distrib1: | 
| 747 | "min (max b c) a = max (min b a) (min c a)" | |
| 748 | by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) | |
| 749 | ||
| 750 | lemma min_max_distrib2: | |
| 751 | "min a (max b c) = max (min a b) (min a c)" | |
| 752 | by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) | |
| 753 | ||
| 754 | lemma max_min_distrib1: | |
| 755 | "max (min b c) a = min (max b a) (max c a)" | |
| 756 | by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) | |
| 757 | ||
| 758 | lemma max_min_distrib2: | |
| 759 | "max a (min b c) = min (max a b) (max a c)" | |
| 760 | by (auto simp add: min_def max_def not_le dest: le_less_trans less_trans intro: antisym) | |
| 761 | ||
| 762 | lemmas min_max_distribs = min_max_distrib1 min_max_distrib2 max_min_distrib1 max_min_distrib2 | |
| 763 | ||
| 54861 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 764 | lemma split_min [no_atp]: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 765 | "P (min i j) \<longleftrightarrow> (i \<le> j \<longrightarrow> P i) \<and> (\<not> i \<le> j \<longrightarrow> P j)" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 766 | by (simp add: min_def) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 767 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 768 | lemma split_max [no_atp]: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 769 | "P (max i j) \<longleftrightarrow> (i \<le> j \<longrightarrow> P j) \<and> (\<not> i \<le> j \<longrightarrow> P i)" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 770 | by (simp add: max_def) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 771 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 772 | lemma min_of_mono: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 773 | fixes f :: "'a \<Rightarrow> 'b\<Colon>linorder" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 774 | shows "mono f \<Longrightarrow> min (f m) (f n) = f (min m n)" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 775 | by (auto simp: mono_def Orderings.min_def min_def intro: Orderings.antisym) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 776 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 777 | lemma max_of_mono: | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 778 | fixes f :: "'a \<Rightarrow> 'b\<Colon>linorder" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 779 | shows "mono f \<Longrightarrow> max (f m) (f n) = f (max m n)" | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 780 | by (auto simp: mono_def Orderings.max_def max_def intro: Orderings.antisym) | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 781 | |
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 782 | end | 
| 
00d551179872
postponed min/max lemmas until abstract lattice is available
 haftmann parents: 
54859diff
changeset | 783 | |
| 51540 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 784 | lemma inf_min: "inf = (min \<Colon> 'a\<Colon>{semilattice_inf, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
 | 
| 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 785 | by (auto intro: antisym simp add: min_def fun_eq_iff) | 
| 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 786 | |
| 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 787 | lemma sup_max: "sup = (max \<Colon> 'a\<Colon>{semilattice_sup, linorder} \<Rightarrow> 'a \<Rightarrow> 'a)"
 | 
| 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 788 | by (auto intro: antisym simp add: max_def fun_eq_iff) | 
| 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 789 | |
| 
eea5c4ca4a0e
explicit sublocale dependency for Min/Max yields more appropriate Min/Max prefix for a couple of facts
 haftmann parents: 
51489diff
changeset | 790 | |
| 22454 | 791 | subsection {* Uniqueness of inf and sup *}
 | 
| 792 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 793 | lemma (in semilattice_inf) inf_unique: | 
| 22454 | 794 | fixes f (infixl "\<triangle>" 70) | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 795 | assumes le1: "\<And>x y. x \<triangle> y \<sqsubseteq> x" and le2: "\<And>x y. x \<triangle> y \<sqsubseteq> y" | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 796 | and greatest: "\<And>x y z. x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<triangle> z" | 
| 22737 | 797 | shows "x \<sqinter> y = x \<triangle> y" | 
| 22454 | 798 | proof (rule antisym) | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 799 | show "x \<triangle> y \<sqsubseteq> x \<sqinter> y" by (rule le_infI) (rule le1, rule le2) | 
| 22454 | 800 | next | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 801 | have leI: "\<And>x y z. x \<sqsubseteq> y \<Longrightarrow> x \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> y \<triangle> z" by (blast intro: greatest) | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 802 | show "x \<sqinter> y \<sqsubseteq> x \<triangle> y" by (rule leI) simp_all | 
| 22454 | 803 | qed | 
| 804 | ||
| 35028 
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
 haftmann parents: 
34973diff
changeset | 805 | lemma (in semilattice_sup) sup_unique: | 
| 22454 | 806 | fixes f (infixl "\<nabla>" 70) | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 807 | assumes ge1 [simp]: "\<And>x y. x \<sqsubseteq> x \<nabla> y" and ge2: "\<And>x y. y \<sqsubseteq> x \<nabla> y" | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 808 | and least: "\<And>x y z. y \<sqsubseteq> x \<Longrightarrow> z \<sqsubseteq> x \<Longrightarrow> y \<nabla> z \<sqsubseteq> x" | 
| 22737 | 809 | shows "x \<squnion> y = x \<nabla> y" | 
| 22454 | 810 | proof (rule antisym) | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 811 | show "x \<squnion> y \<sqsubseteq> x \<nabla> y" by (rule le_supI) (rule ge1, rule ge2) | 
| 22454 | 812 | next | 
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 813 | have leI: "\<And>x y z. x \<sqsubseteq> z \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<nabla> y \<sqsubseteq> z" by (blast intro: least) | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 814 | show "x \<nabla> y \<sqsubseteq> x \<squnion> y" by (rule leI) simp_all | 
| 22454 | 815 | qed | 
| 36008 | 816 | |
| 22454 | 817 | |
| 46631 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 818 | subsection {* Lattice on @{typ bool} *}
 | 
| 22454 | 819 | |
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 820 | instantiation bool :: boolean_algebra | 
| 25510 | 821 | begin | 
| 822 | ||
| 823 | definition | |
| 41080 | 824 | bool_Compl_def [simp]: "uminus = Not" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 825 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 826 | definition | 
| 41080 | 827 | bool_diff_def [simp]: "A - B \<longleftrightarrow> A \<and> \<not> B" | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 828 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 829 | definition | 
| 41080 | 830 | [simp]: "P \<sqinter> Q \<longleftrightarrow> P \<and> Q" | 
| 25510 | 831 | |
| 832 | definition | |
| 41080 | 833 | [simp]: "P \<squnion> Q \<longleftrightarrow> P \<or> Q" | 
| 25510 | 834 | |
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 835 | instance proof | 
| 41080 | 836 | qed auto | 
| 22454 | 837 | |
| 25510 | 838 | end | 
| 839 | ||
| 32781 | 840 | lemma sup_boolI1: | 
| 841 | "P \<Longrightarrow> P \<squnion> Q" | |
| 41080 | 842 | by simp | 
| 32781 | 843 | |
| 844 | lemma sup_boolI2: | |
| 845 | "Q \<Longrightarrow> P \<squnion> Q" | |
| 41080 | 846 | by simp | 
| 32781 | 847 | |
| 848 | lemma sup_boolE: | |
| 849 | "P \<squnion> Q \<Longrightarrow> (P \<Longrightarrow> R) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R" | |
| 41080 | 850 | by auto | 
| 32781 | 851 | |
| 23878 | 852 | |
| 46631 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 853 | subsection {* Lattice on @{typ "_ \<Rightarrow> _"} *}
 | 
| 23878 | 854 | |
| 51387 | 855 | instantiation "fun" :: (type, semilattice_sup) semilattice_sup | 
| 25510 | 856 | begin | 
| 857 | ||
| 858 | definition | |
| 41080 | 859 | "f \<squnion> g = (\<lambda>x. f x \<squnion> g x)" | 
| 860 | ||
| 49769 | 861 | lemma sup_apply [simp, code]: | 
| 41080 | 862 | "(f \<squnion> g) x = f x \<squnion> g x" | 
| 863 | by (simp add: sup_fun_def) | |
| 25510 | 864 | |
| 32780 | 865 | instance proof | 
| 46884 | 866 | qed (simp_all add: le_fun_def) | 
| 23878 | 867 | |
| 25510 | 868 | end | 
| 23878 | 869 | |
| 51387 | 870 | instantiation "fun" :: (type, semilattice_inf) semilattice_inf | 
| 871 | begin | |
| 872 | ||
| 873 | definition | |
| 874 | "f \<sqinter> g = (\<lambda>x. f x \<sqinter> g x)" | |
| 875 | ||
| 876 | lemma inf_apply [simp, code]: | |
| 877 | "(f \<sqinter> g) x = f x \<sqinter> g x" | |
| 878 | by (simp add: inf_fun_def) | |
| 879 | ||
| 880 | instance proof | |
| 881 | qed (simp_all add: le_fun_def) | |
| 882 | ||
| 883 | end | |
| 884 | ||
| 885 | instance "fun" :: (type, lattice) lattice .. | |
| 886 | ||
| 41080 | 887 | instance "fun" :: (type, distrib_lattice) distrib_lattice proof | 
| 46884 | 888 | qed (rule ext, simp add: sup_inf_distrib1) | 
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 889 | |
| 34007 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 890 | instance "fun" :: (type, bounded_lattice) bounded_lattice .. | 
| 
aea892559fc5
tuned lattices theory fragements; generlized some lemmas from sets to lattices
 haftmann parents: 
32781diff
changeset | 891 | |
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 892 | instantiation "fun" :: (type, uminus) uminus | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 893 | begin | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 894 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 895 | definition | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 896 | fun_Compl_def: "- A = (\<lambda>x. - A x)" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 897 | |
| 49769 | 898 | lemma uminus_apply [simp, code]: | 
| 41080 | 899 | "(- A) x = - (A x)" | 
| 900 | by (simp add: fun_Compl_def) | |
| 901 | ||
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 902 | instance .. | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 903 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 904 | end | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 905 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 906 | instantiation "fun" :: (type, minus) minus | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 907 | begin | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 908 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 909 | definition | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 910 | fun_diff_def: "A - B = (\<lambda>x. A x - B x)" | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 911 | |
| 49769 | 912 | lemma minus_apply [simp, code]: | 
| 41080 | 913 | "(A - B) x = A x - B x" | 
| 914 | by (simp add: fun_diff_def) | |
| 915 | ||
| 31991 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 916 | instance .. | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 917 | |
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 918 | end | 
| 
37390299214a
added boolean_algebra type class; tuned lattice duals
 haftmann parents: 
30729diff
changeset | 919 | |
| 41080 | 920 | instance "fun" :: (type, boolean_algebra) boolean_algebra proof | 
| 46884 | 921 | qed (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+ | 
| 26794 | 922 | |
| 46631 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 923 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 924 | subsection {* Lattice on unary and binary predicates *}
 | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 925 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 926 | lemma inf1I: "A x \<Longrightarrow> B x \<Longrightarrow> (A \<sqinter> B) x" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 927 | by (simp add: inf_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 928 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 929 | lemma inf2I: "A x y \<Longrightarrow> B x y \<Longrightarrow> (A \<sqinter> B) x y" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 930 | by (simp add: inf_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 931 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 932 | lemma inf1E: "(A \<sqinter> B) x \<Longrightarrow> (A x \<Longrightarrow> B x \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 933 | by (simp add: inf_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 934 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 935 | lemma inf2E: "(A \<sqinter> B) x y \<Longrightarrow> (A x y \<Longrightarrow> B x y \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 936 | by (simp add: inf_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 937 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 938 | lemma inf1D1: "(A \<sqinter> B) x \<Longrightarrow> A x" | 
| 54857 | 939 | by (rule inf1E) | 
| 46631 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 940 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 941 | lemma inf2D1: "(A \<sqinter> B) x y \<Longrightarrow> A x y" | 
| 54857 | 942 | by (rule inf2E) | 
| 46631 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 943 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 944 | lemma inf1D2: "(A \<sqinter> B) x \<Longrightarrow> B x" | 
| 54857 | 945 | by (rule inf1E) | 
| 46631 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 946 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 947 | lemma inf2D2: "(A \<sqinter> B) x y \<Longrightarrow> B x y" | 
| 54857 | 948 | by (rule inf2E) | 
| 46631 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 949 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 950 | lemma sup1I1: "A x \<Longrightarrow> (A \<squnion> B) x" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 951 | by (simp add: sup_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 952 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 953 | lemma sup2I1: "A x y \<Longrightarrow> (A \<squnion> B) x y" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 954 | by (simp add: sup_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 955 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 956 | lemma sup1I2: "B x \<Longrightarrow> (A \<squnion> B) x" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 957 | by (simp add: sup_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 958 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 959 | lemma sup2I2: "B x y \<Longrightarrow> (A \<squnion> B) x y" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 960 | by (simp add: sup_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 961 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 962 | lemma sup1E: "(A \<squnion> B) x \<Longrightarrow> (A x \<Longrightarrow> P) \<Longrightarrow> (B x \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 963 | by (simp add: sup_fun_def) iprover | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 964 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 965 | lemma sup2E: "(A \<squnion> B) x y \<Longrightarrow> (A x y \<Longrightarrow> P) \<Longrightarrow> (B x y \<Longrightarrow> P) \<Longrightarrow> P" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 966 | by (simp add: sup_fun_def) iprover | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 967 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 968 | text {*
 | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 969 |   \medskip Classical introduction rule: no commitment to @{text A} vs
 | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 970 |   @{text B}.
 | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 971 | *} | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 972 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 973 | lemma sup1CI: "(\<not> B x \<Longrightarrow> A x) \<Longrightarrow> (A \<squnion> B) x" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 974 | by (auto simp add: sup_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 975 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 976 | lemma sup2CI: "(\<not> B x y \<Longrightarrow> A x y) \<Longrightarrow> (A \<squnion> B) x y" | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 977 | by (auto simp add: sup_fun_def) | 
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 978 | |
| 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 979 | |
| 25062 | 980 | no_notation | 
| 46691 | 981 | less_eq (infix "\<sqsubseteq>" 50) and | 
| 982 | less (infix "\<sqsubset>" 50) | |
| 25062 | 983 | |
| 21249 | 984 | end | 
| 46631 
2c5c003cee35
moved lemmas for orderings and lattices on predicates to corresponding theories, retaining declaration order of classical rules; tuned headings; tuned syntax
 haftmann parents: 
46557diff
changeset | 985 |