# HG changeset patch # User haftmann # Date 1387983145 -3600 # Node ID 69b3e46d8fbebb32900ed83c30fae558aec97ff5 # Parent 64ff7f16d5b7ca6845596c2f7502d0ee2801ba3f tuned structure of min/max lemmas diff -r 64ff7f16d5b7 -r 69b3e46d8fbe src/HOL/Orderings.thy --- a/src/HOL/Orderings.thy Wed Dec 25 15:52:25 2013 +0100 +++ b/src/HOL/Orderings.thy Wed Dec 25 15:52:25 2013 +0100 @@ -342,47 +342,6 @@ "class.linorder (op \) (op >)" by (rule class.linorder.intro, rule dual_order) (unfold_locales, rule linear) - -text {* min/max *} - -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_le_iff_disj: - "min x y \ z \ x \ z \ y \ z" -unfolding min_def using linear by (auto intro: order_trans) - -lemma le_max_iff_disj: - "z \ max x y \ z \ x \ z \ y" -unfolding max_def using linear by (auto intro: order_trans) - -lemma min_less_iff_disj: - "min x y < z \ x < z \ y < z" -unfolding min_def le_less using less_linear by (auto intro: less_trans) - -lemma less_max_iff_disj: - "z < max x y \ z < x \ z < y" -unfolding max_def le_less using less_linear by (auto intro: less_trans) - -lemma min_less_iff_conj [simp]: - "z < min x y \ z < x \ z < y" -unfolding min_def le_less using less_linear by (auto intro: less_trans) - -lemma max_less_iff_conj [simp]: - "max x y < z \ x < z \ y < z" -unfolding max_def le_less using less_linear by (auto intro: less_trans) - -lemma split_min [no_atp]: - "P (min i j) \ (i \ j \ P i) \ (\ i \ j \ P j)" -by (simp add: min_def) - -lemma split_max [no_atp]: - "P (max i j) \ (i \ j \ P j) \ (\ i \ j \ P i)" -by (simp add: max_def) - end @@ -1028,7 +987,7 @@ *) -subsection {* Monotonicity, least value operator and min/max *} +subsection {* Monotonicity *} context order begin @@ -1140,6 +1099,52 @@ using assms by (auto simp add: less_le Orderings.less_le strict_mono_eq strict_mono_less_eq) +end + + +subsection {* min and max *} + +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)" + +context linorder +begin + +lemma min_le_iff_disj: + "min x y \ z \ x \ z \ y \ z" +unfolding min_def using linear by (auto intro: order_trans) + +lemma le_max_iff_disj: + "z \ max x y \ z \ x \ z \ y" +unfolding max_def using linear by (auto intro: order_trans) + +lemma min_less_iff_disj: + "min x y < z \ x < z \ y < z" +unfolding min_def le_less using less_linear by (auto intro: less_trans) + +lemma less_max_iff_disj: + "z < max x y \ z < x \ z < y" +unfolding max_def le_less using less_linear by (auto intro: less_trans) + +lemma min_less_iff_conj [simp]: + "z < min x y \ z < x \ z < y" +unfolding min_def le_less using less_linear by (auto intro: less_trans) + +lemma max_less_iff_conj [simp]: + "max x y < z \ x < z \ y < z" +unfolding max_def le_less using less_linear by (auto intro: less_trans) + +lemma split_min [no_atp]: + "P (min i j) \ (i \ j \ P i) \ (\ i \ j \ P j)" +by (simp add: min_def) + +lemma split_max [no_atp]: + "P (max i j) \ (i \ j \ P j) \ (\ i \ j \ P i)" +by (simp add: max_def) + lemma min_of_mono: fixes f :: "'a \ 'b\linorder" shows "mono f \ min (f m) (f n) = f (min m n)" @@ -1503,7 +1508,7 @@ unfolding le_fun_def by simp lemma le_funD: "f \ g \ f x \ g x" - unfolding le_fun_def by simp + by (rule le_funE) subsection {* Order on unary and binary predicates *}