# HG changeset patch # User huffman # Date 1312216270 25200 # Node ID b9839fad3bb62682f4cf8d7fde8a68b65c35f1f0 # Parent 421f8bc19ce4d8bea1c2a584c4576e636e3e079b new theory HOL/Library/Product_Lattice.thy diff -r 421f8bc19ce4 -r b9839fad3bb6 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Sun Jul 31 11:13:38 2011 -0700 +++ b/src/HOL/IsaMakefile Mon Aug 01 09:31:10 2011 -0700 @@ -470,6 +470,7 @@ Library/Polynomial.thy Library/Predicate_Compile_Quickcheck.thy \ Library/Preorder.thy Library/Product_Vector.thy \ Library/Product_ord.thy Library/Product_plus.thy \ + Library/Product_Lattice.thy \ Library/Quickcheck_Types.thy \ Library/Quotient_List.thy Library/Quotient_Option.thy \ Library/Quotient_Product.thy Library/Quotient_Sum.thy \ diff -r 421f8bc19ce4 -r b9839fad3bb6 src/HOL/Library/Product_Lattice.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Product_Lattice.thy Mon Aug 01 09:31:10 2011 -0700 @@ -0,0 +1,198 @@ +(* Title: Product_Lattice.thy + Author: Brian Huffman +*) + +header {* Lattice operations on product types *} + +theory Product_Lattice +imports "~~/src/HOL/Library/Product_plus" +begin + +subsection {* Pointwise ordering *} + +instantiation prod :: (ord, ord) ord +begin + +definition + "x \ y \ fst x \ fst y \ snd x \ snd y" + +definition + "(x::'a \ 'b) < y \ x \ y \ \ y \ x" + +instance .. + +end + +lemma fst_mono: "x \ y \ fst x \ fst y" + unfolding less_eq_prod_def by simp + +lemma snd_mono: "x \ y \ snd x \ snd y" + unfolding less_eq_prod_def by simp + +lemma Pair_mono: "x \ x' \ y \ y' \ (x, y) \ (x', y')" + unfolding less_eq_prod_def by simp + +lemma Pair_le [simp]: "(a, b) \ (c, d) \ a \ c \ b \ d" + unfolding less_eq_prod_def by simp + +instance prod :: (preorder, preorder) preorder +proof + fix x y z :: "'a \ 'b" + show "x < y \ x \ y \ \ y \ x" + by (rule less_prod_def) + show "x \ x" + unfolding less_eq_prod_def + by fast + assume "x \ y" and "y \ z" thus "x \ z" + unfolding less_eq_prod_def + by (fast elim: order_trans) +qed + +instance prod :: (order, order) order + by default auto + + +subsection {* Binary infimum and supremum *} + +instantiation prod :: (semilattice_inf, semilattice_inf) semilattice_inf +begin + +definition + "inf x y = (inf (fst x) (fst y), inf (snd x) (snd y))" + +lemma inf_Pair_Pair [simp]: "inf (a, b) (c, d) = (inf a c, inf b d)" + unfolding inf_prod_def by simp + +lemma fst_inf [simp]: "fst (inf x y) = inf (fst x) (fst y)" + unfolding inf_prod_def by simp + +lemma snd_inf [simp]: "snd (inf x y) = inf (snd x) (snd y)" + unfolding inf_prod_def by simp + +instance + by default auto + +end + +instantiation prod :: (semilattice_sup, semilattice_sup) semilattice_sup +begin + +definition + "sup x y = (sup (fst x) (fst y), sup (snd x) (snd y))" + +lemma sup_Pair_Pair [simp]: "sup (a, b) (c, d) = (sup a c, sup b d)" + unfolding sup_prod_def by simp + +lemma fst_sup [simp]: "fst (sup x y) = sup (fst x) (fst y)" + unfolding sup_prod_def by simp + +lemma snd_sup [simp]: "snd (sup x y) = sup (snd x) (snd y)" + unfolding sup_prod_def by simp + +instance + by default auto + +end + +instance prod :: (lattice, lattice) lattice .. + +instance prod :: (distrib_lattice, distrib_lattice) distrib_lattice + by default (auto simp add: sup_inf_distrib1) + + +subsection {* Top and bottom elements *} + +instantiation prod :: (top, top) top +begin + +definition + "top = (top, top)" + +lemma fst_top [simp]: "fst top = top" + unfolding top_prod_def by simp + +lemma snd_top [simp]: "snd top = top" + unfolding top_prod_def by simp + +lemma Pair_top_top: "(top, top) = top" + unfolding top_prod_def by simp + +instance + by default (auto simp add: top_prod_def) + +end + +instantiation prod :: (bot, bot) bot +begin + +definition + "bot = (bot, bot)" + +lemma fst_bot [simp]: "fst bot = bot" + unfolding bot_prod_def by simp + +lemma snd_bot [simp]: "snd bot = bot" + unfolding bot_prod_def by simp + +lemma Pair_bot_bot: "(bot, bot) = bot" + unfolding bot_prod_def by simp + +instance + by default (auto simp add: bot_prod_def) + +end + +instance prod :: (bounded_lattice, bounded_lattice) bounded_lattice .. + +instance prod :: (boolean_algebra, boolean_algebra) boolean_algebra + by default (auto simp add: prod_eqI inf_compl_bot sup_compl_top diff_eq) + + +subsection {* Complete lattice operations *} + +instantiation prod :: (complete_lattice, complete_lattice) complete_lattice +begin + +definition + "Sup A = (SUP x:A. fst x, SUP x:A. snd x)" + +definition + "Inf A = (INF x:A. fst x, INF x:A. snd x)" + +instance + by default (simp_all add: less_eq_prod_def Inf_prod_def Sup_prod_def + INF_leI le_SUPI le_INF_iff SUP_le_iff) + +end + +lemma fst_Sup: "fst (Sup A) = (SUP x:A. fst x)" + unfolding Sup_prod_def by simp + +lemma snd_Sup: "snd (Sup A) = (SUP x:A. snd x)" + unfolding Sup_prod_def by simp + +lemma fst_Inf: "fst (Inf A) = (INF x:A. fst x)" + unfolding Inf_prod_def by simp + +lemma snd_Inf: "snd (Inf A) = (INF x:A. snd x)" + unfolding Inf_prod_def by simp + +lemma fst_SUP: "fst (SUP x:A. f x) = (SUP x:A. fst (f x))" + by (simp add: SUPR_def fst_Sup image_image) + +lemma snd_SUP: "snd (SUP x:A. f x) = (SUP x:A. snd (f x))" + by (simp add: SUPR_def snd_Sup image_image) + +lemma fst_INF: "fst (INF x:A. f x) = (INF x:A. fst (f x))" + by (simp add: INFI_def fst_Inf image_image) + +lemma snd_INF: "snd (INF x:A. f x) = (INF x:A. snd (f x))" + by (simp add: INFI_def snd_Inf image_image) + +lemma SUP_Pair: "(SUP x:A. (f x, g x)) = (SUP x:A. f x, SUP x:A. g x)" + by (simp add: SUPR_def Sup_prod_def image_image) + +lemma INF_Pair: "(INF x:A. (f x, g x)) = (INF x:A. f x, INF x:A. g x)" + by (simp add: INFI_def Inf_prod_def image_image) + +end diff -r 421f8bc19ce4 -r b9839fad3bb6 src/HOL/Library/ROOT.ML --- a/src/HOL/Library/ROOT.ML Sun Jul 31 11:13:38 2011 -0700 +++ b/src/HOL/Library/ROOT.ML Mon Aug 01 09:31:10 2011 -0700 @@ -2,4 +2,5 @@ (* Classical Higher-order Logic -- batteries included *) use_thys ["Library", "List_Prefix", "List_lexord", "Sublist_Order", + "Product_Lattice", "Code_Char_chr", "Code_Char_ord", "Code_Integer", "Efficient_Nat", "Executable_Set"(*, "Code_Prolog"*)];