diff -r 389ef8b1959c -r f70344294e99 Max.thy --- a/Max.thy Fri Oct 07 21:15:35 2016 +0100 +++ b/Max.thy Fri Oct 21 14:47:01 2016 +0100 @@ -75,4 +75,76 @@ qed +lemma Max_fg_mono: + assumes "finite A" + and "\ a \ A. f a \ g a" + shows "Max (f ` A) \ Max (g ` A)" +proof(cases "A = {}") + case True + thus ?thesis by auto +next + case False + show ?thesis + proof(rule Max.boundedI) + from assms show "finite (f ` A)" by auto + next + from False show "f ` A \ {}" by auto + next + fix fa + assume "fa \ f ` A" + then obtain a where h_fa: "a \ A" "fa = f a" by auto + show "fa \ Max (g ` A)" + proof(rule Max_ge_iff[THEN iffD2]) + from assms show "finite (g ` A)" by auto + next + from False show "g ` A \ {}" by auto + next + from h_fa have "g a \ g ` A" by auto + moreover have "fa \ g a" using h_fa assms(2) by auto + ultimately show "\a\g ` A. fa \ a" by auto + qed + qed +qed + +lemma Max_f_mono: + assumes seq: "A \ B" + and np: "A \ {}" + and fnt: "finite B" + shows "Max (f ` A) \ Max (f ` B)" +proof(rule Max_mono) + from seq show "f ` A \ f ` B" by auto +next + from np show "f ` A \ {}" by auto +next + from fnt and seq show "finite (f ` B)" by auto +qed + +lemma Max_UNION: + assumes "finite A" + and "A \ {}" + and "\ M \ f ` A. finite M" + and "\ M \ f ` A. M \ {}" + shows "Max (\x\ A. f x) = Max (Max ` f ` A)" (is "?L = ?R") + using assms[simp] +proof - + have "?L = Max (\(f ` A))" + by (fold Union_image_eq, simp) + also have "... = ?R" + by (subst Max_Union, simp+) + finally show ?thesis . +qed + +lemma max_Max_eq: + assumes "finite A" + and "A \ {}" + and "x = y" + shows "max x (Max A) = Max ({y} \ A)" (is "?L = ?R") +proof - + have "?R = Max (insert y A)" by simp + also from assms have "... = ?L" + by (subst Max.insert, simp+) + finally show ?thesis by simp +qed + + end