diff -r 405f7fd15f4e -r f94b30fa2b6c src/HOL/Library/Ramsey.thy --- a/src/HOL/Library/Ramsey.thy Sun Nov 03 21:12:50 2024 +0100 +++ b/src/HOL/Library/Ramsey.thy Sun Nov 03 22:29:07 2024 +0100 @@ -5,13 +5,13 @@ section \Ramsey's Theorem\ theory Ramsey - imports Infinite_Set Equipollence FuncSet + imports Infinite_Set Equipollence FuncSet begin subsection \Preliminary definitions\ abbreviation strict_sorted :: "'a::linorder list \ bool" where - "strict_sorted \ sorted_wrt (<)" + "strict_sorted \ sorted_wrt (<)" subsubsection \The $n$-element subsets of a set $A$\ @@ -89,12 +89,13 @@ shows "[A]\<^bsup>4\<^esup> = {U. \u x y z. U = {u,x,y,z} \ u \ A \ x \ A \ y \ A \ z \ A \ u < x \ x < y \ y < z}" (is "_ = Collect ?RHS") proof - - { fix U - assume "U \ [A]\<^bsup>4\<^esup>" - then obtain l where "strict_sorted l" "List.set l = U" "length l = 4" "U \ A" + have "?RHS U" if "U \ [A]\<^bsup>4\<^esup>" for U + proof - + from that obtain l where "strict_sorted l" "List.set l = U" "length l = 4" "U \ A" by (simp add: nsets_def) (metis finite_set_strict_sorted) - then have "?RHS U" - unfolding numeral_nat length_Suc_conv by auto blast } + then show ?thesis + unfolding numeral_nat length_Suc_conv by auto blast + qed moreover have "Collect ?RHS \ [A]\<^bsup>4\<^esup>" apply (clarsimp simp add: nsets_def eval_nat_numeral) @@ -109,13 +110,14 @@ shows "[A]\<^bsup>5\<^esup> = {U. \u v x y z. U = {u,v,x,y,z} \ u \ A \ v \ A \ x \ A \ y \ A \ z \ A \ u < v \ v < x \ x < y \ y < z}" (is "_ = Collect ?RHS") proof - - { fix U - assume "U \ [A]\<^bsup>5\<^esup>" - then obtain l where "strict_sorted l" "List.set l = U" "length l = 5" "U \ A" - apply (simp add: nsets_def) - by (metis finite_set_strict_sorted) - then have "?RHS U" - unfolding numeral_nat length_Suc_conv by auto blast } + have "?RHS U" if "U \ [A]\<^bsup>5\<^esup>" for U + proof - + from that obtain l where "strict_sorted l" "List.set l = U" "length l = 5" "U \ A" + apply (simp add: nsets_def) + by (metis finite_set_strict_sorted) + then show ?thesis + unfolding numeral_nat length_Suc_conv by auto blast + qed moreover have "Collect ?RHS \ [A]\<^bsup>5\<^esup>" apply (clarsimp simp add: nsets_def eval_nat_numeral) @@ -338,7 +340,7 @@ using partn_lst_def monochromatic_def assms by metis lemma partn_lst_less: - assumes M: "partn_lst \ \ n" and eq: "length \' = length \" + assumes M: "partn_lst \ \ n" and eq: "length \' = length \" and le: "\i. i < length \ \ \'!i \ \!i " shows "partn_lst \ \' n" proof (clarsimp simp: partn_lst_def) @@ -385,7 +387,7 @@ text \The Erdős--Szekeres bound, essentially extracted from the proof\ fun ES :: "[nat,nat,nat] \ nat" where "ES 0 k l = max k l" - | "ES (Suc r) k l = + | "ES (Suc r) k l = (if r=0 then k+l-1 else if k=0 \ l=0 then 1 else Suc (ES r (ES (Suc r) (k-1) l) (ES (Suc r) k (l-1))))" @@ -685,7 +687,7 @@ using Suc.prems proof (induct k \ "q1 + q2" arbitrary: q1 q2) case 0 - with partn_lst_0 show ?case by auto + with partn_lst_0 show ?case by auto next case (Suc k) consider "q1 = 0 \ q2 = 0" | "q1 \ 0" "q2 \ 0" by auto