| author | noschinl | 
| Fri, 17 Apr 2015 11:12:19 +0200 | |
| changeset 60109 | 22389d4cdd6b | 
| parent 58881 | b9556a055632 | 
| child 60500 | 903bb1495239 | 
| permissions | -rw-r--r-- | 
| 51115 
7dbd6832a689
consolidation of library theories on product orders
 haftmann parents: 
50573diff
changeset | 1 | (* Title: HOL/Library/Product_Order.thy | 
| 44006 | 2 | Author: Brian Huffman | 
| 3 | *) | |
| 4 | ||
| 58881 | 5 | section {* Pointwise order on product types *}
 | 
| 44006 | 6 | |
| 51115 
7dbd6832a689
consolidation of library theories on product orders
 haftmann parents: 
50573diff
changeset | 7 | theory Product_Order | 
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 8 | imports Product_plus Conditionally_Complete_Lattices | 
| 44006 | 9 | begin | 
| 10 | ||
| 11 | subsection {* Pointwise ordering *}
 | |
| 12 | ||
| 13 | instantiation prod :: (ord, ord) ord | |
| 14 | begin | |
| 15 | ||
| 16 | definition | |
| 17 | "x \<le> y \<longleftrightarrow> fst x \<le> fst y \<and> snd x \<le> snd y" | |
| 18 | ||
| 19 | definition | |
| 20 | "(x::'a \<times> 'b) < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x" | |
| 21 | ||
| 22 | instance .. | |
| 23 | ||
| 24 | end | |
| 25 | ||
| 26 | lemma fst_mono: "x \<le> y \<Longrightarrow> fst x \<le> fst y" | |
| 27 | unfolding less_eq_prod_def by simp | |
| 28 | ||
| 29 | lemma snd_mono: "x \<le> y \<Longrightarrow> snd x \<le> snd y" | |
| 30 | unfolding less_eq_prod_def by simp | |
| 31 | ||
| 32 | lemma Pair_mono: "x \<le> x' \<Longrightarrow> y \<le> y' \<Longrightarrow> (x, y) \<le> (x', y')" | |
| 33 | unfolding less_eq_prod_def by simp | |
| 34 | ||
| 35 | lemma Pair_le [simp]: "(a, b) \<le> (c, d) \<longleftrightarrow> a \<le> c \<and> b \<le> d" | |
| 36 | unfolding less_eq_prod_def by simp | |
| 37 | ||
| 38 | instance prod :: (preorder, preorder) preorder | |
| 39 | proof | |
| 40 | fix x y z :: "'a \<times> 'b" | |
| 41 | show "x < y \<longleftrightarrow> x \<le> y \<and> \<not> y \<le> x" | |
| 42 | by (rule less_prod_def) | |
| 43 | show "x \<le> x" | |
| 44 | unfolding less_eq_prod_def | |
| 45 | by fast | |
| 46 | assume "x \<le> y" and "y \<le> z" thus "x \<le> z" | |
| 47 | unfolding less_eq_prod_def | |
| 48 | by (fast elim: order_trans) | |
| 49 | qed | |
| 50 | ||
| 51 | instance prod :: (order, order) order | |
| 52 | by default auto | |
| 53 | ||
| 54 | ||
| 55 | subsection {* Binary infimum and supremum *}
 | |
| 56 | ||
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 57 | instantiation prod :: (inf, inf) inf | 
| 44006 | 58 | begin | 
| 59 | ||
| 60 | definition | |
| 61 | "inf x y = (inf (fst x) (fst y), inf (snd x) (snd y))" | |
| 62 | ||
| 63 | lemma inf_Pair_Pair [simp]: "inf (a, b) (c, d) = (inf a c, inf b d)" | |
| 64 | unfolding inf_prod_def by simp | |
| 65 | ||
| 66 | lemma fst_inf [simp]: "fst (inf x y) = inf (fst x) (fst y)" | |
| 67 | unfolding inf_prod_def by simp | |
| 68 | ||
| 69 | lemma snd_inf [simp]: "snd (inf x y) = inf (snd x) (snd y)" | |
| 70 | unfolding inf_prod_def by simp | |
| 71 | ||
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 72 | instance proof qed | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 73 | end | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 74 | |
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 75 | instance prod :: (semilattice_inf, semilattice_inf) semilattice_inf | 
| 44006 | 76 | by default auto | 
| 77 | ||
| 78 | ||
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 79 | instantiation prod :: (sup, sup) sup | 
| 44006 | 80 | begin | 
| 81 | ||
| 82 | definition | |
| 83 | "sup x y = (sup (fst x) (fst y), sup (snd x) (snd y))" | |
| 84 | ||
| 85 | lemma sup_Pair_Pair [simp]: "sup (a, b) (c, d) = (sup a c, sup b d)" | |
| 86 | unfolding sup_prod_def by simp | |
| 87 | ||
| 88 | lemma fst_sup [simp]: "fst (sup x y) = sup (fst x) (fst y)" | |
| 89 | unfolding sup_prod_def by simp | |
| 90 | ||
| 91 | lemma snd_sup [simp]: "snd (sup x y) = sup (snd x) (snd y)" | |
| 92 | unfolding sup_prod_def by simp | |
| 93 | ||
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 94 | instance proof qed | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 95 | end | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 96 | |
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 97 | instance prod :: (semilattice_sup, semilattice_sup) semilattice_sup | 
| 44006 | 98 | by default auto | 
| 99 | ||
| 100 | instance prod :: (lattice, lattice) lattice .. | |
| 101 | ||
| 102 | instance prod :: (distrib_lattice, distrib_lattice) distrib_lattice | |
| 103 | by default (auto simp add: sup_inf_distrib1) | |
| 104 | ||
| 105 | ||
| 106 | subsection {* Top and bottom elements *}
 | |
