# HG changeset patch # User haftmann # Date 1360847695 -3600 # Node ID 7dbd6832a689d71bdbdfc845312dcc2232285150 # Parent 3e913a575dc677215aafc07aa63b44d0e7b68e09 consolidation of library theories on product orders diff -r 3e913a575dc6 -r 7dbd6832a689 NEWS --- a/NEWS Thu Feb 14 12:24:56 2013 +0100 +++ b/NEWS Thu Feb 14 14:14:55 2013 +0100 @@ -10,6 +10,15 @@ (lin)order_topology. Allows to generalize theorems about limits and order. Instances are reals and extended reals. +*** HOL *** + +* Consolidation of library theories on product orders: + + Product_Lattice ~> Product_Order -- pointwise order on products + Product_ord ~> Product_Lexorder -- lexicographic order on products + +INCOMPATIBILITY. + New in Isabelle2013 (February 2013) ----------------------------------- diff -r 3e913a575dc6 -r 7dbd6832a689 src/HOL/Library/Finite_Lattice.thy --- a/src/HOL/Library/Finite_Lattice.thy Thu Feb 14 12:24:56 2013 +0100 +++ b/src/HOL/Library/Finite_Lattice.thy Thu Feb 14 14:14:55 2013 +0100 @@ -1,7 +1,7 @@ (* Author: Alessandro Coglio *) theory Finite_Lattice -imports Product_Lattice +imports Product_Order begin text {* A non-empty finite lattice is a complete lattice. diff -r 3e913a575dc6 -r 7dbd6832a689 src/HOL/Library/Product_Lattice.thy --- a/src/HOL/Library/Product_Lattice.thy Thu Feb 14 12:24:56 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,238 +0,0 @@ -(* Title: HOL/Library/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_lower SUP_upper 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: SUP_def fst_Sup image_image) - -lemma snd_SUP: "snd (SUP x:A. f x) = (SUP x:A. snd (f x))" - by (simp add: SUP_def snd_Sup image_image) - -lemma fst_INF: "fst (INF x:A. f x) = (INF x:A. fst (f x))" - by (simp add: INF_def fst_Inf image_image) - -lemma snd_INF: "snd (INF x:A. f x) = (INF x:A. snd (f x))" - by (simp add: INF_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: SUP_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: INF_def Inf_prod_def image_image) - - -text {* Alternative formulations for set infima and suprema over the product -of two complete lattices: *} - -lemma Inf_prod_alt_def: "Inf A = (Inf (fst ` A), Inf (snd ` A))" -by (auto simp: Inf_prod_def INF_def) - -lemma Sup_prod_alt_def: "Sup A = (Sup (fst ` A), Sup (snd ` A))" -by (auto simp: Sup_prod_def SUP_def) - -lemma INFI_prod_alt_def: "INFI A f = (INFI A (fst o f), INFI A (snd o f))" -by (auto simp: INF_def Inf_prod_def image_compose) - -lemma SUPR_prod_alt_def: "SUPR A f = (SUPR A (fst o f), SUPR A (snd o f))" -by (auto simp: SUP_def Sup_prod_def image_compose) - -lemma INF_prod_alt_def: - "(INF x:A. f x) = (INF x:A. fst (f x), INF x:A. snd (f x))" -by (metis fst_INF snd_INF surjective_pairing) - -lemma SUP_prod_alt_def: - "(SUP x:A. f x) = (SUP x:A. fst (f x), SUP x:A. snd (f x))" -by (metis fst_SUP snd_SUP surjective_pairing) - - -subsection {* Complete distributive lattices *} - -(* Contribution: Alessandro Coglio *) - -instance prod :: - (complete_distrib_lattice, complete_distrib_lattice) complete_distrib_lattice -proof - case goal1 thus ?case - by (auto simp: sup_prod_def Inf_prod_def INF_prod_alt_def sup_Inf sup_INF) -next - case goal2 thus ?case - by (auto simp: inf_prod_def Sup_prod_def SUP_prod_alt_def inf_Sup inf_SUP) -qed - - -end diff -r 3e913a575dc6 -r 7dbd6832a689 src/HOL/Library/Product_Lexorder.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Product_Lexorder.thy Thu Feb 14 14:14:55 2013 +0100 @@ -0,0 +1,125 @@ +(* Title: HOL/Library/Product_Lexorder.thy + Author: Norbert Voelker +*) + +header {* Lexicographic order on product types *} + +theory Product_Lexorder +imports Main +begin + +instantiation prod :: (ord, ord) ord +begin + +definition + "x \ y \ fst x < fst y \ fst x \ fst y \ snd x \ snd y" + +definition + "x < y \ fst x < fst y \ fst x \ fst y \ snd x < snd y" + +instance .. + +end + +lemma less_eq_prod_simp [simp, code]: + "(x1, y1) \ (x2, y2) \ x1 < x2 \ x1 \ x2 \ y1 \ y2" + by (simp add: less_eq_prod_def) + +lemma less_prod_simp [simp, code]: + "(x1, y1) < (x2, y2) \ x1 < x2 \ x1 \ x2 \ y1 < y2" + by (simp add: less_prod_def) + +text {* A stronger version for partial orders. *} + +lemma less_prod_def': + fixes x y :: "'a::order \ 'b::ord" + shows "x < y \ fst x < fst y \ fst x = fst y \ snd x < snd y" + by (auto simp add: less_prod_def le_less) + +instance prod :: (preorder, preorder) preorder + by default (auto simp: less_eq_prod_def less_prod_def less_le_not_le intro: order_trans) + +instance prod :: (order, order) order + by default (auto simp add: less_eq_prod_def) + +instance prod :: (linorder, linorder) linorder + by default (auto simp: less_eq_prod_def) + +instantiation prod :: (linorder, linorder) distrib_lattice +begin + +definition + "(inf :: 'a \ 'b \ _ \ _) = min" + +definition + "(sup :: 'a \ 'b \ _ \ _) = max" + +instance + by default (auto simp add: inf_prod_def sup_prod_def min_max.sup_inf_distrib1) + +end + +instantiation prod :: (bot, bot) bot +begin + +definition + "bot = (bot, bot)" + +instance + by default (auto simp add: bot_prod_def) + +end + +instantiation prod :: (top, top) top +begin + +definition + "top = (top, top)" + +instance + by default (auto simp add: top_prod_def) + +end + +instance prod :: (wellorder, wellorder) wellorder +proof + fix P :: "'a \ 'b \ bool" and z :: "'a \ 'b" + assume P: "\x. (\y. y < x \ P y) \ P x" + show "P z" + proof (induct z) + case (Pair a b) + show "P (a, b)" + proof (induct a arbitrary: b rule: less_induct) + case (less a\<^isub>1) note a\<^isub>1 = this + show "P (a\<^isub>1, b)" + proof (induct b rule: less_induct) + case (less b\<^isub>1) note b\<^isub>1 = this + show "P (a\<^isub>1, b\<^isub>1)" + proof (rule P) + fix p assume p: "p < (a\<^isub>1, b\<^isub>1)" + show "P p" + proof (cases "fst p < a\<^isub>1") + case True + then have "P (fst p, snd p)" by (rule a\<^isub>1) + then show ?thesis by simp + next + case False + with p have 1: "a\<^isub>1 = fst p" and 2: "snd p < b\<^isub>1" + by (simp_all add: less_prod_def') + from 2 have "P (a\<^isub>1, snd p)" by (rule b\<^isub>1) + with 1 show ?thesis by simp + qed + qed + qed + qed + qed +qed + +text {* Legacy lemma bindings *} + +lemmas prod_le_def = less_eq_prod_def +lemmas prod_less_def = less_prod_def +lemmas prod_less_eq = less_prod_def' + +end + diff -r 3e913a575dc6 -r 7dbd6832a689 src/HOL/Library/Product_Order.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Library/Product_Order.thy Thu Feb 14 14:14:55 2013 +0100 @@ -0,0 +1,238 @@ +(* Title: HOL/Library/Product_Order.thy + Author: Brian Huffman +*) + +header {* Pointwise order on product types *} + +theory Product_Order +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_lower SUP_upper 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: SUP_def fst_Sup image_image) + +lemma snd_SUP: "snd (SUP x:A. f x) = (SUP x:A. snd (f x))" + by (simp add: SUP_def snd_Sup image_image) + +lemma fst_INF: "fst (INF x:A. f x) = (INF x:A. fst (f x))" + by (simp add: INF_def fst_Inf image_image) + +lemma snd_INF: "snd (INF x:A. f x) = (INF x:A. snd (f x))" + by (simp add: INF_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: SUP_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: INF_def Inf_prod_def image_image) + + +text {* Alternative formulations for set infima and suprema over the product +of two complete lattices: *} + +lemma Inf_prod_alt_def: "Inf A = (Inf (fst ` A), Inf (snd ` A))" +by (auto simp: Inf_prod_def INF_def) + +lemma Sup_prod_alt_def: "Sup A = (Sup (fst ` A), Sup (snd ` A))" +by (auto simp: Sup_prod_def SUP_def) + +lemma INFI_prod_alt_def: "INFI A f = (INFI A (fst o f), INFI A (snd o f))" +by (auto simp: INF_def Inf_prod_def image_compose) + +lemma SUPR_prod_alt_def: "SUPR A f = (SUPR A (fst o f), SUPR A (snd o f))" +by (auto simp: SUP_def Sup_prod_def image_compose) + +lemma INF_prod_alt_def: + "(INF x:A. f x) = (INF x:A. fst (f x), INF x:A. snd (f x))" +by (metis fst_INF snd_INF surjective_pairing) + +lemma SUP_prod_alt_def: + "(SUP x:A. f x) = (SUP x:A. fst (f x), SUP x:A. snd (f x))" +by (metis fst_SUP snd_SUP surjective_pairing) + + +subsection {* Complete distributive lattices *} + +(* Contribution: Alessandro Coglio *) + +instance prod :: + (complete_distrib_lattice, complete_distrib_lattice) complete_distrib_lattice +proof + case goal1 thus ?case + by (auto simp: sup_prod_def Inf_prod_def INF_prod_alt_def sup_Inf sup_INF) +next + case goal2 thus ?case + by (auto simp: inf_prod_def Sup_prod_def SUP_prod_alt_def inf_Sup inf_SUP) +qed + +end + diff -r 3e913a575dc6 -r 7dbd6832a689 src/HOL/Library/Product_ord.thy --- a/src/HOL/Library/Product_ord.thy Thu Feb 14 12:24:56 2013 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,115 +0,0 @@ -(* Title: HOL/Library/Product_ord.thy - Author: Norbert Voelker -*) - -header {* Order on product types *} - -theory Product_ord -imports Main -begin - -instantiation prod :: (ord, ord) ord -begin - -definition - prod_le_def: "x \ y \ fst x < fst y \ fst x \ fst y \ snd x \ snd y" - -definition - prod_less_def: "x < y \ fst x < fst y \ fst x \ fst y \ snd x < snd y" - -instance .. - -end - -lemma [code]: - "(x1::'a::{ord, equal}, y1) \ (x2, y2) \ x1 < x2 \ x1 \ x2 \ y1 \ y2" - "(x1::'a::{ord, equal}, y1) < (x2, y2) \ x1 < x2 \ x1 \ x2 \ y1 < y2" - unfolding prod_le_def prod_less_def by simp_all - -instance prod :: (preorder, preorder) preorder - by default (auto simp: prod_le_def prod_less_def less_le_not_le intro: order_trans) - -instance prod :: (order, order) order - by default (auto simp add: prod_le_def) - -instance prod :: (linorder, linorder) linorder - by default (auto simp: prod_le_def) - -instantiation prod :: (linorder, linorder) distrib_lattice -begin - -definition - inf_prod_def: "(inf :: 'a \ 'b \ _ \ _) = min" - -definition - sup_prod_def: "(sup :: 'a \ 'b \ _ \ _) = max" - -instance - by default (auto simp add: inf_prod_def sup_prod_def min_max.sup_inf_distrib1) - -end - -instantiation prod :: (bot, bot) bot -begin - -definition - bot_prod_def: "bot = (bot, bot)" - -instance - by default (auto simp add: bot_prod_def prod_le_def) - -end - -instantiation prod :: (top, top) top -begin - -definition - top_prod_def: "top = (top, top)" - -instance - by default (auto simp add: top_prod_def prod_le_def) - -end - -text {* A stronger version of the definition holds for partial orders. *} - -lemma prod_less_eq: - fixes x y :: "'a::order \ 'b::ord" - shows "x < y \ fst x < fst y \ (fst x = fst y \ snd x < snd y)" - unfolding prod_less_def fst_conv snd_conv le_less by auto - -instance prod :: (wellorder, wellorder) wellorder -proof - fix P :: "'a \ 'b \ bool" and z :: "'a \ 'b" - assume P: "\x. (\y. y < x \ P y) \ P x" - show "P z" - proof (induct z) - case (Pair a b) - show "P (a, b)" - proof (induct a arbitrary: b rule: less_induct) - case (less a\<^isub>1) note a\<^isub>1 = this - show "P (a\<^isub>1, b)" - proof (induct b rule: less_induct) - case (less b\<^isub>1) note b\<^isub>1 = this - show "P (a\<^isub>1, b\<^isub>1)" - proof (rule P) - fix p assume p: "p < (a\<^isub>1, b\<^isub>1)" - show "P p" - proof (cases "fst p < a\<^isub>1") - case True - then have "P (fst p, snd p)" by (rule a\<^isub>1) - then show ?thesis by simp - next - case False - with p have 1: "a\<^isub>1 = fst p" and 2: "snd p < b\<^isub>1" - by (simp_all add: prod_less_eq) - from 2 have "P (a\<^isub>1, snd p)" by (rule b\<^isub>1) - with 1 show ?thesis by simp - qed - qed - qed - qed - qed -qed - -end diff -r 3e913a575dc6 -r 7dbd6832a689 src/HOL/Library/RBT_Set.thy --- a/src/HOL/Library/RBT_Set.thy Thu Feb 14 12:24:56 2013 +0100 +++ b/src/HOL/Library/RBT_Set.thy Thu Feb 14 14:14:55 2013 +0100 @@ -5,7 +5,7 @@ header {* Implementation of sets using RBT trees *} theory RBT_Set -imports RBT Product_ord +imports RBT Product_Lexorder begin (* diff -r 3e913a575dc6 -r 7dbd6832a689 src/HOL/ROOT --- a/src/HOL/ROOT Thu Feb 14 12:24:56 2013 +0100 +++ b/src/HOL/ROOT Thu Feb 14 14:14:55 2013 +0100 @@ -26,6 +26,8 @@ Finite_Lattice Code_Char_chr Code_Char_ord + Product_Lexorder + Product_Order Code_Integer Efficient_Nat (* Code_Prolog FIXME cf. 76965c356d2a *)