diff -r f1b39d77db00 -r ad57323fd4d6 PrioG.thy --- a/PrioG.thy Thu Dec 03 14:34:29 2015 +0800 +++ b/PrioG.thy Tue Dec 15 21:45:46 2015 +0800 @@ -2192,11 +2192,9 @@ hence eq_max: "Max ((\th. preced th s) ` ({th1} \ dependants (wq s) th1)) = Max ((\th. preced th s) ` ({th2} \ dependants (wq s) th2))" (is "Max (?f ` ?A) = Max (?f ` ?B)") - thm cp_def image_Collect unfolding cp_eq_cpreced unfolding cpreced_def . obtain th1' where th1_in: "th1' \ ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)" - thm Max_in proof - have h1: "finite (?f ` ?A)" proof - @@ -2231,7 +2229,6 @@ have "?A \ {}" by simp thus ?thesis by simp qed - thm Max_in from Max_in [OF h1 h2] have "Max (?f ` ?A) \ (?f ` ?A)" . thus ?thesis