| 107 | ||
| 108 | instantiation prod :: (top, top) top | |
| 109 | begin | |
| 110 | ||
| 111 | definition | |
| 112 | "top = (top, top)" | |
| 113 | ||
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 114 | instance .. | 
| 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 115 | |
| 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 116 | end | 
| 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 117 | |
| 44006 | 118 | lemma fst_top [simp]: "fst top = top" | 
| 119 | unfolding top_prod_def by simp | |
| 120 | ||
| 121 | lemma snd_top [simp]: "snd top = top" | |
| 122 | unfolding top_prod_def by simp | |
| 123 | ||
| 124 | lemma Pair_top_top: "(top, top) = top" | |
| 125 | unfolding top_prod_def by simp | |
| 126 | ||
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 127 | instance prod :: (order_top, order_top) order_top | 
| 44006 | 128 | by default (auto simp add: top_prod_def) | 
| 129 | ||
| 130 | instantiation prod :: (bot, bot) bot | |
| 131 | begin | |
| 132 | ||
| 133 | definition | |
| 134 | "bot = (bot, bot)" | |
| 135 | ||
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 136 | instance .. | 
| 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 137 | |
| 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 138 | end | 
| 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 139 | |
| 44006 | 140 | lemma fst_bot [simp]: "fst bot = bot" | 
| 141 | unfolding bot_prod_def by simp | |
| 142 | ||
| 143 | lemma snd_bot [simp]: "snd bot = bot" | |
| 144 | unfolding bot_prod_def by simp | |
| 145 | ||
| 146 | lemma Pair_bot_bot: "(bot, bot) = bot" | |
| 147 | unfolding bot_prod_def by simp | |
| 148 | ||
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 149 | instance prod :: (order_bot, order_bot) order_bot | 
| 44006 | 150 | by default (auto simp add: bot_prod_def) | 
| 151 | ||
| 152 | instance prod :: (bounded_lattice, bounded_lattice) bounded_lattice .. | |
| 153 | ||
| 154 | instance prod :: (boolean_algebra, boolean_algebra) boolean_algebra | |
| 155 | by default (auto simp add: prod_eqI inf_compl_bot sup_compl_top diff_eq) | |
| 156 | ||
| 157 | ||
| 158 | subsection {* Complete lattice operations *}
 | |
| 159 | ||
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 160 | instantiation prod :: (Inf, Inf) Inf | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 161 | begin | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 162 | |
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 163 | definition | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 164 | "Inf A = (INF x:A. fst x, INF x:A. snd x)" | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 165 | |
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 166 | instance proof qed | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 167 | end | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 168 | |
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 169 | instantiation prod :: (Sup, Sup) Sup | 
| 44006 | 170 | begin | 
| 171 | ||
| 172 | definition | |
| 173 | "Sup A = (SUP x:A. fst x, SUP x:A. snd x)" | |
| 174 | ||
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 175 | instance proof qed | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 176 | end | 
| 44006 | 177 | |
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 178 | instance prod :: (conditionally_complete_lattice, conditionally_complete_lattice) | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 179 | conditionally_complete_lattice | 
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 180 | by default (force simp: less_eq_prod_def Inf_prod_def Sup_prod_def bdd_below_def bdd_above_def | 
| 56166 | 181 | INF_def SUP_def simp del: Inf_image_eq Sup_image_eq intro!: cInf_lower cSup_upper cInf_greatest cSup_least)+ | 
| 54776 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 182 | |
| 
db890d9fc5c2
ordered_euclidean_space compatible with more standard pointwise ordering on products; conditionally complete lattice with product order
 immler parents: 
