diff -r fdfbbb92dadf -r c9e39eafc7a0 src/HOL/Library/Product_ord.thy --- a/src/HOL/Library/Product_ord.thy Fri Dec 07 15:07:56 2007 +0100 +++ b/src/HOL/Library/Product_ord.thy Fri Dec 07 15:07:59 2007 +0100 @@ -9,9 +9,18 @@ imports Main begin -instance "*" :: (ord, ord) ord +instantiation "*" :: (ord, ord) ord +begin + +definition prod_le_def [code func del]: "x \ y \ fst x < fst y \ fst x = fst y \ snd x \ snd y" - prod_less_def [code func del]: "x < y \ fst x < fst y \ fst x = fst y \ snd x < snd y" .. + +definition + prod_less_def [code func del]: "x < y \ fst x < fst y \ fst x = fst y \ snd x < snd y" + +instance .. + +end lemma [code func]: "(x1\'a\{ord, eq}, y1) \ (x2, y2) \ x1 < x2 \ x1 = x2 \ y1 \ y2" @@ -29,10 +38,19 @@ instance * :: (linorder, linorder) linorder by default (auto simp: prod_le_def) -instance * :: (linorder, linorder) distrib_lattice +instantiation * :: (linorder, linorder) distrib_lattice +begin + +definition inf_prod_def: "(inf \ 'a \ 'b \ _ \ _) = min" + +definition sup_prod_def: "(sup \ 'a \ 'b \ _ \ _) = max" + +instance by intro_classes (auto simp add: inf_prod_def sup_prod_def min_max.sup_inf_distrib1) end + +end