| author | bulwahn | 
| Fri, 03 Dec 2010 08:40:47 +0100 | |
| changeset 40921 | a516fbdd9931 | 
| parent 38857 | 97775f3e8722 | 
| child 44063 | 4588597ba37e | 
| permissions | -rw-r--r-- | 
| 15737 | 1 | (* Title: HOL/Library/Product_ord.thy | 
| 2 | Author: Norbert Voelker | |
| 3 | *) | |
| 4 | ||
| 17200 | 5 | header {* Order on product types *}
 | 
| 15737 | 6 | |
| 7 | theory Product_ord | |
| 30738 | 8 | imports Main | 
| 15737 | 9 | begin | 
| 10 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
31040diff
changeset | 11 | instantiation prod :: (ord, ord) ord | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 12 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 13 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 14 | definition | 
| 37765 | 15 | prod_le_def: "x \<le> y \<longleftrightarrow> fst x < fst y \<or> fst x \<le> fst y \<and> snd x \<le> snd y" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 16 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 17 | definition | 
| 37765 | 18 | prod_less_def: "x < y \<longleftrightarrow> fst x < fst y \<or> fst x \<le> fst y \<and> snd x < snd y" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 19 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 20 | instance .. | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 21 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 22 | end | 
| 15737 | 23 | |
| 28562 | 24 | lemma [code]: | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 25 |   "(x1\<Colon>'a\<Colon>{ord, equal}, y1) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 \<le> y2"
 | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
37765diff
changeset | 26 |   "(x1\<Colon>'a\<Colon>{ord, equal}, y1) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 \<le> x2 \<and> y1 < y2"
 | 
| 25502 | 27 | unfolding prod_le_def prod_less_def by simp_all | 
| 22177 | 28 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
31040diff
changeset | 29 | instance prod :: (preorder, preorder) preorder proof | 
| 31040 | 30 | qed (auto simp: prod_le_def prod_less_def less_le_not_le intro: order_trans) | 
| 15737 | 31 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
31040diff
changeset | 32 | instance prod :: (order, order) order proof | 
| 31040 | 33 | qed (auto simp add: prod_le_def) | 
| 34 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
31040diff
changeset | 35 | instance prod :: (linorder, linorder) linorder proof | 
| 31040 | 36 | qed (auto simp: prod_le_def) | 
| 15737 | 37 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
31040diff
changeset | 38 | instantiation prod :: (linorder, linorder) distrib_lattice | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 39 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 40 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 41 | definition | 
| 25502 | 42 | inf_prod_def: "(inf \<Colon> 'a \<times> 'b \<Rightarrow> _ \<Rightarrow> _) = min" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 43 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 44 | definition | 
| 25502 | 45 | sup_prod_def: "(sup \<Colon> 'a \<times> 'b \<Rightarrow> _ \<Rightarrow> _) = max" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 46 | |
| 31040 | 47 | instance proof | 
| 48 | qed (auto simp add: inf_prod_def sup_prod_def min_max.sup_inf_distrib1) | |
| 49 | ||
| 50 | end | |
| 51 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
31040diff
changeset | 52 | instantiation prod :: (bot, bot) bot | 
| 31040 | 53 | begin | 
| 54 | ||
| 55 | definition | |
| 56 | bot_prod_def: "bot = (bot, bot)" | |
| 57 | ||
| 58 | instance proof | |
| 59 | qed (auto simp add: bot_prod_def prod_le_def) | |
| 60 | ||
| 61 | end | |
| 62 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
31040diff
changeset | 63 | instantiation prod :: (top, top) top | 
| 31040 | 64 | begin | 
| 65 | ||
| 66 | definition | |
| 67 | top_prod_def: "top = (top, top)" | |
| 68 | ||
| 69 | instance proof | |
| 70 | qed (auto simp add: top_prod_def prod_le_def) | |
| 22483 | 71 | |
| 19736 | 72 | end | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 73 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 74 | end |