52729diff
changeset | 183 | instance prod :: (complete_lattice, complete_lattice) complete_lattice | 
| 44006 | 184 | by default (simp_all add: less_eq_prod_def Inf_prod_def Sup_prod_def | 
| 52729 
412c9e0381a1
factored syntactic type classes for bot and top (by Alessandro Coglio)
 haftmann parents: 
51542diff
changeset | 185 | INF_lower SUP_upper le_INF_iff SUP_le_iff bot_prod_def top_prod_def) | 
| 44006 | 186 | |
| 187 | lemma fst_Sup: "fst (Sup A) = (SUP x:A. fst x)" | |
| 188 | unfolding Sup_prod_def by simp | |
| 189 | ||
| 190 | lemma snd_Sup: "snd (Sup A) = (SUP x:A. snd x)" | |
| 191 | unfolding Sup_prod_def by simp | |
| 192 | ||
| 193 | lemma fst_Inf: "fst (Inf A) = (INF x:A. fst x)" | |
| 194 | unfolding Inf_prod_def by simp | |
| 195 | ||
| 196 | lemma snd_Inf: "snd (Inf A) = (INF x:A. snd x)" | |
| 197 | unfolding Inf_prod_def by simp | |
| 198 | ||
| 199 | lemma fst_SUP: "fst (SUP x:A. f x) = (SUP x:A. fst (f x))" | |
| 56166 | 200 | using fst_Sup [of "f ` A", symmetric] by (simp add: comp_def) | 
| 44006 | 201 | |
| 202 | lemma snd_SUP: "snd (SUP x:A. f x) = (SUP x:A. snd (f x))" | |
| 56166 | 203 | using snd_Sup [of "f ` A", symmetric] by (simp add: comp_def) | 
| 44006 | 204 | |
| 205 | lemma fst_INF: "fst (INF x:A. f x) = (INF x:A. fst (f x))" | |
| 56166 | 206 | using fst_Inf [of "f ` A", symmetric] by (simp add: comp_def) | 
| 44006 | 207 | |
| 208 | lemma snd_INF: "snd (INF x:A. f x) = (INF x:A. snd (f x))" | |
| 56166 | 209 | using snd_Inf [of "f ` A", symmetric] by (simp add: comp_def) | 
| 44006 | 210 | |
| 211 | lemma SUP_Pair: "(SUP x:A. (f x, g x)) = (SUP x:A. f x, SUP x:A. g x)" | |
| 56166 | 212 | unfolding SUP_def Sup_prod_def by (simp add: comp_def) | 
| 44006 | 213 | |
| 214 | lemma INF_Pair: "(INF x:A. (f x, g x)) = (INF x:A. f x, INF x:A. g x)" | |
| 56166 | 215 | unfolding INF_def Inf_prod_def by (simp add: comp_def) | 
| 44006 | 216 | |
| 50535 | 217 | |
| 218 | text {* Alternative formulations for set infima and suprema over the product
 | |
| 219 | of two complete lattices: *} | |
| 220 | ||
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56166diff
changeset | 221 | lemma INF_prod_alt_def: | 
| 56218 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 haftmann parents: 
56212diff
changeset | 222 | "INFIMUM A f = (INFIMUM A (fst o f), INFIMUM A (snd o f))" | 
| 56166 | 223 | unfolding INF_def Inf_prod_def by simp | 
| 50535 | 224 | |
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56166diff
changeset | 225 | lemma SUP_prod_alt_def: | 
| 56218 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 haftmann parents: 
56212diff
changeset | 226 | "SUPREMUM A f = (SUPREMUM A (fst o f), SUPREMUM A (snd o f))" | 
| 56166 | 227 | unfolding SUP_def Sup_prod_def by simp | 
| 50535 | 228 | |
| 229 | ||
| 230 | subsection {* Complete distributive lattices *}
 | |
| 231 | ||
| 50573 | 232 | (* Contribution: Alessandro Coglio *) | 
| 50535 | 233 | |
| 234 | instance prod :: | |
| 235 | (complete_distrib_lattice, complete_distrib_lattice) complete_distrib_lattice | |
| 236 | proof | |
| 237 | case goal1 thus ?case | |
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56166diff
changeset | 238 | by (auto simp: sup_prod_def Inf_prod_def INF_prod_alt_def sup_Inf sup_INF comp_def) | 
| 50535 | 239 | next | 
| 240 | case goal2 thus ?case | |
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56166diff
changeset | 241 | by (auto simp: inf_prod_def Sup_prod_def SUP_prod_alt_def inf_Sup inf_SUP comp_def) | 
| 50535 | 242 | qed | 
| 243 | ||
| 51115 
7dbd6832a689
consolidation of library theories on product orders
 haftmann parents: 
50573diff
changeset | 244 | end | 
| 50535 | 245 |