| author | wenzelm | 
| Sun, 06 Jan 2008 18:09:34 +0100 | |
| changeset 25856 | 890c51553b33 | 
| parent 25691 | 8f8d83af100a | 
| child 26993 | b952df8d505b | 
| permissions | -rw-r--r-- | 
| 15737 | 1 | (* Title: HOL/Library/Product_ord.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Norbert Voelker | |
| 4 | *) | |
| 5 | ||
| 17200 | 6 | header {* Order on product types *}
 | 
| 15737 | 7 | |
| 8 | theory Product_ord | |
| 25691 | 9 | imports ATP_Linkup | 
| 15737 | 10 | begin | 
| 11 | ||
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 12 | instantiation "*" :: (ord, ord) ord | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 13 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 14 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 15 | definition | 
| 25502 | 16 | prod_le_def [code func del]: "x \<le> y \<longleftrightarrow> fst x < fst y \<or> fst x = fst y \<and> snd x \<le> snd y" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 17 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 18 | definition | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 19 | prod_less_def [code func del]: "x < y \<longleftrightarrow> fst x < fst y \<or> fst x = fst y \<and> snd x < snd y" | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 20 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 21 | instance .. | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 22 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 23 | end | 
| 15737 | 24 | |
| 22177 | 25 | lemma [code func]: | 
| 26 |   "(x1\<Colon>'a\<Colon>{ord, eq}, y1) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 = x2 \<and> y1 \<le> y2"
 | |
| 27 |   "(x1\<Colon>'a\<Colon>{ord, eq}, y1) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 = x2 \<and> y1 < y2"
 | |
| 25502 | 28 | unfolding prod_le_def prod_less_def by simp_all | 
| 22177 | 29 | |
| 21458 | 30 | lemma [code]: | 
| 31 | "(x1, y1) \<le> (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 = x2 \<and> y1 \<le> y2" | |
| 32 | "(x1, y1) < (x2, y2) \<longleftrightarrow> x1 < x2 \<or> x1 = x2 \<and> y1 < y2" | |
| 25502 | 33 | unfolding prod_le_def prod_less_def by simp_all | 
| 21458 | 34 | |
| 19736 | 35 | instance * :: (order, order) order | 
| 25502 | 36 | by default (auto simp: prod_le_def prod_less_def intro: order_less_trans) | 
| 15737 | 37 | |
| 19736 | 38 | instance * :: (linorder, linorder) linorder | 
| 39 | by default (auto simp: prod_le_def) | |
| 15737 | 40 | |
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 41 | instantiation * :: (linorder, linorder) distrib_lattice | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 42 | begin | 
| 
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 | inf_prod_def: "(inf \<Colon> 'a \<times> 'b \<Rightarrow> _ \<Rightarrow> _) = min" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 46 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 47 | definition | 
| 25502 | 48 | sup_prod_def: "(sup \<Colon> 'a \<times> 'b \<Rightarrow> _ \<Rightarrow> _) = max" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 49 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 50 | instance | 
| 22483 | 51 | by intro_classes | 
| 52 | (auto simp add: inf_prod_def sup_prod_def min_max.sup_inf_distrib1) | |
| 53 | ||
| 19736 | 54 | end | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 55 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25502diff
changeset | 56 | end |