diff --git a/src/HOL/Conditionally_Complete_Lattices.thy b/src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy +++ b/src/HOL/Conditionally_Complete_Lattices.thy @@ -1,743 +1,799 @@ (* Title: HOL/Conditionally_Complete_Lattices.thy Author: Amine Chaieb and L C Paulson, University of Cambridge Author: Johannes Hölzl, TU München Author: Luke S. Serafin, Carnegie Mellon University *) section \Conditionally-complete Lattices\ theory Conditionally_Complete_Lattices imports Finite_Set Lattices_Big Set_Interval begin +locale preordering_bdd = preordering +begin + +definition bdd :: \'a set \ bool\ + where unfold: \bdd A \ (\M. \x \ A. x \<^bold>\ M)\ + +lemma empty [simp, intro]: + \bdd {}\ + by (simp add: unfold) + +lemma I [intro]: + \bdd A\ if \\x. x \ A \ x \<^bold>\ M\ + using that by (auto simp add: unfold) + +lemma E: + assumes \bdd A\ + obtains M where \\x. x \ A \ x \<^bold>\ M\ + using assms that by (auto simp add: unfold) + +lemma I2: + \bdd (f ` A)\ if \\x. x \ A \ f x \<^bold>\ M\ + using that by (auto simp add: unfold) + +lemma mono: + \bdd A\ if \bdd B\ \A \ B\ + using that by (auto simp add: unfold) + +lemma Int1 [simp]: + \bdd (A \ B)\ if \bdd A\ + using mono that by auto + +lemma Int2 [simp]: + \bdd (A \ B)\ if \bdd B\ + using mono that by auto + +end + context preorder begin -definition "bdd_above A \ (\M. \x \ A. x \ M)" -definition "bdd_below A \ (\m. \x \ A. m \ x)" +sublocale bdd_above: preordering_bdd \(\)\ \(<)\ + defines bdd_above_primitive_def: bdd_above = bdd_above.bdd .. -lemma bdd_aboveI[intro]: "(\x. x \ A \ x \ M) \ bdd_above A" - by (auto simp: bdd_above_def) +sublocale bdd_below: preordering_bdd \(\)\ \(>)\ + defines bdd_below_primitive_def: bdd_below = bdd_below.bdd .. -lemma bdd_belowI[intro]: "(\x. x \ A \ m \ x) \ bdd_below A" - by (auto simp: bdd_below_def) +lemma bdd_above_def: \bdd_above A \ (\M. \x \ A. x \ M)\ + by (fact bdd_above.unfold) + +lemma bdd_below_def: \bdd_below A \ (\M. \x \ A. M \ x)\ + by (fact bdd_below.unfold) + +lemma bdd_aboveI: "(\x. x \ A \ x \ M) \ bdd_above A" + by (fact bdd_above.I) + +lemma bdd_belowI: "(\x. x \ A \ m \ x) \ bdd_below A" + by (fact bdd_below.I) lemma bdd_aboveI2: "(\x. x \ A \ f x \ M) \ bdd_above (f`A)" - by force + by (fact bdd_above.I2) lemma bdd_belowI2: "(\x. x \ A \ m \ f x) \ bdd_below (f`A)" - by force + by (fact bdd_below.I2) -lemma bdd_above_empty [simp, intro]: "bdd_above {}" - unfolding bdd_above_def by auto +lemma bdd_above_empty: "bdd_above {}" + by (fact bdd_above.empty) -lemma bdd_below_empty [simp, intro]: "bdd_below {}" - unfolding bdd_below_def by auto +lemma bdd_below_empty: "bdd_below {}" + by (fact bdd_below.empty) lemma bdd_above_mono: "bdd_above B \ A \ B \ bdd_above A" - by (metis (full_types) bdd_above_def order_class.le_neq_trans psubsetD) + by (fact bdd_above.mono) lemma bdd_below_mono: "bdd_below B \ A \ B \ bdd_below A" - by (metis bdd_below_def order_class.le_neq_trans psubsetD) - -lemma bdd_above_Int1 [simp]: "bdd_above A \ bdd_above (A \ B)" - using bdd_above_mono by auto + by (fact bdd_below.mono) -lemma bdd_above_Int2 [simp]: "bdd_above B \ bdd_above (A \ B)" - using bdd_above_mono by auto +lemma bdd_above_Int1: "bdd_above A \ bdd_above (A \ B)" + by (fact bdd_above.Int1) -lemma bdd_below_Int1 [simp]: "bdd_below A \ bdd_below (A \ B)" - using bdd_below_mono by auto +lemma bdd_above_Int2: "bdd_above B \ bdd_above (A \ B)" + by (fact bdd_above.Int2) -lemma bdd_below_Int2 [simp]: "bdd_below B \ bdd_below (A \ B)" - using bdd_below_mono by auto +lemma bdd_below_Int1: "bdd_below A \ bdd_below (A \ B)" + by (fact bdd_below.Int1) + +lemma bdd_below_Int2: "bdd_below B \ bdd_below (A \ B)" + by (fact bdd_below.Int2) lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}" by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}" by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le) lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}" by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le) lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}" by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}" by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le) lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}" by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le) end -lemma (in order_top) bdd_above_top[simp, intro!]: "bdd_above A" - by (rule bdd_aboveI[of _ top]) simp +context order_top +begin -lemma (in order_bot) bdd_above_bot[simp, intro!]: "bdd_below A" - by (rule bdd_belowI[of _ bot]) simp +lemma bdd_above_top [simp, intro!]: "bdd_above A" + by (rule bdd_aboveI [of _ top]) simp + +end + +context order_bot +begin + +lemma bdd_below_bot [simp, intro!]: "bdd_below A" + by (rule bdd_belowI [of _ bot]) simp + +end lemma bdd_above_image_mono: "mono f \ bdd_above A \ bdd_above (f`A)" by (auto simp: bdd_above_def mono_def) lemma bdd_below_image_mono: "mono f \ bdd_below A \ bdd_below (f`A)" by (auto simp: bdd_below_def mono_def) lemma bdd_above_image_antimono: "antimono f \ bdd_below A \ bdd_above (f`A)" by (auto simp: bdd_above_def bdd_below_def antimono_def) lemma bdd_below_image_antimono: "antimono f \ bdd_above A \ bdd_below (f`A)" by (auto simp: bdd_above_def bdd_below_def antimono_def) lemma fixes X :: "'a::ordered_ab_group_add set" shows bdd_above_uminus[simp]: "bdd_above (uminus ` X) \ bdd_below X" and bdd_below_uminus[simp]: "bdd_below (uminus ` X) \ bdd_above X" using bdd_above_image_antimono[of uminus X] bdd_below_image_antimono[of uminus "uminus`X"] using bdd_below_image_antimono[of uminus X] bdd_above_image_antimono[of uminus "uminus`X"] by (auto simp: antimono_def image_image) context lattice begin lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A" by (auto simp: bdd_above_def intro: le_supI2 sup_ge1) lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A" by (auto simp: bdd_below_def intro: le_infI2 inf_le1) lemma bdd_finite [simp]: assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A" using assms by (induct rule: finite_induct, auto) lemma bdd_above_Un [simp]: "bdd_above (A \ B) = (bdd_above A \ bdd_above B)" proof assume "bdd_above (A \ B)" thus "bdd_above A \ bdd_above B" unfolding bdd_above_def by auto next assume "bdd_above A \ bdd_above B" then obtain a b where "\x\A. x \ a" "\x\B. x \ b" unfolding bdd_above_def by auto hence "\x \ A \ B. x \ sup a b" by (auto intro: Un_iff le_supI1 le_supI2) thus "bdd_above (A \ B)" unfolding bdd_above_def .. qed lemma bdd_below_Un [simp]: "bdd_below (A \ B) = (bdd_below A \ bdd_below B)" proof assume "bdd_below (A \ B)" thus "bdd_below A \ bdd_below B" unfolding bdd_below_def by auto next assume "bdd_below A \ bdd_below B" then obtain a b where "\x\A. a \ x" "\x\B. b \ x" unfolding bdd_below_def by auto hence "\x \ A \ B. inf a b \ x" by (auto intro: Un_iff le_infI1 le_infI2) thus "bdd_below (A \ B)" unfolding bdd_below_def .. qed lemma bdd_above_image_sup[simp]: "bdd_above ((\x. sup (f x) (g x)) ` A) \ bdd_above (f`A) \ bdd_above (g`A)" by (auto simp: bdd_above_def intro: le_supI1 le_supI2) lemma bdd_below_image_inf[simp]: "bdd_below ((\x. inf (f x) (g x)) ` A) \ bdd_below (f`A) \ bdd_below (g`A)" by (auto simp: bdd_below_def intro: le_infI1 le_infI2) lemma bdd_below_UN[simp]: "finite I \ bdd_below (\i\I. A i) = (\i \ I. bdd_below (A i))" by (induction I rule: finite.induct) auto lemma bdd_above_UN[simp]: "finite I \ bdd_above (\i\I. A i) = (\i \ I. bdd_above (A i))" by (induction I rule: finite.induct) auto end text \ To avoid name classes with the \<^class>\complete_lattice\-class we prefix \<^const>\Sup\ and \<^const>\Inf\ in theorem names with c. \ class conditionally_complete_lattice = lattice + Sup + Inf + assumes cInf_lower: "x \ X \ bdd_below X \ Inf X \ x" and cInf_greatest: "X \ {} \ (\x. x \ X \ z \ x) \ z \ Inf X" assumes cSup_upper: "x \ X \ bdd_above X \ x \ Sup X" and cSup_least: "X \ {} \ (\x. x \ X \ x \ z) \ Sup X \ z" begin lemma cSup_upper2: "x \ X \ y \ x \ bdd_above X \ y \ Sup X" by (metis cSup_upper order_trans) lemma cInf_lower2: "x \ X \ x \ y \ bdd_below X \ Inf X \ y" by (metis cInf_lower order_trans) lemma cSup_mono: "B \ {} \ bdd_above A \ (\b. b \ B \ \a\A. b \ a) \ Sup B \ Sup A" by (metis cSup_least cSup_upper2) lemma cInf_mono: "B \ {} \ bdd_below A \ (\b. b \ B \ \a\A. a \ b) \ Inf A \ Inf B" by (metis cInf_greatest cInf_lower2) lemma cSup_subset_mono: "A \ {} \ bdd_above B \ A \ B \ Sup A \ Sup B" by (metis cSup_least cSup_upper subsetD) lemma cInf_superset_mono: "A \ {} \ bdd_below B \ A \ B \ Inf B \ Inf A" by (metis cInf_greatest cInf_lower subsetD) lemma cSup_eq_maximum: "z \ X \ (\x. x \ X \ x \ z) \ Sup X = z" by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto lemma cInf_eq_minimum: "z \ X \ (\x. x \ X \ z \ x) \ Inf X = z" by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto lemma cSup_le_iff: "S \ {} \ bdd_above S \ Sup S \ a \ (\x\S. x \ a)" by (metis order_trans cSup_upper cSup_least) lemma le_cInf_iff: "S \ {} \ bdd_below S \ a \ Inf S \ (\x\S. a \ x)" by (metis order_trans cInf_lower cInf_greatest) lemma cSup_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ x \ a" assumes 3: "\y. (\x. x \ X \ x \ y) \ a \ y" shows "Sup X = a" by (intro 3 1 antisym cSup_least) (auto intro: 2 1 cSup_upper) lemma cInf_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ a \ x" assumes 3: "\y. (\x. x \ X \ y \ x) \ y \ a" shows "Inf X = a" by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower) lemma cInf_cSup: "S \ {} \ bdd_below S \ Inf S = Sup {x. \s\S. x \ s}" by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def) lemma cSup_cInf: "S \ {} \ bdd_above S \ Sup S = Inf {x. \s\S. s \ x}" by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def) lemma cSup_insert: "X \ {} \ bdd_above X \ Sup (insert a X) = sup a (Sup X)" by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least) lemma cInf_insert: "X \ {} \ bdd_below X \ Inf (insert a X) = inf a (Inf X)" by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest) lemma cSup_singleton [simp]: "Sup {x} = x" by (intro cSup_eq_maximum) auto lemma cInf_singleton [simp]: "Inf {x} = x" by (intro cInf_eq_minimum) auto lemma cSup_insert_If: "bdd_above X \ Sup (insert a X) = (if X = {} then a else sup a (Sup X))" using cSup_insert[of X] by simp lemma cInf_insert_If: "bdd_below X \ Inf (insert a X) = (if X = {} then a else inf a (Inf X))" using cInf_insert[of X] by simp lemma le_cSup_finite: "finite X \ x \ X \ x \ Sup X" proof (induct X arbitrary: x rule: finite_induct) case (insert x X y) then show ?case by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2) qed simp lemma cInf_le_finite: "finite X \ x \ X \ Inf X \ x" proof (induct X arbitrary: x rule: finite_induct) case (insert x X y) then show ?case by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2) qed simp lemma cSup_eq_Sup_fin: "finite X \ X \ {} \ Sup X = Sup_fin X" by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert) lemma cInf_eq_Inf_fin: "finite X \ X \ {} \ Inf X = Inf_fin X" by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert) lemma cSup_atMost[simp]: "Sup {..x} = x" by (auto intro!: cSup_eq_maximum) lemma cSup_greaterThanAtMost[simp]: "y < x \ Sup {y<..x} = x" by (auto intro!: cSup_eq_maximum) lemma cSup_atLeastAtMost[simp]: "y \ x \ Sup {y..x} = x" by (auto intro!: cSup_eq_maximum) lemma cInf_atLeast[simp]: "Inf {x..} = x" by (auto intro!: cInf_eq_minimum) lemma cInf_atLeastLessThan[simp]: "y < x \ Inf {y.. x \ Inf {y..x} = y" by (auto intro!: cInf_eq_minimum) lemma cINF_lower: "bdd_below (f ` A) \ x \ A \ \(f ` A) \ f x" using cInf_lower [of _ "f ` A"] by simp lemma cINF_greatest: "A \ {} \ (\x. x \ A \ m \ f x) \ m \ \(f ` A)" using cInf_greatest [of "f ` A"] by auto lemma cSUP_upper: "x \ A \ bdd_above (f ` A) \ f x \ \(f ` A)" using cSup_upper [of _ "f ` A"] by simp lemma cSUP_least: "A \ {} \ (\x. x \ A \ f x \ M) \ \(f ` A) \ M" using cSup_least [of "f ` A"] by auto lemma cINF_lower2: "bdd_below (f ` A) \ x \ A \ f x \ u \ \(f ` A) \ u" by (auto intro: cINF_lower order_trans) lemma cSUP_upper2: "bdd_above (f ` A) \ x \ A \ u \ f x \ u \ \(f ` A)" by (auto intro: cSUP_upper order_trans) lemma cSUP_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro antisym cSUP_least) (auto intro: cSUP_upper) lemma cINF_const [simp]: "A \ {} \ (\x\A. c) = c" by (intro antisym cINF_greatest) (auto intro: cINF_lower) lemma le_cINF_iff: "A \ {} \ bdd_below (f ` A) \ u \ \(f ` A) \ (\x\A. u \ f x)" by (metis cINF_greatest cINF_lower order_trans) lemma cSUP_le_iff: "A \ {} \ bdd_above (f ` A) \ \(f ` A) \ u \ (\x\A. f x \ u)" by (metis cSUP_least cSUP_upper order_trans) lemma less_cINF_D: "bdd_below (f`A) \ y < (\i\A. f i) \ i \ A \ y < f i" by (metis cINF_lower less_le_trans) lemma cSUP_lessD: "bdd_above (f`A) \ (\i\A. f i) < y \ i \ A \ f i < y" by (metis cSUP_upper le_less_trans) lemma cINF_insert: "A \ {} \ bdd_below (f ` A) \ \(f ` insert a A) = inf (f a) (\(f ` A))" by (simp add: cInf_insert) lemma cSUP_insert: "A \ {} \ bdd_above (f ` A) \ \(f ` insert a A) = sup (f a) (\(f ` A))" by (simp add: cSup_insert) lemma cINF_mono: "B \ {} \ bdd_below (f ` A) \ (\m. m \ B \ \n\A. f n \ g m) \ \(f ` A) \ \(g ` B)" using cInf_mono [of "g ` B" "f ` A"] by auto lemma cSUP_mono: "A \ {} \ bdd_above (g ` B) \ (\n. n \ A \ \m\B. f n \ g m) \ \(f ` A) \ \(g ` B)" using cSup_mono [of "f ` A" "g ` B"] by auto lemma cINF_superset_mono: "A \ {} \ bdd_below (g ` B) \ A \ B \ (\x. x \ B \ g x \ f x) \ \(g ` B) \ \(f ` A)" by (rule cINF_mono) auto lemma cSUP_subset_mono: "\A \ {}; bdd_above (g ` B); A \ B; \x. x \ A \ f x \ g x\ \ \ (f ` A) \ \ (g ` B)" by (rule cSUP_mono) auto lemma less_eq_cInf_inter: "bdd_below A \ bdd_below B \ A \ B \ {} \ inf (Inf A) (Inf B) \ Inf (A \ B)" by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1) lemma cSup_inter_less_eq: "bdd_above A \ bdd_above B \ A \ B \ {} \ Sup (A \ B) \ sup (Sup A) (Sup B) " by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1) lemma cInf_union_distrib: "A \ {} \ bdd_below A \ B \ {} \ bdd_below B \ Inf (A \ B) = inf (Inf A) (Inf B)" by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower) lemma cINF_union: "A \ {} \ bdd_below (f ` A) \ B \ {} \ bdd_below (f ` B) \ \ (f ` (A \ B)) = \ (f ` A) \ \ (f ` B)" using cInf_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un) lemma cSup_union_distrib: "A \ {} \ bdd_above A \ B \ {} \ bdd_above B \ Sup (A \ B) = sup (Sup A) (Sup B)" by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper) lemma cSUP_union: "A \ {} \ bdd_above (f ` A) \ B \ {} \ bdd_above (f ` B) \ \ (f ` (A \ B)) = \ (f ` A) \ \ (f ` B)" using cSup_union_distrib [of "f ` A" "f ` B"] by (simp add: image_Un) lemma cINF_inf_distrib: "A \ {} \ bdd_below (f`A) \ bdd_below (g`A) \ \ (f ` A) \ \ (g ` A) = (\a\A. inf (f a) (g a))" by (intro antisym le_infI cINF_greatest cINF_lower2) (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI) lemma SUP_sup_distrib: "A \ {} \ bdd_above (f`A) \ bdd_above (g`A) \ \ (f ` A) \ \ (g ` A) = (\a\A. sup (f a) (g a))" by (intro antisym le_supI cSUP_least cSUP_upper2) (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI) lemma cInf_le_cSup: "A \ {} \ bdd_above A \ bdd_below A \ Inf A \ Sup A" by (auto intro!: cSup_upper2[of "SOME a. a \ A"] intro: someI cInf_lower) end instance complete_lattice \ conditionally_complete_lattice by standard (auto intro: Sup_upper Sup_least Inf_lower Inf_greatest) lemma cSup_eq: fixes a :: "'a :: {conditionally_complete_lattice, no_bot}" assumes upper: "\x. x \ X \ x \ a" assumes least: "\y. (\x. x \ X \ x \ y) \ a \ y" shows "Sup X = a" proof cases assume "X = {}" with lt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cSup_eq_non_empty assms) lemma cInf_eq: fixes a :: "'a :: {conditionally_complete_lattice, no_top}" assumes upper: "\x. x \ X \ a \ x" assumes least: "\y. (\x. x \ X \ y \ x) \ y \ a" shows "Inf X = a" proof cases assume "X = {}" with gt_ex[of a] least show ?thesis by (auto simp: less_le_not_le) qed (intro cInf_eq_non_empty assms) class conditionally_complete_linorder = conditionally_complete_lattice + linorder begin lemma less_cSup_iff: "X \ {} \ bdd_above X \ y < Sup X \ (\x\X. y < x)" by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans) lemma cInf_less_iff: "X \ {} \ bdd_below X \ Inf X < y \ (\x\X. x < y)" by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans) lemma cINF_less_iff: "A \ {} \ bdd_below (f`A) \ (\i\A. f i) < a \ (\x\A. f x < a)" using cInf_less_iff[of "f`A"] by auto lemma less_cSUP_iff: "A \ {} \ bdd_above (f`A) \ a < (\i\A. f i) \ (\x\A. a < f x)" using less_cSup_iff[of "f`A"] by auto lemma less_cSupE: assumes "y < Sup X" "X \ {}" obtains x where "x \ X" "y < x" by (metis cSup_least assms not_le that) lemma less_cSupD: "X \ {} \ z < Sup X \ \x\X. z < x" by (metis less_cSup_iff not_le_imp_less bdd_above_def) lemma cInf_lessD: "X \ {} \ Inf X < z \ \x\X. x < z" by (metis cInf_less_iff not_le_imp_less bdd_below_def) lemma complete_interval: assumes "a < b" and "P a" and "\ P b" shows "\c. a \ c \ c \ b \ (\x. a \ x \ x < c \ P x) \ (\d. (\x. a \ x \ x < d \ P x) \ d \ c)" proof (rule exI [where x = "Sup {d. \x. a \ x \ x < d \ P x}"], auto) show "a \ Sup {d. \c. a \ c \ c < d \ P c}" by (rule cSup_upper, auto simp: bdd_above_def) (metis \a < b\ \\ P b\ linear less_le) next show "Sup {d. \c. a \ c \ c < d \ P c} \ b" apply (rule cSup_least) apply auto apply (metis less_le_not_le) apply (metis \a \\ P b\ linear less_le) done next fix x assume x: "a \ x" and lt: "x < Sup {d. \c. a \ c \ c < d \ P c}" show "P x" apply (rule less_cSupE [OF lt], auto) apply (metis less_le_not_le) apply (metis x) done next fix d assume 0: "\x. a \ x \ x < d \ P x" thus "d \ Sup {d. \c. a \ c \ c < d \ P c}" by (rule_tac cSup_upper, auto simp: bdd_above_def) (metis \a \\ P b\ linear less_le) qed end instance complete_linorder < conditionally_complete_linorder .. lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Sup X = Max X" using cSup_eq_Sup_fin[of X] by (simp add: Sup_fin_Max) lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \ X \ {} \ Inf X = Min X" using cInf_eq_Inf_fin[of X] by (simp add: Inf_fin_Min) lemma cSup_lessThan[simp]: "Sup {.. Sup {y<.. Sup {y.. Inf {y<..x::'a::{conditionally_complete_linorder, dense_linorder}} = y" by (auto intro!: cInf_eq_non_empty intro: dense_ge_bounded) lemma cInf_greaterThanLessThan[simp]: "y < x \ Inf {y<.. Inf (insert x S) = (if S = {} then x else min x (Inf S))" by (simp add: cInf_eq_Min) lemma Sup_insert_finite: fixes S :: "'a::conditionally_complete_linorder set" shows "finite S \ Sup (insert x S) = (if S = {} then x else max x (Sup S))" by (simp add: cSup_insert sup_max) lemma finite_imp_less_Inf: fixes a :: "'a::conditionally_complete_linorder" shows "\finite X; x \ X; \x. x\X \ a < x\ \ a < Inf X" by (induction X rule: finite_induct) (simp_all add: cInf_eq_Min Inf_insert_finite) lemma finite_less_Inf_iff: fixes a :: "'a :: conditionally_complete_linorder" shows "\finite X; X \ {}\ \ a < Inf X \ (\x \ X. a < x)" by (auto simp: cInf_eq_Min) lemma finite_imp_Sup_less: fixes a :: "'a::conditionally_complete_linorder" shows "\finite X; x \ X; \x. x\X \ a > x\ \ a > Sup X" by (induction X rule: finite_induct) (simp_all add: cSup_eq_Max Sup_insert_finite) lemma finite_Sup_less_iff: fixes a :: "'a :: conditionally_complete_linorder" shows "\finite X; X \ {}\ \ a > Sup X \ (\x \ X. a > x)" by (auto simp: cSup_eq_Max) class linear_continuum = conditionally_complete_linorder + dense_linorder + assumes UNIV_not_singleton: "\a b::'a. a \ b" begin lemma ex_gt_or_lt: "\b. a < b \ b < a" by (metis UNIV_not_singleton neq_iff) end instantiation nat :: conditionally_complete_linorder begin definition "Sup (X::nat set) = (if X={} then 0 else Max X)" definition "Inf (X::nat set) = (LEAST n. n \ X)" lemma bdd_above_nat: "bdd_above X \ finite (X::nat set)" proof assume "bdd_above X" then obtain z where "X \ {.. z}" by (auto simp: bdd_above_def) then show "finite X" by (rule finite_subset) simp qed simp instance proof fix x :: nat fix X :: "nat set" show "Inf X \ x" if "x \ X" "bdd_below X" using that by (simp add: Inf_nat_def Least_le) show "x \ Inf X" if "X \ {}" "\y. y \ X \ x \ y" using that unfolding Inf_nat_def ex_in_conv[symmetric] by (rule LeastI2_ex) show "x \ Sup X" if "x \ X" "bdd_above X" using that by (auto simp add: Sup_nat_def bdd_above_nat) show "Sup X \ x" if "X \ {}" "\y. y \ X \ y \ x" proof - from that have "bdd_above X" by (auto simp: bdd_above_def) with that show ?thesis by (simp add: Sup_nat_def bdd_above_nat) qed qed end lemma Inf_nat_def1: fixes K::"nat set" assumes "K \ {}" shows "Inf K \ K" by (auto simp add: Min_def Inf_nat_def) (meson LeastI assms bot.extremum_unique subsetI) lemma Sup_nat_empty [simp]: "Sup {} = (0::nat)" by (auto simp add: Sup_nat_def) instantiation int :: conditionally_complete_linorder begin definition "Sup (X::int set) = (THE x. x \ X \ (\y\X. y \ x))" definition "Inf (X::int set) = - (Sup (uminus ` X))" instance proof { fix x :: int and X :: "int set" assume "X \ {}" "bdd_above X" then obtain x y where "X \ {..y}" "x \ X" by (auto simp: bdd_above_def) then have *: "finite (X \ {x..y})" "X \ {x..y} \ {}" and "x \ y" by (auto simp: subset_eq) have "\!x\X. (\y\X. y \ x)" proof { fix z assume "z \ X" have "z \ Max (X \ {x..y})" proof cases assume "x \ z" with \z \ X\ \X \ {..y}\ *(1) show ?thesis by (auto intro!: Max_ge) next assume "\ x \ z" then have "z < x" by simp also have "x \ Max (X \ {x..y})" using \x \ X\ *(1) \x \ y\ by (intro Max_ge) auto finally show ?thesis by simp qed } note le = this with Max_in[OF *] show ex: "Max (X \ {x..y}) \ X \ (\z\X. z \ Max (X \ {x..y}))" by auto fix z assume *: "z \ X \ (\y\X. y \ z)" with le have "z \ Max (X \ {x..y})" by auto moreover have "Max (X \ {x..y}) \ z" using * ex by auto ultimately show "z = Max (X \ {x..y})" by auto qed then have "Sup X \ X \ (\y\X. y \ Sup X)" unfolding Sup_int_def by (rule theI') } note Sup_int = this { fix x :: int and X :: "int set" assume "x \ X" "bdd_above X" then show "x \ Sup X" using Sup_int[of X] by auto } note le_Sup = this { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ y \ x" then show "Sup X \ x" using Sup_int[of X] by (auto simp: bdd_above_def) } note Sup_le = this { fix x :: int and X :: "int set" assume "x \ X" "bdd_below X" then show "Inf X \ x" using le_Sup[of "-x" "uminus ` X"] by (auto simp: Inf_int_def) } { fix x :: int and X :: "int set" assume "X \ {}" "\y. y \ X \ x \ y" then show "x \ Inf X" using Sup_le[of "uminus ` X" "-x"] by (force simp: Inf_int_def) } qed end lemma interval_cases: fixes S :: "'a :: conditionally_complete_linorder set" assumes ivl: "\a b x. a \ S \ b \ S \ a \ x \ x \ b \ x \ S" shows "\a b. S = {} \ S = UNIV \ S = {.. S = {..b} \ S = {a<..} \ S = {a..} \ S = {a<.. S = {a<..b} \ S = {a.. S = {a..b}" proof - define lower upper where "lower = {x. \s\S. s \ x}" and "upper = {x. \s\S. x \ s}" with ivl have "S = lower \ upper" by auto moreover have "\a. upper = UNIV \ upper = {} \ upper = {.. a} \ upper = {..< a}" proof cases assume *: "bdd_above S \ S \ {}" from * have "upper \ {.. Sup S}" by (auto simp: upper_def intro: cSup_upper2) moreover from * have "{..< Sup S} \ upper" by (force simp add: less_cSup_iff upper_def subset_eq Ball_def) ultimately have "upper = {.. Sup S} \ upper = {..< Sup S}" unfolding ivl_disj_un(2)[symmetric] by auto then show ?thesis by auto next assume "\ (bdd_above S \ S \ {})" then have "upper = UNIV \ upper = {}" by (auto simp: upper_def bdd_above_def not_le dest: less_imp_le) then show ?thesis by auto qed moreover have "\b. lower = UNIV \ lower = {} \ lower = {b ..} \ lower = {b <..}" proof cases assume *: "bdd_below S \ S \ {}" from * have "lower \ {Inf S ..}" by (auto simp: lower_def intro: cInf_lower2) moreover from * have "{Inf S <..} \ lower" by (force simp add: cInf_less_iff lower_def subset_eq Ball_def) ultimately have "lower = {Inf S ..} \ lower = {Inf S <..}" unfolding ivl_disj_un(1)[symmetric] by auto then show ?thesis by auto next assume "\ (bdd_below S \ S \ {})" then have "lower = UNIV \ lower = {}" by (auto simp: lower_def bdd_below_def not_le dest: less_imp_le) then show ?thesis by auto qed ultimately show ?thesis unfolding greaterThanAtMost_def greaterThanLessThan_def atLeastAtMost_def atLeastLessThan_def by (metis inf_bot_left inf_bot_right inf_top.left_neutral inf_top.right_neutral) qed lemma cSUP_eq_cINF_D: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes eq: "(\x\A. f x) = (\x\A. f x)" and bdd: "bdd_above (f ` A)" "bdd_below (f ` A)" and a: "a \ A" shows "f a = (\x\A. f x)" apply (rule antisym) using a bdd apply (auto simp: cINF_lower) apply (metis eq cSUP_upper) done lemma cSUP_UNION: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_above (\x\A. f ` B x)" shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_above (f ` B x)" using bdd_UN by (meson UN_upper bdd_above_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_above_def) then have bdd2: "bdd_above ((\x. \z\B x. f z) ` A)" unfolding bdd_above_def by (force simp: bdd cSUP_le_iff ne(2)) have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper2 simp: bdd2 bdd) moreover have "(\x\A. \z\B x. f z) \ (\ z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper simp: image_UN bdd_UN) ultimately show ?thesis by (rule order_antisym) qed lemma cINF_UNION: fixes f :: "_ \ 'b::conditionally_complete_lattice" assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" and bdd_UN: "bdd_below (\x\A. f ` B x)" shows "(\z \ \x\A. B x. f z) = (\x\A. \z\B x. f z)" proof - have bdd: "\x. x \ A \ bdd_below (f ` B x)" using bdd_UN by (meson UN_upper bdd_below_mono) obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" using bdd_UN by (auto simp: bdd_below_def) then have bdd2: "bdd_below ((\x. \z\B x. f z) ` A)" unfolding bdd_below_def by (force simp: bdd le_cINF_iff ne(2)) have "(\z \ \x\A. B x. f z) \ (\x\A. \z\B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower simp: bdd2 bdd) moreover have "(\x\A. \z\B x. f z) \ (\z \ \x\A. B x. f z)" using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower2 simp: bdd bdd_UN bdd2) ultimately show ?thesis by (rule order_antisym) qed lemma cSup_abs_le: fixes S :: "('a::{linordered_idom,conditionally_complete_linorder}) set" shows "S \ {} \ (\x. x\S \ \x\ \ a) \ \Sup S\ \ a" apply (auto simp add: abs_le_iff intro: cSup_least) by (metis bdd_aboveI cSup_upper neg_le_iff_le order_trans) end diff --git a/src/HOL/Orderings.thy b/src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy +++ b/src/HOL/Orderings.thy @@ -1,1745 +1,1801 @@ (* Title: HOL/Orderings.thy Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson *) section \Abstract orderings\ theory Orderings imports HOL keywords "print_orders" :: diag begin ML_file \~~/src/Provers/order.ML\ subsection \Abstract ordering\ -locale ordering = - fixes less_eq :: "'a \ 'a \ bool" (infix "\<^bold>\" 50) - and less :: "'a \ 'a \ bool" (infix "\<^bold><" 50) - assumes strict_iff_order: "a \<^bold>< b \ a \<^bold>\ b \ a \ b" - assumes refl: "a \<^bold>\ a" \ \not \iff\: makes problems due to multiple (dual) interpretations\ - and antisym: "a \<^bold>\ b \ b \<^bold>\ a \ a = b" - and trans: "a \<^bold>\ b \ b \<^bold>\ c \ a \<^bold>\ c" +locale partial_preordering = + fixes less_eq :: \'a \ 'a \ bool\ (infix \\<^bold>\\ 50) + assumes refl: \a \<^bold>\ a\ \ \not \iff\: makes problems due to multiple (dual) interpretations\ + and trans: \a \<^bold>\ b \ b \<^bold>\ c \ a \<^bold>\ c\ + +locale preordering = partial_preordering + + fixes less :: \'a \ 'a \ bool\ (infix \\<^bold><\ 50) + assumes strict_iff_not: \a \<^bold>< b \ a \<^bold>\ b \ \ b \<^bold>\ a\ begin lemma strict_implies_order: - "a \<^bold>< b \ a \<^bold>\ b" - by (simp add: strict_iff_order) + \a \<^bold>< b \ a \<^bold>\ b\ + by (simp add: strict_iff_not) + +lemma irrefl: \ \not \iff\: makes problems due to multiple (dual) interpretations\ + \\ a \<^bold>< a\ + by (simp add: strict_iff_not) + +lemma asym: + \a \<^bold>< b \ b \<^bold>< a \ False\ + by (auto simp add: strict_iff_not) + +lemma strict_trans1: + \a \<^bold>\ b \ b \<^bold>< c \ a \<^bold>< c\ + by (auto simp add: strict_iff_not intro: trans) + +lemma strict_trans2: + \a \<^bold>< b \ b \<^bold>\ c \ a \<^bold>< c\ + by (auto simp add: strict_iff_not intro: trans) + +lemma strict_trans: + \a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c\ + by (auto intro: strict_trans1 strict_implies_order) + +end + +lemma preordering_strictI: \ \Alternative introduction rule with bias towards strict order\ + fixes less_eq (infix \\<^bold>\\ 50) + and less (infix \\<^bold><\ 50) + assumes less_eq_less: \\a b. a \<^bold>\ b \ a \<^bold>< b \ a = b\ + assumes asym: \\a b. a \<^bold>< b \ \ b \<^bold>< a\ + assumes irrefl: \\a. \ a \<^bold>< a\ + assumes trans: \\a b c. a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c\ + shows \preordering (\<^bold>\) (\<^bold><)\ +proof + fix a b + show \a \<^bold>< b \ a \<^bold>\ b \ \ b \<^bold>\ a\ + by (auto simp add: less_eq_less asym irrefl) +next + fix a + show \a \<^bold>\ a\ + by (auto simp add: less_eq_less) +next + fix a b c + assume \a \<^bold>\ b\ and \b \<^bold>\ c\ then show \a \<^bold>\ c\ + by (auto simp add: less_eq_less intro: trans) +qed + +lemma preordering_dualI: + fixes less_eq (infix \\<^bold>\\ 50) + and less (infix \\<^bold><\ 50) + assumes \preordering (\a b. b \<^bold>\ a) (\a b. b \<^bold>< a)\ + shows \preordering (\<^bold>\) (\<^bold><)\ +proof - + from assms interpret preordering \\a b. b \<^bold>\ a\ \\a b. b \<^bold>< a\ . + show ?thesis + by standard (auto simp: strict_iff_not refl intro: trans) +qed + +locale ordering = partial_preordering + + fixes less :: \'a \ 'a \ bool\ (infix \\<^bold><\ 50) + assumes strict_iff_order: \a \<^bold>< b \ a \<^bold>\ b \ a \ b\ + assumes antisym: \a \<^bold>\ b \ b \<^bold>\ a \ a = b\ +begin + +sublocale preordering \(\<^bold>\)\ \(\<^bold><)\ +proof + show \a \<^bold>< b \ a \<^bold>\ b \ \ b \<^bold>\ a\ for a b + by (auto simp add: strict_iff_order intro: antisym) +qed lemma strict_implies_not_eq: - "a \<^bold>< b \ a \ b" + \a \<^bold>< b \ a \ b\ by (simp add: strict_iff_order) lemma not_eq_order_implies_strict: - "a \ b \ a \<^bold>\ b \ a \<^bold>< b" + \a \ b \ a \<^bold>\ b \ a \<^bold>< b\ by (simp add: strict_iff_order) lemma order_iff_strict: - "a \<^bold>\ b \ a \<^bold>< b \ a = b" + \a \<^bold>\ b \ a \<^bold>< b \ a = b\ by (auto simp add: strict_iff_order refl) -lemma irrefl: \ \not \iff\: makes problems due to multiple (dual) interpretations\ - "\ a \<^bold>< a" - by (simp add: strict_iff_order) - -lemma asym: - "a \<^bold>< b \ b \<^bold>< a \ False" - by (auto simp add: strict_iff_order intro: antisym) - -lemma strict_trans1: - "a \<^bold>\ b \ b \<^bold>< c \ a \<^bold>< c" - by (auto simp add: strict_iff_order intro: trans antisym) - -lemma strict_trans2: - "a \<^bold>< b \ b \<^bold>\ c \ a \<^bold>< c" - by (auto simp add: strict_iff_order intro: trans antisym) - -lemma strict_trans: - "a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c" - by (auto intro: strict_trans1 strict_implies_order) - -lemma eq_iff: "a = b \ a \<^bold>\ b \ b \<^bold>\ a" +lemma eq_iff: \a = b \ a \<^bold>\ b \ b \<^bold>\ a\ by (auto simp add: refl intro: antisym) end -text \Alternative introduction rule with bias towards strict order\ - -lemma ordering_strictI: - fixes less_eq (infix "\<^bold>\" 50) - and less (infix "\<^bold><" 50) - assumes less_eq_less: "\a b. a \<^bold>\ b \ a \<^bold>< b \ a = b" - assumes asym: "\a b. a \<^bold>< b \ \ b \<^bold>< a" - assumes irrefl: "\a. \ a \<^bold>< a" - assumes trans: "\a b c. a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c" - shows "ordering less_eq less" +lemma ordering_strictI: \ \Alternative introduction rule with bias towards strict order\ + fixes less_eq (infix \\<^bold>\\ 50) + and less (infix \\<^bold><\ 50) + assumes less_eq_less: \\a b. a \<^bold>\ b \ a \<^bold>< b \ a = b\ + assumes asym: \\a b. a \<^bold>< b \ \ b \<^bold>< a\ + assumes irrefl: \\a. \ a \<^bold>< a\ + assumes trans: \\a b c. a \<^bold>< b \ b \<^bold>< c \ a \<^bold>< c\ + shows \ordering (\<^bold>\) (\<^bold><)\ proof fix a b - show "a \<^bold>< b \ a \<^bold>\ b \ a \ b" + show \a \<^bold>< b \ a \<^bold>\ b \ a \ b\ by (auto simp add: less_eq_less asym irrefl) next fix a - show "a \<^bold>\ a" + show \a \<^bold>\ a\ by (auto simp add: less_eq_less) next fix a b c - assume "a \<^bold>\ b" and "b \<^bold>\ c" then show "a \<^bold>\ c" + assume \a \<^bold>\ b\ and \b \<^bold>\ c\ then show \a \<^bold>\ c\ by (auto simp add: less_eq_less intro: trans) next fix a b - assume "a \<^bold>\ b" and "b \<^bold>\ a" then show "a = b" + assume \a \<^bold>\ b\ and \b \<^bold>\ a\ then show \a = b\ by (auto simp add: less_eq_less asym) qed lemma ordering_dualI: - fixes less_eq (infix "\<^bold>\" 50) - and less (infix "\<^bold><" 50) - assumes "ordering (\a b. b \<^bold>\ a) (\a b. b \<^bold>< a)" - shows "ordering less_eq less" + fixes less_eq (infix \\<^bold>\\ 50) + and less (infix \\<^bold><\ 50) + assumes \ordering (\a b. b \<^bold>\ a) (\a b. b \<^bold>< a)\ + shows \ordering (\<^bold>\) (\<^bold><)\ proof - - from assms interpret ordering "\a b. b \<^bold>\ a" "\a b. b \<^bold>< a" . + from assms interpret ordering \\a b. b \<^bold>\ a\ \\a b. b \<^bold>< a\ . show ?thesis by standard (auto simp: strict_iff_order refl intro: antisym trans) qed locale ordering_top = ordering + - fixes top :: "'a" ("\<^bold>\") - assumes extremum [simp]: "a \<^bold>\ \<^bold>\" + fixes top :: \'a\ (\\<^bold>\\) + assumes extremum [simp]: \a \<^bold>\ \<^bold>\\ begin lemma extremum_uniqueI: - "\<^bold>\ \<^bold>\ a \ a = \<^bold>\" + \\<^bold>\ \<^bold>\ a \ a = \<^bold>\\ by (rule antisym) auto lemma extremum_unique: - "\<^bold>\ \<^bold>\ a \ a = \<^bold>\" + \\<^bold>\ \<^bold>\ a \ a = \<^bold>\\ by (auto intro: antisym) lemma extremum_strict [simp]: - "\ (\<^bold>\ \<^bold>< a)" + \\ (\<^bold>\ \<^bold>< a)\ using extremum [of a] by (auto simp add: order_iff_strict intro: asym irrefl) lemma not_eq_extremum: - "a \ \<^bold>\ \ a \<^bold>< \<^bold>\" + \a \ \<^bold>\ \ a \<^bold>< \<^bold>\\ by (auto simp add: order_iff_strict intro: not_eq_order_implies_strict extremum) end subsection \Syntactic orders\ class ord = fixes less_eq :: "'a \ 'a \ bool" and less :: "'a \ 'a \ bool" begin notation less_eq ("'(\')") and less_eq ("(_/ \ _)" [51, 51] 50) and less ("'(<')") and less ("(_/ < _)" [51, 51] 50) abbreviation (input) greater_eq (infix "\" 50) where "x \ y \ y \ x" abbreviation (input) greater (infix ">" 50) where "x > y \ y < x" notation (ASCII) less_eq ("'(<=')") and less_eq ("(_/ <= _)" [51, 51] 50) notation (input) greater_eq (infix ">=" 50) end subsection \Quasi orders\ class preorder = ord + assumes less_le_not_le: "x < y \ x \ y \ \ (y \ x)" and order_refl [iff]: "x \ x" and order_trans: "x \ y \ y \ z \ x \ z" begin +sublocale order: preordering less_eq less + dual_order: preordering greater_eq greater +proof - + interpret preordering less_eq less + by standard (auto intro: order_trans simp add: less_le_not_le) + show \preordering less_eq less\ + by (fact preordering_axioms) + then show \preordering greater_eq greater\ + by (rule preordering_dualI) +qed + text \Reflexivity.\ lemma eq_refl: "x = y \ x \ y" \ \This form is useful with the classical reasoner.\ by (erule ssubst) (rule order_refl) lemma less_irrefl [iff]: "\ x < x" by (simp add: less_le_not_le) lemma less_imp_le: "x < y \ x \ y" by (simp add: less_le_not_le) text \Asymmetry.\ lemma less_not_sym: "x < y \ \ (y < x)" by (simp add: less_le_not_le) lemma less_asym: "x < y \ (\ P \ y < x) \ P" by (drule less_not_sym, erule contrapos_np) simp text \Transitivity.\ lemma less_trans: "x < y \ y < z \ x < z" by (auto simp add: less_le_not_le intro: order_trans) lemma le_less_trans: "x \ y \ y < z \ x < z" by (auto simp add: less_le_not_le intro: order_trans) lemma less_le_trans: "x < y \ y \ z \ x < z" by (auto simp add: less_le_not_le intro: order_trans) text \Useful for simplification, but too risky to include by default.\ lemma less_imp_not_less: "x < y \ (\ y < x) \ True" by (blast elim: less_asym) lemma less_imp_triv: "x < y \ (y < x \ P) \ True" by (blast elim: less_asym) text \Transitivity rules for calculational reasoning\ lemma less_asym': "a < b \ b < a \ P" by (rule less_asym) text \Dual order\ lemma dual_preorder: - "class.preorder (\) (>)" + \class.preorder (\) (>)\ by standard (auto simp add: less_le_not_le intro: order_trans) end subsection \Partial orders\ class order = preorder + assumes antisym: "x \ y \ y \ x \ x = y" begin lemma less_le: "x < y \ x \ y \ x \ y" by (auto simp add: less_le_not_le intro: antisym) sublocale order: ordering less_eq less + dual_order: ordering greater_eq greater proof - interpret ordering less_eq less by standard (auto intro: antisym order_trans simp add: less_le) show "ordering less_eq less" by (fact ordering_axioms) then show "ordering greater_eq greater" by (rule ordering_dualI) qed text \Reflexivity.\ lemma le_less: "x \ y \ x < y \ x = y" \ \NOT suitable for iff, since it can cause PROOF FAILED.\ by (fact order.order_iff_strict) lemma le_imp_less_or_eq: "x \ y \ x < y \ x = y" by (simp add: less_le) text \Useful for simplification, but too risky to include by default.\ lemma less_imp_not_eq: "x < y \ (x = y) \ False" by auto lemma less_imp_not_eq2: "x < y \ (y = x) \ False" by auto text \Transitivity rules for calculational reasoning\ lemma neq_le_trans: "a \ b \ a \ b \ a < b" by (fact order.not_eq_order_implies_strict) lemma le_neq_trans: "a \ b \ a \ b \ a < b" by (rule order.not_eq_order_implies_strict) text \Asymmetry.\ lemma eq_iff: "x = y \ x \ y \ y \ x" by (fact order.eq_iff) lemma antisym_conv: "y \ x \ x \ y \ x = y" by (simp add: eq_iff) lemma less_imp_neq: "x < y \ x \ y" by (fact order.strict_implies_not_eq) lemma antisym_conv1: "\ x < y \ x \ y \ x = y" by (simp add: local.le_less) lemma antisym_conv2: "x \ y \ \ x < y \ x = y" by (simp add: local.less_le) lemma leD: "y \ x \ \ x < y" by (auto simp: less_le antisym) text \Least value operator\ definition (in ord) Least :: "('a \ bool) \ 'a" (binder "LEAST " 10) where "Least P = (THE x. P x \ (\y. P y \ x \ y))" lemma Least_equality: assumes "P x" and "\y. P y \ x \ y" shows "Least P = x" unfolding Least_def by (rule the_equality) (blast intro: assms antisym)+ lemma LeastI2_order: assumes "P x" and "\y. P y \ x \ y" and "\x. P x \ \y. P y \ x \ y \ Q x" shows "Q (Least P)" unfolding Least_def by (rule theI2) (blast intro: assms antisym)+ lemma Least_ex1: assumes "\!x. P x \ (\y. P y \ x \ y)" shows Least1I: "P (Least P)" and Least1_le: "P z \ Least P \ z" using theI'[OF assms] unfolding Least_def by auto text \Greatest value operator\ definition Greatest :: "('a \ bool) \ 'a" (binder "GREATEST " 10) where "Greatest P = (THE x. P x \ (\y. P y \ x \ y))" lemma GreatestI2_order: "\ P x; \y. P y \ x \ y; \x. \ P x; \y. P y \ x \ y \ \ Q x \ \ Q (Greatest P)" unfolding Greatest_def by (rule theI2) (blast intro: antisym)+ lemma Greatest_equality: "\ P x; \y. P y \ x \ y \ \ Greatest P = x" unfolding Greatest_def by (rule the_equality) (blast intro: antisym)+ end lemma ordering_orderI: fixes less_eq (infix "\<^bold>\" 50) and less (infix "\<^bold><" 50) assumes "ordering less_eq less" shows "class.order less_eq less" proof - from assms interpret ordering less_eq less . show ?thesis by standard (auto intro: antisym trans simp add: refl strict_iff_order) qed lemma order_strictI: fixes less (infix "\" 50) and less_eq (infix "\" 50) assumes "\a b. a \ b \ a \ b \ a = b" assumes "\a b. a \ b \ \ b \ a" assumes "\a. \ a \ a" assumes "\a b c. a \ b \ b \ c \ a \ c" shows "class.order less_eq less" by (rule ordering_orderI) (rule ordering_strictI, (fact assms)+) context order begin text \Dual order\ lemma dual_order: "class.order (\) (>)" using dual_order.ordering_axioms by (rule ordering_orderI) end subsection \Linear (total) orders\ class linorder = order + assumes linear: "x \ y \ y \ x" begin lemma less_linear: "x < y \ x = y \ y < x" unfolding less_le using less_le linear by blast lemma le_less_linear: "x \ y \ y < x" by (simp add: le_less less_linear) lemma le_cases [case_names le ge]: "(x \ y \ P) \ (y \ x \ P) \ P" using linear by blast lemma (in linorder) le_cases3: "\\x \ y; y \ z\ \ P; \y \ x; x \ z\ \ P; \x \ z; z \ y\ \ P; \z \ y; y \ x\ \ P; \y \ z; z \ x\ \ P; \z \ x; x \ y\ \ P\ \ P" by (blast intro: le_cases) lemma linorder_cases [case_names less equal greater]: "(x < y \ P) \ (x = y \ P) \ (y < x \ P) \ P" using less_linear by blast lemma linorder_wlog[case_names le sym]: "(\a b. a \ b \ P a b) \ (\a b. P b a \ P a b) \ P a b" by (cases rule: le_cases[of a b]) blast+ lemma not_less: "\ x < y \ y \ x" unfolding less_le using linear by (blast intro: antisym) lemma not_less_iff_gr_or_eq: "\(x < y) \ (x > y \ x = y)" by (auto simp add:not_less le_less) lemma not_le: "\ x \ y \ y < x" unfolding less_le using linear by (blast intro: antisym) lemma neq_iff: "x \ y \ x < y \ y < x" by (cut_tac x = x and y = y in less_linear, auto) lemma neqE: "x \ y \ (x < y \ R) \ (y < x \ R) \ R" by (simp add: neq_iff) blast lemma antisym_conv3: "\ y < x \ \ x < y \ x = y" by (blast intro: antisym dest: not_less [THEN iffD1]) lemma leI: "\ x < y \ y \ x" unfolding not_less . lemma not_le_imp_less: "\ y \ x \ x < y" unfolding not_le . lemma linorder_less_wlog[case_names less refl sym]: "\\a b. a < b \ P a b; \a. P a a; \a b. P b a \ P a b\ \ P a b" using antisym_conv3 by blast text \Dual order\ lemma dual_linorder: "class.linorder (\) (>)" by (rule class.linorder.intro, rule dual_order) (unfold_locales, rule linear) end text \Alternative introduction rule with bias towards strict order\ lemma linorder_strictI: fixes less_eq (infix "\<^bold>\" 50) and less (infix "\<^bold><" 50) assumes "class.order less_eq less" assumes trichotomy: "\a b. a \<^bold>< b \ a = b \ b \<^bold>< a" shows "class.linorder less_eq less" proof - interpret order less_eq less by (fact \class.order less_eq less\) show ?thesis proof fix a b show "a \<^bold>\ b \ b \<^bold>\ a" using trichotomy by (auto simp add: le_less) qed qed subsection \Reasoning tools setup\ ML \ signature ORDERS = sig val print_structures: Proof.context -> unit val order_tac: Proof.context -> thm list -> int -> tactic val add_struct: string * term list -> string -> attribute val del_struct: string * term list -> attribute end; structure Orders: ORDERS = struct (* context data *) fun struct_eq ((s1: string, ts1), (s2, ts2)) = s1 = s2 andalso eq_list (op aconv) (ts1, ts2); structure Data = Generic_Data ( type T = ((string * term list) * Order_Tac.less_arith) list; (* Order structures: identifier of the structure, list of operations and record of theorems needed to set up the transitivity reasoner, identifier and operations identify the structure uniquely. *) val empty = []; val extend = I; fun merge data = AList.join struct_eq (K fst) data; ); fun print_structures ctxt = let val structs = Data.get (Context.Proof ctxt); fun pretty_term t = Pretty.block [Pretty.quote (Syntax.pretty_term ctxt t), Pretty.brk 1, Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt (type_of t))]; fun pretty_struct ((s, ts), _) = Pretty.block [Pretty.str s, Pretty.str ":", Pretty.brk 1, Pretty.enclose "(" ")" (Pretty.breaks (map pretty_term ts))]; in Pretty.writeln (Pretty.big_list "order structures:" (map pretty_struct structs)) end; val _ = Outer_Syntax.command \<^command_keyword>\print_orders\ "print order structures available to transitivity reasoner" (Scan.succeed (Toplevel.keep (print_structures o Toplevel.context_of))); (* tactics *) fun struct_tac ((s, ops), thms) ctxt facts = let val [eq, le, less] = ops; fun decomp thy (\<^const>\Trueprop\ $ t) = let fun excluded t = (* exclude numeric types: linear arithmetic subsumes transitivity *) let val T = type_of t in T = HOLogic.natT orelse T = HOLogic.intT orelse T = HOLogic.realT end; fun rel (bin_op $ t1 $ t2) = if excluded t1 then NONE else if Pattern.matches thy (eq, bin_op) then SOME (t1, "=", t2) else if Pattern.matches thy (le, bin_op) then SOME (t1, "<=", t2) else if Pattern.matches thy (less, bin_op) then SOME (t1, "<", t2) else NONE | rel _ = NONE; fun dec (Const (\<^const_name>\Not\, _) $ t) = (case rel t of NONE => NONE | SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2)) | dec x = rel x; in dec t end | decomp _ _ = NONE; in (case s of "order" => Order_Tac.partial_tac decomp thms ctxt facts | "linorder" => Order_Tac.linear_tac decomp thms ctxt facts | _ => error ("Unknown order kind " ^ quote s ^ " encountered in transitivity reasoner")) end fun order_tac ctxt facts = FIRST' (map (fn s => CHANGED o struct_tac s ctxt facts) (Data.get (Context.Proof ctxt))); (* attributes *) fun add_struct s tag = Thm.declaration_attribute (fn thm => Data.map (AList.map_default struct_eq (s, Order_Tac.empty TrueI) (Order_Tac.update tag thm))); fun del_struct s = Thm.declaration_attribute (fn _ => Data.map (AList.delete struct_eq s)); end; \ attribute_setup order = \ Scan.lift ((Args.add -- Args.name >> (fn (_, s) => SOME s) || Args.del >> K NONE) --| Args.colon (* FIXME || Scan.succeed true *) ) -- Scan.lift Args.name -- Scan.repeat Args.term >> (fn ((SOME tag, n), ts) => Orders.add_struct (n, ts) tag | ((NONE, n), ts) => Orders.del_struct (n, ts)) \ "theorems controlling transitivity reasoner" method_setup order = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (Orders.order_tac ctxt [])) \ "transitivity reasoner" text \Declarations to set up transitivity reasoner of partial and linear orders.\ context order begin (* The type constraint on @{term (=}) below is necessary since the operation is not a parameter of the locale. *) declare less_irrefl [THEN notE, order add less_reflE: order "(=) :: 'a \ 'a \ bool" "(<=)" "(<)"] declare order_refl [order add le_refl: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_imp_le [order add less_imp_le: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare antisym [order add eqI: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare eq_refl [order add eqD1: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare sym [THEN eq_refl, order add eqD2: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_trans [order add less_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_le_trans [order add less_le_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare le_less_trans [order add le_less_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare order_trans [order add le_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare le_neq_trans [order add le_neq_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare neq_le_trans [order add neq_le_trans: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_imp_neq [order add less_imp_neq: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare eq_neq_eq_imp_neq [order add eq_neq_eq_imp_neq: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_sym [order add not_sym: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"] end context linorder begin declare [[order del: order "(=) :: 'a => 'a => bool" "(<=)" "(<)"]] declare less_irrefl [THEN notE, order add less_reflE: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare order_refl [order add le_refl: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_imp_le [order add less_imp_le: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_less [THEN iffD2, order add not_lessI: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_le [THEN iffD2, order add not_leI: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_less [THEN iffD1, order add not_lessD: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_le [THEN iffD1, order add not_leD: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare antisym [order add eqI: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare eq_refl [order add eqD1: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare sym [THEN eq_refl, order add eqD2: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_trans [order add less_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_le_trans [order add less_le_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare le_less_trans [order add le_less_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare order_trans [order add le_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare le_neq_trans [order add le_neq_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare neq_le_trans [order add neq_le_trans: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare less_imp_neq [order add less_imp_neq: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare eq_neq_eq_imp_neq [order add eq_neq_eq_imp_neq: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] declare not_sym [order add not_sym: linorder "(=) :: 'a => 'a => bool" "(<=)" "(<)"] end setup \ map_theory_simpset (fn ctxt0 => ctxt0 addSolver mk_solver "Transitivity" (fn ctxt => Orders.order_tac ctxt (Simplifier.prems_of ctxt))) (*Adding the transitivity reasoners also as safe solvers showed a slight speed up, but the reasoning strength appears to be not higher (at least no breaking of additional proofs in the entire HOL distribution, as of 5 March 2004, was observed).*) \ ML \ local fun prp t thm = Thm.prop_of thm = t; (* FIXME proper aconv!? *) in fun antisym_le_simproc ctxt ct = (case Thm.term_of ct of (le as Const (_, T)) $ r $ s => (let val prems = Simplifier.prems_of ctxt; val less = Const (\<^const_name>\less\, T); val t = HOLogic.mk_Trueprop(le $ s $ r); in (case find_first (prp t) prems of NONE => let val t = HOLogic.mk_Trueprop(HOLogic.Not $ (less $ r $ s)) in (case find_first (prp t) prems of NONE => NONE | SOME thm => SOME(mk_meta_eq(thm RS @{thm antisym_conv1}))) end | SOME thm => SOME (mk_meta_eq (thm RS @{thm order_class.antisym_conv}))) end handle THM _ => NONE) | _ => NONE); fun antisym_less_simproc ctxt ct = (case Thm.term_of ct of NotC $ ((less as Const(_,T)) $ r $ s) => (let val prems = Simplifier.prems_of ctxt; val le = Const (\<^const_name>\less_eq\, T); val t = HOLogic.mk_Trueprop(le $ r $ s); in (case find_first (prp t) prems of NONE => let val t = HOLogic.mk_Trueprop (NotC $ (less $ s $ r)) in (case find_first (prp t) prems of NONE => NONE | SOME thm => SOME (mk_meta_eq(thm RS @{thm linorder_class.antisym_conv3}))) end | SOME thm => SOME (mk_meta_eq (thm RS @{thm antisym_conv2}))) end handle THM _ => NONE) | _ => NONE); end; \ simproc_setup antisym_le ("(x::'a::order) \ y") = "K antisym_le_simproc" simproc_setup antisym_less ("\ (x::'a::linorder) < y") = "K antisym_less_simproc" subsection \Bounded quantifiers\ syntax (ASCII) "_All_less" :: "[idt, 'a, bool] => bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3EX _<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) "_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) "_All_greater" :: "[idt, 'a, bool] => bool" ("(3ALL _>_./ _)" [0, 0, 10] 10) "_Ex_greater" :: "[idt, 'a, bool] => bool" ("(3EX _>_./ _)" [0, 0, 10] 10) "_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3ALL _>=_./ _)" [0, 0, 10] 10) "_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3EX _>=_./ _)" [0, 0, 10] 10) "_All_neq" :: "[idt, 'a, bool] => bool" ("(3ALL _~=_./ _)" [0, 0, 10] 10) "_Ex_neq" :: "[idt, 'a, bool] => bool" ("(3EX _~=_./ _)" [0, 0, 10] 10) syntax "_All_less" :: "[idt, 'a, bool] => bool" ("(3\_<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3\_<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_All_greater" :: "[idt, 'a, bool] => bool" ("(3\_>_./ _)" [0, 0, 10] 10) "_Ex_greater" :: "[idt, 'a, bool] => bool" ("(3\_>_./ _)" [0, 0, 10] 10) "_All_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_greater_eq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_All_neq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Ex_neq" :: "[idt, 'a, bool] => bool" ("(3\_\_./ _)" [0, 0, 10] 10) syntax (input) "_All_less" :: "[idt, 'a, bool] => bool" ("(3! _<_./ _)" [0, 0, 10] 10) "_Ex_less" :: "[idt, 'a, bool] => bool" ("(3? _<_./ _)" [0, 0, 10] 10) "_All_less_eq" :: "[idt, 'a, bool] => bool" ("(3! _<=_./ _)" [0, 0, 10] 10) "_Ex_less_eq" :: "[idt, 'a, bool] => bool" ("(3? _<=_./ _)" [0, 0, 10] 10) "_All_neq" :: "[idt, 'a, bool] => bool" ("(3! _~=_./ _)" [0, 0, 10] 10) "_Ex_neq" :: "[idt, 'a, bool] => bool" ("(3? _~=_./ _)" [0, 0, 10] 10) translations "\x "\x. x < y \ P" "\x "\x. x < y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x>y. P" \ "\x. x > y \ P" "\x>y. P" \ "\x. x > y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" "\x\y. P" \ "\x. x \ y \ P" print_translation \ let val All_binder = Mixfix.binder_name \<^const_syntax>\All\; val Ex_binder = Mixfix.binder_name \<^const_syntax>\Ex\; val impl = \<^const_syntax>\HOL.implies\; val conj = \<^const_syntax>\HOL.conj\; val less = \<^const_syntax>\less\; val less_eq = \<^const_syntax>\less_eq\; val trans = [((All_binder, impl, less), (\<^syntax_const>\_All_less\, \<^syntax_const>\_All_greater\)), ((All_binder, impl, less_eq), (\<^syntax_const>\_All_less_eq\, \<^syntax_const>\_All_greater_eq\)), ((Ex_binder, conj, less), (\<^syntax_const>\_Ex_less\, \<^syntax_const>\_Ex_greater\)), ((Ex_binder, conj, less_eq), (\<^syntax_const>\_Ex_less_eq\, \<^syntax_const>\_Ex_greater_eq\))]; fun matches_bound v t = (case t of Const (\<^syntax_const>\_bound\, _) $ Free (v', _) => v = v' | _ => false); fun contains_var v = Term.exists_subterm (fn Free (x, _) => x = v | _ => false); fun mk x c n P = Syntax.const c $ Syntax_Trans.mark_bound_body x $ n $ P; fun tr' q = (q, fn _ => (fn [Const (\<^syntax_const>\_bound\, _) $ Free (v, T), Const (c, _) $ (Const (d, _) $ t $ u) $ P] => (case AList.lookup (=) trans (q, c, d) of NONE => raise Match | SOME (l, g) => if matches_bound v t andalso not (contains_var v u) then mk (v, T) l u P else if matches_bound v u andalso not (contains_var v t) then mk (v, T) g t P else raise Match) | _ => raise Match)); in [tr' All_binder, tr' Ex_binder] end \ subsection \Transitivity reasoning\ context ord begin lemma ord_le_eq_trans: "a \ b \ b = c \ a \ c" by (rule subst) lemma ord_eq_le_trans: "a = b \ b \ c \ a \ c" by (rule ssubst) lemma ord_less_eq_trans: "a < b \ b = c \ a < c" by (rule subst) lemma ord_eq_less_trans: "a = b \ b < c \ a < c" by (rule ssubst) end lemma order_less_subst2: "(a::'a::order) < b ==> f b < (c::'c::order) ==> (!!x y. x < y ==> f x < f y) ==> f a < c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a < b" hence "f a < f b" by (rule r) also assume "f b < c" finally (less_trans) show ?thesis . qed lemma order_less_subst1: "(a::'a::order) < f b ==> (b::'b::order) < c ==> (!!x y. x < y ==> f x < f y) ==> a < f c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a < f b" also assume "b < c" hence "f b < f c" by (rule r) finally (less_trans) show ?thesis . qed lemma order_le_less_subst2: "(a::'a::order) <= b ==> f b < (c::'c::order) ==> (!!x y. x <= y ==> f x <= f y) ==> f a < c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a <= b" hence "f a <= f b" by (rule r) also assume "f b < c" finally (le_less_trans) show ?thesis . qed lemma order_le_less_subst1: "(a::'a::order) <= f b ==> (b::'b::order) < c ==> (!!x y. x < y ==> f x < f y) ==> a < f c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a <= f b" also assume "b < c" hence "f b < f c" by (rule r) finally (le_less_trans) show ?thesis . qed lemma order_less_le_subst2: "(a::'a::order) < b ==> f b <= (c::'c::order) ==> (!!x y. x < y ==> f x < f y) ==> f a < c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a < b" hence "f a < f b" by (rule r) also assume "f b <= c" finally (less_le_trans) show ?thesis . qed lemma order_less_le_subst1: "(a::'a::order) < f b ==> (b::'b::order) <= c ==> (!!x y. x <= y ==> f x <= f y) ==> a < f c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a < f b" also assume "b <= c" hence "f b <= f c" by (rule r) finally (less_le_trans) show ?thesis . qed lemma order_subst1: "(a::'a::order) <= f b ==> (b::'b::order) <= c ==> (!!x y. x <= y ==> f x <= f y) ==> a <= f c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a <= f b" also assume "b <= c" hence "f b <= f c" by (rule r) finally (order_trans) show ?thesis . qed lemma order_subst2: "(a::'a::order) <= b ==> f b <= (c::'c::order) ==> (!!x y. x <= y ==> f x <= f y) ==> f a <= c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a <= b" hence "f a <= f b" by (rule r) also assume "f b <= c" finally (order_trans) show ?thesis . qed lemma ord_le_eq_subst: "a <= b ==> f b = c ==> (!!x y. x <= y ==> f x <= f y) ==> f a <= c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a <= b" hence "f a <= f b" by (rule r) also assume "f b = c" finally (ord_le_eq_trans) show ?thesis . qed lemma ord_eq_le_subst: "a = f b ==> b <= c ==> (!!x y. x <= y ==> f x <= f y) ==> a <= f c" proof - assume r: "!!x y. x <= y ==> f x <= f y" assume "a = f b" also assume "b <= c" hence "f b <= f c" by (rule r) finally (ord_eq_le_trans) show ?thesis . qed lemma ord_less_eq_subst: "a < b ==> f b = c ==> (!!x y. x < y ==> f x < f y) ==> f a < c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a < b" hence "f a < f b" by (rule r) also assume "f b = c" finally (ord_less_eq_trans) show ?thesis . qed lemma ord_eq_less_subst: "a = f b ==> b < c ==> (!!x y. x < y ==> f x < f y) ==> a < f c" proof - assume r: "!!x y. x < y ==> f x < f y" assume "a = f b" also assume "b < c" hence "f b < f c" by (rule r) finally (ord_eq_less_trans) show ?thesis . qed text \ Note that this list of rules is in reverse order of priorities. \ lemmas [trans] = order_less_subst2 order_less_subst1 order_le_less_subst2 order_le_less_subst1 order_less_le_subst2 order_less_le_subst1 order_subst2 order_subst1 ord_le_eq_subst ord_eq_le_subst ord_less_eq_subst ord_eq_less_subst forw_subst back_subst rev_mp mp lemmas (in order) [trans] = neq_le_trans le_neq_trans lemmas (in preorder) [trans] = less_trans less_asym' le_less_trans less_le_trans order_trans lemmas (in order) [trans] = antisym lemmas (in ord) [trans] = ord_le_eq_trans ord_eq_le_trans ord_less_eq_trans ord_eq_less_trans lemmas [trans] = trans lemmas order_trans_rules = order_less_subst2 order_less_subst1 order_le_less_subst2 order_le_less_subst1 order_less_le_subst2 order_less_le_subst1 order_subst2 order_subst1 ord_le_eq_subst ord_eq_le_subst ord_less_eq_subst ord_eq_less_subst forw_subst back_subst rev_mp mp neq_le_trans le_neq_trans less_trans less_asym' le_less_trans less_le_trans order_trans antisym ord_le_eq_trans ord_eq_le_trans ord_less_eq_trans ord_eq_less_trans trans text \These support proving chains of decreasing inequalities a >= b >= c ... in Isar proofs.\ lemma xt1 [no_atp]: "a = b \ b > c \ a > c" "a > b \ b = c \ a > c" "a = b \ b \ c \ a \ c" "a \ b \ b = c \ a \ c" "(x::'a::order) \ y \ y \ x \ x = y" "(x::'a::order) \ y \ y \ z \ x \ z" "(x::'a::order) > y \ y \ z \ x > z" "(x::'a::order) \ y \ y > z \ x > z" "(a::'a::order) > b \ b > a \ P" "(x::'a::order) > y \ y > z \ x > z" "(a::'a::order) \ b \ a \ b \ a > b" "(a::'a::order) \ b \ a \ b \ a > b" "a = f b \ b > c \ (\x y. x > y \ f x > f y) \ a > f c" "a > b \ f b = c \ (\x y. x > y \ f x > f y) \ f a > c" "a = f b \ b \ c \ (\x y. x \ y \ f x \ f y) \ a \ f c" "a \ b \ f b = c \ (\x y. x \ y \ f x \ f y) \ f a \ c" by auto lemma xt2 [no_atp]: "(a::'a::order) >= f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a >= f c" by (subgoal_tac "f b >= f c", force, force) lemma xt3 [no_atp]: "(a::'a::order) >= b ==> (f b::'b::order) >= c ==> (!!x y. x >= y ==> f x >= f y) ==> f a >= c" by (subgoal_tac "f a >= f b", force, force) lemma xt4 [no_atp]: "(a::'a::order) > f b ==> (b::'b::order) >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a > f c" by (subgoal_tac "f b >= f c", force, force) lemma xt5 [no_atp]: "(a::'a::order) > b ==> (f b::'b::order) >= c==> (!!x y. x > y ==> f x > f y) ==> f a > c" by (subgoal_tac "f a > f b", force, force) lemma xt6 [no_atp]: "(a::'a::order) >= f b ==> b > c ==> (!!x y. x > y ==> f x > f y) ==> a > f c" by (subgoal_tac "f b > f c", force, force) lemma xt7 [no_atp]: "(a::'a::order) >= b ==> (f b::'b::order) > c ==> (!!x y. x >= y ==> f x >= f y) ==> f a > c" by (subgoal_tac "f a >= f b", force, force) lemma xt8 [no_atp]: "(a::'a::order) > f b ==> (b::'b::order) > c ==> (!!x y. x > y ==> f x > f y) ==> a > f c" by (subgoal_tac "f b > f c", force, force) lemma xt9 [no_atp]: "(a::'a::order) > b ==> (f b::'b::order) > c ==> (!!x y. x > y ==> f x > f y) ==> f a > c" by (subgoal_tac "f a > f b", force, force) lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9 (* Since "a >= b" abbreviates "b <= a", the abbreviation "..." stands for the wrong thing in an Isar proof. The extra transitivity rules can be used as follows: lemma "(a::'a::order) > z" proof - have "a >= b" (is "_ >= ?rhs") sorry also have "?rhs >= c" (is "_ >= ?rhs") sorry also (xtrans) have "?rhs = d" (is "_ = ?rhs") sorry also (xtrans) have "?rhs >= e" (is "_ >= ?rhs") sorry also (xtrans) have "?rhs > f" (is "_ > ?rhs") sorry also (xtrans) have "?rhs > z" sorry finally (xtrans) show ?thesis . qed Alternatively, one can use "declare xtrans [trans]" and then leave out the "(xtrans)" above. *) subsection \Monotonicity\ context order begin definition mono :: "('a \ 'b::order) \ bool" where "mono f \ (\x y. x \ y \ f x \ f y)" lemma monoI [intro?]: fixes f :: "'a \ 'b::order" shows "(\x y. x \ y \ f x \ f y) \ mono f" unfolding mono_def by iprover lemma monoD [dest?]: fixes f :: "'a \ 'b::order" shows "mono f \ x \ y \ f x \ f y" unfolding mono_def by iprover lemma monoE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: mono_def) qed definition antimono :: "('a \ 'b::order) \ bool" where "antimono f \ (\x y. x \ y \ f x \ f y)" lemma antimonoI [intro?]: fixes f :: "'a \ 'b::order" shows "(\x y. x \ y \ f x \ f y) \ antimono f" unfolding antimono_def by iprover lemma antimonoD [dest?]: fixes f :: "'a \ 'b::order" shows "antimono f \ x \ y \ f x \ f y" unfolding antimono_def by iprover lemma antimonoE: fixes f :: "'a \ 'b::order" assumes "antimono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: antimono_def) qed definition strict_mono :: "('a \ 'b::order) \ bool" where "strict_mono f \ (\x y. x < y \ f x < f y)" lemma strict_monoI [intro?]: assumes "\x y. x < y \ f x < f y" shows "strict_mono f" using assms unfolding strict_mono_def by auto lemma strict_monoD [dest?]: "strict_mono f \ x < y \ f x < f y" unfolding strict_mono_def by auto lemma strict_mono_mono [dest?]: assumes "strict_mono f" shows "mono f" proof (rule monoI) fix x y assume "x \ y" show "f x \ f y" proof (cases "x = y") case True then show ?thesis by simp next case False with \x \ y\ have "x < y" by simp with assms strict_monoD have "f x < f y" by auto then show ?thesis by simp qed qed end context linorder begin lemma mono_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x \ y" proof show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma mono_strict_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x < y" proof show "x < y" proof (rule ccontr) assume "\ x < y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma strict_mono_eq: assumes "strict_mono f" shows "f x = f y \ x = y" proof assume "f x = f y" show "x = y" proof (cases x y rule: linorder_cases) case less with assms strict_monoD have "f x < f y" by auto with \f x = f y\ show ?thesis by simp next case equal then show ?thesis . next case greater with assms strict_monoD have "f y < f x" by auto with \f x = f y\ show ?thesis by simp qed qed simp lemma strict_mono_less_eq: assumes "strict_mono f" shows "f x \ f y \ x \ y" proof assume "x \ y" with assms strict_mono_mono monoD show "f x \ f y" by auto next assume "f x \ f y" show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y < x" by simp with assms strict_monoD have "f y < f x" by auto with \f x \ f y\ show False by simp qed qed lemma strict_mono_less: assumes "strict_mono f" shows "f x < f y \ x < y" using assms by (auto simp add: less_le Orderings.less_le strict_mono_eq strict_mono_less_eq) end subsection \min and max -- fundamental\ definition (in ord) min :: "'a \ 'a \ 'a" where "min a b = (if a \ b then a else b)" definition (in ord) max :: "'a \ 'a \ 'a" where "max a b = (if a \ b then b else a)" lemma min_absorb1: "x \ y \ min x y = x" by (simp add: min_def) lemma max_absorb2: "x \ y \ max x y = y" by (simp add: max_def) lemma min_absorb2: "(y::'a::order) \ x \ min x y = y" by (simp add:min_def) lemma max_absorb1: "(y::'a::order) \ x \ max x y = x" by (simp add: max_def) lemma max_min_same [simp]: fixes x y :: "'a :: linorder" shows "max x (min x y) = x" "max (min x y) x = x" "max (min x y) y = y" "max y (min x y) = y" by(auto simp add: max_def min_def) subsection \(Unique) top and bottom elements\ class bot = fixes bot :: 'a ("\") class order_bot = order + bot + assumes bot_least: "\ \ a" begin sublocale bot: ordering_top greater_eq greater bot by standard (fact bot_least) lemma le_bot: "a \ \ \ a = \" by (fact bot.extremum_uniqueI) lemma bot_unique: "a \ \ \ a = \" by (fact bot.extremum_unique) lemma not_less_bot: "\ a < \" by (fact bot.extremum_strict) lemma bot_less: "a \ \ \ \ < a" by (fact bot.not_eq_extremum) lemma max_bot[simp]: "max bot x = x" by(simp add: max_def bot_unique) lemma max_bot2[simp]: "max x bot = x" by(simp add: max_def bot_unique) lemma min_bot[simp]: "min bot x = bot" by(simp add: min_def bot_unique) lemma min_bot2[simp]: "min x bot = bot" by(simp add: min_def bot_unique) end class top = fixes top :: 'a ("\") class order_top = order + top + assumes top_greatest: "a \ \" begin sublocale top: ordering_top less_eq less top by standard (fact top_greatest) lemma top_le: "\ \ a \ a = \" by (fact top.extremum_uniqueI) lemma top_unique: "\ \ a \ a = \" by (fact top.extremum_unique) lemma not_top_less: "\ \ < a" by (fact top.extremum_strict) lemma less_top: "a \ \ \ a < \" by (fact top.not_eq_extremum) lemma max_top[simp]: "max top x = top" by(simp add: max_def top_unique) lemma max_top2[simp]: "max x top = top" by(simp add: max_def top_unique) lemma min_top[simp]: "min top x = x" by(simp add: min_def top_unique) lemma min_top2[simp]: "min x top = x" by(simp add: min_def top_unique) end subsection \Dense orders\ class dense_order = order + assumes dense: "x < y \ (\z. x < z \ z < y)" class dense_linorder = linorder + dense_order begin lemma dense_le: fixes y z :: 'a assumes "\x. x < y \ x \ z" shows "y \ z" proof (rule ccontr) assume "\ ?thesis" hence "z < y" by simp from dense[OF this] obtain x where "x < y" and "z < x" by safe moreover have "x \ z" using assms[OF \x < y\] . ultimately show False by auto qed lemma dense_le_bounded: fixes x y z :: 'a assumes "x < y" assumes *: "\w. \ x < w ; w < y \ \ w \ z" shows "y \ z" proof (rule dense_le) fix w assume "w < y" from dense[OF \x < y\] obtain u where "x < u" "u < y" by safe from linear[of u w] show "w \ z" proof (rule disjE) assume "u \ w" from less_le_trans[OF \x < u\ \u \ w\] \w < y\ show "w \ z" by (rule *) next assume "w \ u" from \w \ u\ *[OF \x < u\ \u < y\] show "w \ z" by (rule order_trans) qed qed lemma dense_ge: fixes y z :: 'a assumes "\x. z < x \ y \ x" shows "y \ z" proof (rule ccontr) assume "\ ?thesis" hence "z < y" by simp from dense[OF this] obtain x where "x < y" and "z < x" by safe moreover have "y \ x" using assms[OF \z < x\] . ultimately show False by auto qed lemma dense_ge_bounded: fixes x y z :: 'a assumes "z < x" assumes *: "\w. \ z < w ; w < x \ \ y \ w" shows "y \ z" proof (rule dense_ge) fix w assume "z < w" from dense[OF \z < x\] obtain u where "z < u" "u < x" by safe from linear[of u w] show "y \ w" proof (rule disjE) assume "w \ u" from \z < w\ le_less_trans[OF \w \ u\ \u < x\] show "y \ w" by (rule *) next assume "u \ w" from *[OF \z < u\ \u < x\] \u \ w\ show "y \ w" by (rule order_trans) qed qed end class no_top = order + assumes gt_ex: "\y. x < y" class no_bot = order + assumes lt_ex: "\y. y < x" class unbounded_dense_linorder = dense_linorder + no_top + no_bot subsection \Wellorders\ class wellorder = linorder + assumes less_induct [case_names less]: "(\x. (\y. y < x \ P y) \ P x) \ P a" begin lemma wellorder_Least_lemma: fixes k :: 'a assumes "P k" shows LeastI: "P (LEAST x. P x)" and Least_le: "(LEAST x. P x) \ k" proof - have "P (LEAST x. P x) \ (LEAST x. P x) \ k" using assms proof (induct k rule: less_induct) case (less x) then have "P x" by simp show ?case proof (rule classical) assume assm: "\ (P (LEAST a. P a) \ (LEAST a. P a) \ x)" have "\y. P y \ x \ y" proof (rule classical) fix y assume "P y" and "\ x \ y" with less have "P (LEAST a. P a)" and "(LEAST a. P a) \ y" by (auto simp add: not_le) with assm have "x < (LEAST a. P a)" and "(LEAST a. P a) \ y" by auto then show "x \ y" by auto qed with \P x\ have Least: "(LEAST a. P a) = x" by (rule Least_equality) with \P x\ show ?thesis by simp qed qed then show "P (LEAST x. P x)" and "(LEAST x. P x) \ k" by auto qed \ \The following 3 lemmas are due to Brian Huffman\ lemma LeastI_ex: "\x. P x \ P (Least P)" by (erule exE) (erule LeastI) lemma LeastI2: "P a \ (\x. P x \ Q x) \ Q (Least P)" by (blast intro: LeastI) lemma LeastI2_ex: "\a. P a \ (\x. P x \ Q x) \ Q (Least P)" by (blast intro: LeastI_ex) lemma LeastI2_wellorder: assumes "P a" and "\a. \ P a; \b. P b \ a \ b \ \ Q a" shows "Q (Least P)" proof (rule LeastI2_order) show "P (Least P)" using \P a\ by (rule LeastI) next fix y assume "P y" thus "Least P \ y" by (rule Least_le) next fix x assume "P x" "\y. P y \ x \ y" thus "Q x" by (rule assms(2)) qed lemma LeastI2_wellorder_ex: assumes "\x. P x" and "\a. \ P a; \b. P b \ a \ b \ \ Q a" shows "Q (Least P)" using assms by clarify (blast intro!: LeastI2_wellorder) lemma not_less_Least: "k < (LEAST x. P x) \ \ P k" apply (simp add: not_le [symmetric]) apply (erule contrapos_nn) apply (erule Least_le) done lemma exists_least_iff: "(\n. P n) \ (\n. P n \ (\m < n. \ P m))" (is "?lhs \ ?rhs") proof assume ?rhs thus ?lhs by blast next assume H: ?lhs then obtain n where n: "P n" by blast let ?x = "Least P" { fix m assume m: "m < ?x" from not_less_Least[OF m] have "\ P m" . } with LeastI_ex[OF H] show ?rhs by blast qed end subsection \Order on \<^typ>\bool\\ instantiation bool :: "{order_bot, order_top, linorder}" begin definition le_bool_def [simp]: "P \ Q \ P \ Q" definition [simp]: "(P::bool) < Q \ \ P \ Q" definition [simp]: "\ \ False" definition [simp]: "\ \ True" instance proof qed auto end lemma le_boolI: "(P \ Q) \ P \ Q" by simp lemma le_boolI': "P \ Q \ P \ Q" by simp lemma le_boolE: "P \ Q \ P \ (Q \ R) \ R" by simp lemma le_boolD: "P \ Q \ P \ Q" by simp lemma bot_boolE: "\ \ P" by simp lemma top_boolI: \ by simp lemma [code]: "False \ b \ True" "True \ b \ b" "False < b \ b" "True < b \ False" by simp_all subsection \Order on \<^typ>\_ \ _\\ instantiation "fun" :: (type, ord) ord begin definition le_fun_def: "f \ g \ (\x. f x \ g x)" definition "(f::'a \ 'b) < g \ f \ g \ \ (g \ f)" instance .. end instance "fun" :: (type, preorder) preorder proof qed (auto simp add: le_fun_def less_fun_def intro: order_trans antisym) instance "fun" :: (type, order) order proof qed (auto simp add: le_fun_def intro: antisym) instantiation "fun" :: (type, bot) bot begin definition "\ = (\x. \)" instance .. end instantiation "fun" :: (type, order_bot) order_bot begin lemma bot_apply [simp, code]: "\ x = \" by (simp add: bot_fun_def) instance proof qed (simp add: le_fun_def) end instantiation "fun" :: (type, top) top begin definition [no_atp]: "\ = (\x. \)" instance .. end instantiation "fun" :: (type, order_top) order_top begin lemma top_apply [simp, code]: "\ x = \" by (simp add: top_fun_def) instance proof qed (simp add: le_fun_def) end lemma le_funI: "(\x. f x \ g x) \ f \ g" unfolding le_fun_def by simp lemma le_funE: "f \ g \ (f x \ g x \ P) \ P" unfolding le_fun_def by simp lemma le_funD: "f \ g \ f x \ g x" by (rule le_funE) lemma mono_compose: "mono Q \ mono (\i x. Q i (f x))" unfolding mono_def le_fun_def by auto subsection \Order on unary and binary predicates\ lemma predicate1I: assumes PQ: "\x. P x \ Q x" shows "P \ Q" apply (rule le_funI) apply (rule le_boolI) apply (rule PQ) apply assumption done lemma predicate1D: "P \ Q \ P x \ Q x" apply (erule le_funE) apply (erule le_boolE) apply assumption+ done lemma rev_predicate1D: "P x \ P \ Q \ Q x" by (rule predicate1D) lemma predicate2I: assumes PQ: "\x y. P x y \ Q x y" shows "P \ Q" apply (rule le_funI)+ apply (rule le_boolI) apply (rule PQ) apply assumption done lemma predicate2D: "P \ Q \ P x y \ Q x y" apply (erule le_funE)+ apply (erule le_boolE) apply assumption+ done lemma rev_predicate2D: "P x y \ P \ Q \ Q x y" by (rule predicate2D) lemma bot1E [no_atp]: "\ x \ P" by (simp add: bot_fun_def) lemma bot2E: "\ x y \ P" by (simp add: bot_fun_def) lemma top1I: "\ x" by (simp add: top_fun_def) lemma top2I: "\ x y" by (simp add: top_fun_def) subsection \Name duplicates\ lemmas order_eq_refl = preorder_class.eq_refl lemmas order_less_irrefl = preorder_class.less_irrefl lemmas order_less_imp_le = preorder_class.less_imp_le lemmas order_less_not_sym = preorder_class.less_not_sym lemmas order_less_asym = preorder_class.less_asym lemmas order_less_trans = preorder_class.less_trans lemmas order_le_less_trans = preorder_class.le_less_trans lemmas order_less_le_trans = preorder_class.less_le_trans lemmas order_less_imp_not_less = preorder_class.less_imp_not_less lemmas order_less_imp_triv = preorder_class.less_imp_triv lemmas order_less_asym' = preorder_class.less_asym' lemmas order_less_le = order_class.less_le lemmas order_le_less = order_class.le_less lemmas order_le_imp_less_or_eq = order_class.le_imp_less_or_eq lemmas order_less_imp_not_eq = order_class.less_imp_not_eq lemmas order_less_imp_not_eq2 = order_class.less_imp_not_eq2 lemmas order_neq_le_trans = order_class.neq_le_trans lemmas order_le_neq_trans = order_class.le_neq_trans lemmas order_antisym = order_class.antisym lemmas order_eq_iff = order_class.eq_iff lemmas order_antisym_conv = order_class.antisym_conv lemmas linorder_linear = linorder_class.linear lemmas linorder_less_linear = linorder_class.less_linear lemmas linorder_le_less_linear = linorder_class.le_less_linear lemmas linorder_le_cases = linorder_class.le_cases lemmas linorder_not_less = linorder_class.not_less lemmas linorder_not_le = linorder_class.not_le lemmas linorder_neq_iff = linorder_class.neq_iff lemmas linorder_neqE = linorder_class.neqE end