merged
authorwenzelm
Wed, 26 Feb 2020 16:08:05 +0100
changeset 71483 6de04d21c26b
parent 71472 c213d067e60f (diff)
parent 71482 aa7b0a5e9fe3 (current diff)
child 71484 bb82dd4d19f6
merged
src/Tools/VSCode/extension/src/content_provider.ts
src/Tools/VSCode/extension/src/preview.ts
src/Tools/VSCode/src/vscode_javascript.scala
--- a/src/HOL/Analysis/Borel_Space.thy	Wed Feb 26 16:03:29 2020 +0100
+++ b/src/HOL/Analysis/Borel_Space.thy	Wed Feb 26 16:08:05 2020 +0100
@@ -28,78 +28,6 @@
     by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis)
 qed
 
-definition\<^marker>\<open>tag important\<close> "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
-
-lemma mono_onI:
-  "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on f A"
-  unfolding mono_on_def by simp
-
-lemma mono_onD:
-  "\<lbrakk>mono_on f A; r \<in> A; s \<in> A; r \<le> s\<rbrakk> \<Longrightarrow> f r \<le> f s"
-  unfolding mono_on_def by simp
-
-lemma mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
-  unfolding mono_def mono_on_def by auto
-
-lemma mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
-  unfolding mono_on_def by auto
-
-definition\<^marker>\<open>tag important\<close> "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
-
-lemma strict_mono_onI:
-  "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on f A"
-  unfolding strict_mono_on_def by simp
-
-lemma strict_mono_onD:
-  "\<lbrakk>strict_mono_on f A; r \<in> A; s \<in> A; r < s\<rbrakk> \<Longrightarrow> f r < f s"
-  unfolding strict_mono_on_def by simp
-
-lemma mono_on_greaterD:
-  assumes "mono_on g A" "x \<in> A" "y \<in> A" "g x > (g (y::_::linorder) :: _ :: linorder)"
-  shows "x > y"
-proof (rule ccontr)
-  assume "\<not>x > y"
-  hence "x \<le> y" by (simp add: not_less)
-  from assms(1-3) and this have "g x \<le> g y" by (rule mono_onD)
-  with assms(4) show False by simp
-qed
-
-lemma strict_mono_inv:
-  fixes f :: "('a::linorder) \<Rightarrow> ('b::linorder)"
-  assumes "strict_mono f" and "surj f" and inv: "\<And>x. g (f x) = x"
-  shows "strict_mono g"
-proof
-  fix x y :: 'b assume "x < y"
-  from \<open>surj f\<close> obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast
-  with \<open>x < y\<close> and \<open>strict_mono f\<close> have "x' < y'" by (simp add: strict_mono_less)
-  with inv show "g x < g y" by simp
-qed
-
-lemma strict_mono_on_imp_inj_on:
-  assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> (_ :: preorder)) A"
-  shows "inj_on f A"
-proof (rule inj_onI)
-  fix x y assume "x \<in> A" "y \<in> A" "f x = f y"
-  thus "x = y"
-    by (cases x y rule: linorder_cases)
-       (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x])
-qed
-
-lemma strict_mono_on_leD:
-  assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A" "x \<in> A" "y \<in> A" "x \<le> y"
-  shows "f x \<le> f y"
-proof (insert le_less_linear[of y x], elim disjE)
-  assume "x < y"
-  with assms have "f x < f y" by (rule_tac strict_mono_onD[OF assms(1)]) simp_all
-  thus ?thesis by (rule less_imp_le)
-qed (insert assms, simp)
-
-lemma strict_mono_on_eqD:
-  fixes f :: "(_ :: linorder) \<Rightarrow> (_ :: preorder)"
-  assumes "strict_mono_on f A" "f x = f y" "x \<in> A" "y \<in> A"
-  shows "y = x"
-  using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD)
-
 proposition mono_on_imp_deriv_nonneg:
   assumes mono: "mono_on f A" and deriv: "(f has_real_derivative D) (at x)"
   assumes "x \<in> interior A"
@@ -127,10 +55,6 @@
   qed
 qed simp
 
-lemma strict_mono_on_imp_mono_on:
-  "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A \<Longrightarrow> mono_on f A"
-  by (rule mono_onI, rule strict_mono_on_leD)
-
 proposition mono_on_ctble_discont:
   fixes f :: "real \<Rightarrow> real"
   fixes A :: "real set"
--- a/src/HOL/Fun.thy	Wed Feb 26 16:03:29 2020 +0100
+++ b/src/HOL/Fun.thy	Wed Feb 26 16:08:05 2020 +0100
@@ -913,6 +913,84 @@
   then show False by blast
 qed
 
+subsection \<open>Monotonic functions over a set\<close>
+
+definition "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
+
+lemma mono_onI:
+  "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on f A"
+  unfolding mono_on_def by simp
+
+lemma mono_onD:
+  "\<lbrakk>mono_on f A; r \<in> A; s \<in> A; r \<le> s\<rbrakk> \<Longrightarrow> f r \<le> f s"
+  unfolding mono_on_def by simp
+
+lemma mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
+  unfolding mono_def mono_on_def by auto
+
+lemma mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
+  unfolding mono_on_def by auto
+
+definition "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
+
+lemma strict_mono_onI:
+  "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on f A"
+  unfolding strict_mono_on_def by simp
+
+lemma strict_mono_onD:
+  "\<lbrakk>strict_mono_on f A; r \<in> A; s \<in> A; r < s\<rbrakk> \<Longrightarrow> f r < f s"
+  unfolding strict_mono_on_def by simp
+
+lemma mono_on_greaterD:
+  assumes "mono_on g A" "x \<in> A" "y \<in> A" "g x > (g (y::_::linorder) :: _ :: linorder)"
+  shows "x > y"
+proof (rule ccontr)
+  assume "\<not>x > y"
+  hence "x \<le> y" by (simp add: not_less)
+  from assms(1-3) and this have "g x \<le> g y" by (rule mono_onD)
+  with assms(4) show False by simp
+qed
+
+lemma strict_mono_inv:
+  fixes f :: "('a::linorder) \<Rightarrow> ('b::linorder)"
+  assumes "strict_mono f" and "surj f" and inv: "\<And>x. g (f x) = x"
+  shows "strict_mono g"
+proof
+  fix x y :: 'b assume "x < y"
+  from \<open>surj f\<close> obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast
+  with \<open>x < y\<close> and \<open>strict_mono f\<close> have "x' < y'" by (simp add: strict_mono_less)
+  with inv show "g x < g y" by simp
+qed
+
+lemma strict_mono_on_imp_inj_on:
+  assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> (_ :: preorder)) A"
+  shows "inj_on f A"
+proof (rule inj_onI)
+  fix x y assume "x \<in> A" "y \<in> A" "f x = f y"
+  thus "x = y"
+    by (cases x y rule: linorder_cases)
+       (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x])
+qed
+
+lemma strict_mono_on_leD:
+  assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A" "x \<in> A" "y \<in> A" "x \<le> y"
+  shows "f x \<le> f y"
+proof (insert le_less_linear[of y x], elim disjE)
+  assume "x < y"
+  with assms have "f x < f y" by (rule_tac strict_mono_onD[OF assms(1)]) simp_all
+  thus ?thesis by (rule less_imp_le)
+qed (insert assms, simp)
+
+lemma strict_mono_on_eqD:
+  fixes f :: "(_ :: linorder) \<Rightarrow> (_ :: preorder)"
+  assumes "strict_mono_on f A" "f x = f y" "x \<in> A" "y \<in> A"
+  shows "y = x"
+  using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD)
+
+lemma strict_mono_on_imp_mono_on:
+  "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A \<Longrightarrow> mono_on f A"
+  by (rule mono_onI, rule strict_mono_on_leD)
+
 
 subsection \<open>Setup\<close>
 
--- a/src/HOL/Library/Ramsey.thy	Wed Feb 26 16:03:29 2020 +0100
+++ b/src/HOL/Library/Ramsey.thy	Wed Feb 26 16:08:05 2020 +0100
@@ -40,25 +40,25 @@
     by (metis insert_iff singletonD)
 qed
 
-lemma nsets_disjoint_2: 
+lemma nsets_disjoint_2:
   "X \<inter> Y = {} \<Longrightarrow> [X \<union> Y]\<^bsup>2\<^esup> = [X]\<^bsup>2\<^esup> \<union> [Y]\<^bsup>2\<^esup> \<union> (\<Union>x\<in>X. \<Union>y\<in>Y. {{x,y}})"
   by (fastforce simp: nsets_2_eq Set.doubleton_eq_iff)
 
-lemma ordered_nsets_2_eq: 
-  fixes A :: "'a::linorder set" 
+lemma ordered_nsets_2_eq:
+  fixes A :: "'a::linorder set"
   shows "nsets A 2 = {{x,y} | x y. x \<in> A \<and> y \<in> A \<and> x<y}"
      (is "_ = ?rhs")
 proof
   show "nsets A 2 \<subseteq> ?rhs"
-    unfolding numeral_nat 
+    unfolding numeral_nat
     apply (clarsimp simp add: nsets_def card_Suc_eq Set.doubleton_eq_iff not_less)
-    by (metis antisym) 
+    by (metis antisym)
   show "?rhs \<subseteq> nsets A 2"
     unfolding numeral_nat by (auto simp: nsets_def card_Suc_eq)
 qed
 
-lemma ordered_nsets_3_eq: 
-  fixes A :: "'a::linorder set" 
+lemma ordered_nsets_3_eq:
+  fixes A :: "'a::linorder set"
   shows "nsets A 3 = {{x,y,z} | x y z. x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> x<y \<and> y<z}"
      (is "_ = ?rhs")
 proof
@@ -70,8 +70,8 @@
   by (metis empty_iff insert_iff not_less_iff_gr_or_eq)
 qed
 
-lemma ordered_nsets_4_eq: 
-  fixes A :: "'a::linorder set" 
+lemma ordered_nsets_4_eq:
+  fixes A :: "'a::linorder set"
   shows "[A]\<^bsup>4\<^esup> = {U. \<exists>u x y z. U = {u,x,y,z} \<and> u \<in> A \<and> x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> u < x \<and> x < y \<and> y < z}"
     (is "_ = Collect ?RHS")
 proof -
@@ -90,8 +90,8 @@
     by auto
 qed
 
-lemma ordered_nsets_5_eq: 
-  fixes A :: "'a::linorder set" 
+lemma ordered_nsets_5_eq:
+  fixes A :: "'a::linorder set"
   shows "[A]\<^bsup>5\<^esup> = {U. \<exists>u v x y z. U = {u,v,x,y,z} \<and> u \<in> A \<and> v \<in> A \<and> x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> u < v \<and> v < x \<and> x < y \<and> y < z}"
     (is "_ = Collect ?RHS")
 proof -
@@ -122,7 +122,7 @@
   show "finite A"
     using infinite_arbitrarily_large that by auto
   then have "\<not> r \<le> card A"
-    using that by (simp add: set_eq_iff) (metis finite_subset get_smaller_card [of A r])
+    using that by (simp add: set_eq_iff) (metis obtain_subset_with_card_n)
   then show "card A < r"
     using not_less by blast
 next
@@ -189,7 +189,7 @@
   have "(\<lambda>X. g ` X) \<in> [S]\<^bsup>k\<^esup> \<rightarrow> [T]\<^bsup>k\<^esup>"
     using assms by (simp add: nsets_image_funcset)
   then show ?thesis
-    using f by fastforce 
+    using f by fastforce
 qed
 
 subsubsection \<open>Partition predicates\<close>
@@ -198,7 +198,7 @@
   where "partn \<beta> \<alpha> \<gamma> \<delta> \<equiv> \<forall>f \<in> nsets \<beta> \<gamma>  \<rightarrow>  \<delta>. \<exists>H \<in> nsets \<beta> \<alpha>. \<exists>\<xi>\<in>\<delta>. f ` (nsets H \<gamma>) \<subseteq> {\<xi>}"
 
 definition partn_lst :: "'a set \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> bool"
-  where "partn_lst \<beta> \<alpha> \<gamma> \<equiv> \<forall>f \<in> nsets \<beta> \<gamma>  \<rightarrow>  {..<length \<alpha>}. 
+  where "partn_lst \<beta> \<alpha> \<gamma> \<equiv> \<forall>f \<in> nsets \<beta> \<gamma>  \<rightarrow>  {..<length \<alpha>}.
               \<exists>i < length \<alpha>. \<exists>H \<in> nsets \<beta> (\<alpha>!i). f ` (nsets H \<gamma>) \<subseteq> {i}"
 
 lemma partn_lst_greater_resource:
@@ -247,8 +247,8 @@
 lemma partn_lst_eq_partn: "partn_lst {..<n} [m,m] 2 = partn {..<n} m 2 {..<2::nat}"
   apply (simp add: partn_lst_def partn_def numeral_2_eq_2)
   by (metis less_2_cases numeral_2_eq_2 lessThan_iff nth_Cons_0 nth_Cons_Suc)
-  
- 
+
+
 subsection \<open>Finite versions of Ramsey's theorem\<close>
 
 text \<open>
@@ -266,7 +266,7 @@
 lemma ramsey1: "\<exists>N::nat. partn_lst {..<N} [q0,q1] 1"
 proof -
   have "\<exists>i<Suc (Suc 0). \<exists>H\<in>nsets {..<Suc (q0 + q1)} ([q0, q1] ! i). f ` nsets H (Suc 0) \<subseteq> {i}"
-    if "f \<in> nsets {..<Suc (q0 + q1)} (Suc 0) \<rightarrow> {..<Suc (Suc 0)}" for f 
+    if "f \<in> nsets {..<Suc (q0 + q1)} (Suc 0) \<rightarrow> {..<Suc (Suc 0)}" for f
   proof -
     define A where "A \<equiv> \<lambda>i. {q. q \<le> q0+q1 \<and> f {q} = i}"
     have "A 0 \<union> A 1 = {..q0 + q1}"
@@ -280,7 +280,7 @@
     then obtain i where "i < Suc (Suc 0)" "card (A i) \<ge> [q0, q1] ! i"
       by (metis One_nat_def lessI nth_Cons_0 nth_Cons_Suc zero_less_Suc)
     then obtain B where "B \<subseteq> A i" "card B = [q0, q1] ! i" "finite B"
-      by (meson finite_subset get_smaller_card infinite_arbitrarily_large)
+      by (meson obtain_subset_with_card_n)
     then have "B \<in> nsets {..<Suc (q0 + q1)} ([q0, q1] ! i) \<and> f ` nsets B (Suc 0) \<subseteq> {i}"
       by (auto simp: A_def nsets_def card_1_singleton_iff)
     then show ?thesis
@@ -297,11 +297,11 @@
 proof (induction r arbitrary: q1 q2)
   case 0
   then show ?case
-    by (simp add: ramsey0) 
+    by (simp add: ramsey0)
 next
   case (Suc r)
   note outer = this
-  show ?case 
+  show ?case
   proof (cases "r = 0")
     case True
     then show ?thesis
@@ -314,7 +314,7 @@
       using Suc.prems
     proof (induct k \<equiv> "q1 + q2" arbitrary: q1 q2)
       case 0
-      show ?case 
+      show ?case
       proof
         show "partn_lst {..<1::nat} [q1, q2] (Suc r)"
           using nsets_empty_iff subset_insert 0
@@ -323,7 +323,7 @@
     next
       case (Suc k)
       consider "q1 = 0 \<or> q2 = 0" | "q1 \<noteq> 0" "q2 \<noteq> 0" by auto
-      then show ?case 
+      then show ?case
       proof cases
         case 1
         then have "partn_lst {..< Suc 0} [q1, q2] (Suc r)"
@@ -335,12 +335,12 @@
         with Suc have "k = (q1 - 1) + q2" "k = q1 + (q2 - 1)" by auto
         then obtain p1 p2::nat where  p1: "partn_lst {..<p1} [q1-1,q2] (Suc r)" and p2: "partn_lst {..<p2} [q1,q2-1] (Suc r)"
           using Suc.hyps by blast
-        then obtain p::nat where p: "partn_lst {..<p} [p1,p2] r" 
+        then obtain p::nat where p: "partn_lst {..<p} [p1,p2] r"
           using outer Suc.prems by auto
         show ?thesis
         proof (intro exI conjI)
           have "\<exists>i<Suc (Suc 0). \<exists>H\<in>nsets {..p} ([q1,q2] ! i). f ` nsets H (Suc r) \<subseteq> {i}"
-            if f: "f \<in> nsets {..p} (Suc r) \<rightarrow> {..<Suc (Suc 0)}" for f 
+            if f: "f \<in> nsets {..p} (Suc r) \<rightarrow> {..<Suc (Suc 0)}" for f
           proof -
             define g where "g \<equiv> \<lambda>R. f (insert p R)"
             have "f (insert p i) \<in> {..<Suc (Suc 0)}" if "i \<in> nsets {..<p} r" for i
@@ -348,7 +348,7 @@
             then have g: "g \<in> nsets {..<p} r \<rightarrow> {..<Suc (Suc 0)}"
               by (force simp: g_def PiE_iff)
             then obtain i U where i: "i < Suc (Suc 0)" and gi: "g ` nsets U r \<subseteq> {i}"
-              and U: "U \<in> nsets {..<p} ([p1, p2] ! i)" 
+              and U: "U \<in> nsets {..<p} ([p1, p2] ! i)"
               using p by (auto simp: partn_lst_def)
             then have Usub: "U \<subseteq> {..<p}"
               by (auto simp: nsets_def)
@@ -359,7 +359,7 @@
               case izero
               then have "U \<in> nsets {..<p} p1"
                 using U by simp
-              then obtain u where u: "bij_betw u {..<p1} U" 
+              then obtain u where u: "bij_betw u {..<p1} U"
                 using ex_bij_betw_nat_finite lessThan_atLeast0 by (fastforce simp add: nsets_def)
               have u_nsets: "u ` X \<in> nsets {..p} n" if "X \<in> nsets {..<p1} n" for X n
               proof -
@@ -373,7 +373,7 @@
               have "h \<in> nsets {..<p1} (Suc r) \<rightarrow> {..<Suc (Suc 0)}"
                 unfolding h_def using f u_nsets by auto
               then obtain j V where j: "j <Suc (Suc 0)" and hj: "h ` nsets V (Suc r) \<subseteq> {j}"
-                and V: "V \<in> nsets {..<p1} ([q1 - Suc 0, q2] ! j)" 
+                and V: "V \<in> nsets {..<p1} ([q1 - Suc 0, q2] ! j)"
                 using p1 by (auto simp: partn_lst_def)
               then have Vsub: "V \<subseteq> {..<p1}"
                 by (auto simp: nsets_def)
@@ -389,11 +389,11 @@
                   using V by simp
                 then have "u ` V \<in> nsets {..<p} (q1 - Suc 0)"
                   using u_nsets [of _ "q1 - Suc 0"] nsets_mono [OF Vsub] Usub u
-                  unfolding bij_betw_def nsets_def 
+                  unfolding bij_betw_def nsets_def
                   by (fastforce elim!: subsetD)
                 then have inq1: "?W \<in> nsets {..p} q1"
                   unfolding nsets_def using \<open>q1 \<noteq> 0\<close> card_insert_if by fastforce
-                have invu_nsets: "inv_into {..<p1} u ` X \<in> nsets V r" 
+                have invu_nsets: "inv_into {..<p1} u ` X \<in> nsets V r"
                   if "X \<in> nsets (u ` V) r" for X r
                 proof -
                   have "X \<subseteq> u ` V \<and> finite X \<and> card X = r"
@@ -418,13 +418,13 @@
                     using gi True \<open>X - {p} \<in> nsets U r\<close> insert_Diff
                     by (fastforce simp add: g_def image_subset_iff)
                   then show ?thesis
-                    by (simp add: \<open>f X = i\<close> \<open>g (X - {p}) = i\<close>) 
+                    by (simp add: \<open>f X = i\<close> \<open>g (X - {p}) = i\<close>)
                 next
                   case False
                   then have Xim: "X \<in> nsets (u ` V) (Suc r)"
                     using X by (auto simp: nsets_def subset_insert)
                   then have "u ` inv_into {..<p1} u ` X = X"
-                    using Vsub bij_betw_imp_inj_on u 
+                    using Vsub bij_betw_imp_inj_on u
                     by (fastforce simp: nsets_def image_mono invinv_eq subset_trans)
                   then show ?thesis
                     using izero jzero hj Xim invu_nsets unfolding h_def
@@ -433,22 +433,22 @@
                 moreover have "insert p (u ` V) \<in> nsets {..p} q1"
                   by (simp add: izero inq1)
                 ultimately show ?thesis
-                  by (metis izero image_subsetI insertI1 nth_Cons_0 zero_less_Suc) 
+                  by (metis izero image_subsetI insertI1 nth_Cons_0 zero_less_Suc)
               next
                 case jone
                 then have "u ` V \<in> nsets {..p} q2"
                   using V u_nsets by auto
                 moreover have "f ` nsets (u ` V) (Suc r) \<subseteq> {j}"
-                  using hj 
+                  using hj
                   by (force simp add: h_def image_subset_iff nsets_def subset_image_inj card_image dest: finite_imageD)
                 ultimately show ?thesis
                   using jone not_less_eq by fastforce
               qed
-            next 
+            next
               case ione
               then have "U \<in> nsets {..<p} p2"
                 using U by simp
-              then obtain u where u: "bij_betw u {..<p2} U" 
+              then obtain u where u: "bij_betw u {..<p2} U"
                 using ex_bij_betw_nat_finite lessThan_atLeast0 by (fastforce simp add: nsets_def)
               have u_nsets: "u ` X \<in> nsets {..p} n" if "X \<in> nsets {..<p2} n" for X n
               proof -
@@ -462,7 +462,7 @@
               have "h \<in> nsets {..<p2} (Suc r) \<rightarrow> {..<Suc (Suc 0)}"
                 unfolding h_def using f u_nsets by auto
               then obtain j V where j: "j <Suc (Suc 0)" and hj: "h ` nsets V (Suc r) \<subseteq> {j}"
-                and V: "V \<in> nsets {..<p2} ([q1, q2 - Suc 0] ! j)" 
+                and V: "V \<in> nsets {..<p2} ([q1, q2 - Suc 0] ! j)"
                 using p2 by (auto simp: partn_lst_def)
               then have Vsub: "V \<subseteq> {..<p2}"
                 by (auto simp: nsets_def)
@@ -478,11 +478,11 @@
                   using V by simp
                 then have "u ` V \<in> nsets {..<p} (q2 - Suc 0)"
                   using u_nsets [of _ "q2 - Suc 0"] nsets_mono [OF Vsub] Usub u
-                  unfolding bij_betw_def nsets_def 
+                  unfolding bij_betw_def nsets_def
                   by (fastforce elim!: subsetD)
                 then have inq1: "?W \<in> nsets {..p} q2"
                   unfolding nsets_def using \<open>q2 \<noteq> 0\<close> card_insert_if by fastforce
-                have invu_nsets: "inv_into {..<p2} u ` X \<in> nsets V r" 
+                have invu_nsets: "inv_into {..<p2} u ` X \<in> nsets V r"
                   if "X \<in> nsets (u ` V) r" for X r
                 proof -
                   have "X \<subseteq> u ` V \<and> finite X \<and> card X = r"
@@ -507,13 +507,13 @@
                     using gi True \<open>X - {p} \<in> nsets U r\<close> insert_Diff
                     by (fastforce simp add: g_def image_subset_iff)
                   then show ?thesis
-                    by (simp add: \<open>f X = i\<close> \<open>g (X - {p}) = i\<close>) 
+                    by (simp add: \<open>f X = i\<close> \<open>g (X - {p}) = i\<close>)
                 next
                   case False
                   then have Xim: "X \<in> nsets (u ` V) (Suc r)"
                     using X by (auto simp: nsets_def subset_insert)
                   then have "u ` inv_into {..<p2} u ` X = X"
-                    using Vsub bij_betw_imp_inj_on u 
+                    using Vsub bij_betw_imp_inj_on u
                     by (fastforce simp: nsets_def image_mono invinv_eq subset_trans)
                   then show ?thesis
                     using ione jone hj Xim invu_nsets unfolding h_def
@@ -528,7 +528,7 @@
                 then have "u ` V \<in> nsets {..p} q1"
                   using V u_nsets by auto
                 moreover have "f ` nsets (u ` V) (Suc r) \<subseteq> {j}"
-                  using hj 
+                  using hj
                   apply (clarsimp simp add: h_def image_subset_iff nsets_def)
                   by (metis Zero_not_Suc card_eq_0_iff card_image subset_image_inj)
                 ultimately show ?thesis
@@ -566,9 +566,9 @@
     then obtain q1 q2 l where qs: "qs = q1#q2#l"
       by (metis Suc.hyps(2) length_Suc_conv)
     then obtain q::nat where q: "partn_lst {..<q} [q1,q2] r"
-      using ramsey2_full by blast 
+      using ramsey2_full by blast
     then obtain p::nat where p: "partn_lst {..<p} (q#l) r"
-      using IH \<open>qs = q1 # q2 # l\<close> by fastforce 
+      using IH \<open>qs = q1 # q2 # l\<close> by fastforce
     have keq: "Suc (length l) = k"
       using IH qs by auto
     show ?thesis
@@ -582,21 +582,21 @@
           unfolding g_def using f Suc IH
           by (auto simp: Pi_def not_less)
         then obtain i U where i: "i < k" and gi: "g ` nsets U r \<subseteq> {i}"
-                and U: "U \<in> nsets {..<p} ((q#l) ! i)" 
+                and U: "U \<in> nsets {..<p} ((q#l) ! i)"
           using p keq by (auto simp: partn_lst_def)
         show "\<exists>i<length qs. \<exists>H\<in>nsets {..<p} (qs ! i). f ` nsets H r \<subseteq> {i}"
         proof (cases "i = 0")
           case True
           then have "U \<in> nsets {..<p} q" and f01: "f ` nsets U r \<subseteq> {0, Suc 0}"
             using U gi unfolding g_def by (auto simp: image_subset_iff)
-          then obtain u where u: "bij_betw u {..<q} U" 
+          then obtain u where u: "bij_betw u {..<q} U"
             using ex_bij_betw_nat_finite lessThan_atLeast0 by (fastforce simp add: nsets_def)
           then have Usub: "U \<subseteq> {..<p}"
             by (smt \<open>U \<in> nsets {..<p} q\<close> mem_Collect_eq nsets_def)
           have u_nsets: "u ` X \<in> nsets {..<p} n" if "X \<in> nsets {..<q} n" for X n
           proof -
             have "inj_on u X"
-              using u that bij_betw_imp_inj_on inj_on_subset              
+              using u that bij_betw_imp_inj_on inj_on_subset
               by (force simp: nsets_def)
             then show ?thesis
               using Usub u that bij_betwE
@@ -611,9 +611,9 @@
               using f01 by auto
           qed
           then have "h \<in> nsets {..<q} r \<rightarrow> {..<Suc (Suc 0)}"
-            unfolding h_def by blast  
+            unfolding h_def by blast
           then obtain j V where j: "j < Suc (Suc 0)" and hj: "h ` nsets V r \<subseteq> {j}"
-            and V: "V \<in> nsets {..<q} ([q1,q2] ! j)" 
+            and V: "V \<in> nsets {..<q} ([q1,q2] ! j)"
             using q by (auto simp: partn_lst_def)
           show ?thesis
           proof (intro exI conjI bexI)
@@ -663,7 +663,7 @@
 proof -
   obtain N where "N \<ge> Suc 0" and N: "partn_lst {..<N} [m,n] 2"
     using ramsey2_full nat_le_linear partn_lst_greater_resource by blast
-  have "\<exists>R\<subseteq>V. card R = m \<and> clique R E \<or> card R = n \<and> indep R E" 
+  have "\<exists>R\<subseteq>V. card R = m \<and> clique R E \<or> card R = n \<and> indep R E"
     if "finite V" "N \<le> card V" for V :: "'a set" and E :: "'a set set"
   proof -
     from that
@@ -673,7 +673,7 @@
     have f: "f \<in> nsets {..<N} 2 \<rightarrow> {..<Suc (Suc 0)}"
       by (simp add: f_def)
     then obtain i U where i: "i < 2" and gi: "f ` nsets U 2 \<subseteq> {i}"
-      and U: "U \<in> nsets {..<N} ([m,n] ! i)" 
+      and U: "U \<in> nsets {..<N} ([m,n] ! i)"
       using N numeral_2_eq_2 by (auto simp: partn_lst_def)
     show ?thesis
     proof (intro exI conjI)
--- a/src/HOL/Set_Interval.thy	Wed Feb 26 16:03:29 2020 +0100
+++ b/src/HOL/Set_Interval.thy	Wed Feb 26 16:08:05 2020 +0100
@@ -17,10 +17,10 @@
 
 (* Belongs in Finite_Set but 2 is not available there *)
 lemma card_2_iff: "card S = 2 \<longleftrightarrow> (\<exists>x y. S = {x,y} \<and> x \<noteq> y)"
-by (auto simp: card_Suc_eq numeral_eq_Suc)
+  by (auto simp: card_Suc_eq numeral_eq_Suc)
 
 lemma card_2_iff': "card S = 2 \<longleftrightarrow> (\<exists>x\<in>S. \<exists>y\<in>S. x \<noteq> y \<and> (\<forall>z\<in>S. z = x \<or> z = y))"
-by (auto simp: card_Suc_eq numeral_eq_Suc)
+  by (auto simp: card_Suc_eq numeral_eq_Suc)
 
 context ord
 begin
@@ -1651,20 +1651,21 @@
    apply (force intro: leI)+
   done
 
-lemma get_smaller_card:
-  assumes "finite A" "k \<le> card A"
-  obtains B where "B \<subseteq> A" "card B = k"
+lemma obtain_subset_with_card_n:
+  assumes "n \<le> card S"
+  obtains T where "T \<subseteq> S" "card T = n" "finite T"
 proof -
-  obtain h where h: "bij_betw h {0..<card A} A"
-    using \<open>finite A\<close> ex_bij_betw_nat_finite by blast
-  show thesis
-  proof
-    show "h ` {0..<k} \<subseteq> A"
-      by (metis \<open>k \<le> card A\<close> bij_betw_def h image_mono ivl_subset zero_le)
-    have "inj_on h {0..<k}"
-      by (meson \<open>k \<le> card A\<close> bij_betw_def h inj_on_subset ivl_subset zero_le)
-    then show "card (h ` {0..<k}) = k"
-      by (simp add: card_image)
+  obtain n' where "card S = n + n'" 
+    by (metis assms le_add_diff_inverse)
+  with that show thesis
+  proof (induct n' arbitrary: S)
+    case 0 
+    then show ?case
+      by (cases "finite S") auto
+  next
+    case Suc 
+    then show ?case 
+      by (simp add: card_Suc_eq) (metis subset_insertI2)
   qed
 qed
 
--- a/src/HOL/ex/Erdoes_Szekeres.thy	Wed Feb 26 16:03:29 2020 +0100
+++ b/src/HOL/ex/Erdoes_Szekeres.thy	Wed Feb 26 16:08:05 2020 +0100
@@ -1,69 +1,34 @@
-(*   Title: HOL/ex/Erdoes_Szekeres.thy
+(*   Title: HOL/ex/Erdős_Szekeres.thy
      Author: Lukas Bulwahn <lukas.bulwahn-at-gmail.com>
 *)
 
-section \<open>The Erdoes-Szekeres Theorem\<close>
+section \<open>The Erdös-Szekeres Theorem\<close>
 
 theory Erdoes_Szekeres
 imports Main
 begin
 
-subsection \<open>Addition to \<^theory>\<open>HOL.Lattices_Big\<close> Theory\<close>
-
-lemma Max_gr:
-  assumes "finite A"
-  assumes "a \<in> A" "a > x"
-  shows "x < Max A"
-using assms Max_ge less_le_trans by blast
-
-subsection \<open>Additions to \<^theory>\<open>HOL.Finite_Set\<close> Theory\<close>
-
-lemma obtain_subset_with_card_n:
-  assumes "n \<le> card S"
-  obtains T where "T \<subseteq> S" "card T = n"
-proof -
-  from assms obtain n' where "card S = n + n'" by (metis le_add_diff_inverse)
-  from this that show ?thesis
-  proof (induct n' arbitrary: S)
-    case 0 from this show ?case by auto
-  next
-    case Suc from this show ?case by (simp add: card_Suc_eq) (metis subset_insertI2)
-  qed
-qed
-
 lemma exists_set_with_max_card:
   assumes "finite S" "S \<noteq> {}"
   shows "\<exists>s \<in> S. card s = Max (card ` S)"
-using assms
-proof (induct S rule: finite.induct)
-  case (insertI S' s')
-  show ?case
-  proof (cases "S' \<noteq> {}")
-    case True
-    from this insertI.hyps(2) obtain s where s: "s \<in> S'" "card s = Max (card ` S')" by auto
-    from this(1) have that: "(if card s \<ge> card s' then s else s') \<in> insert s' S'" by auto
-    have "card (if card s \<ge> card s' then s else s') = Max (card ` insert s' S')"
-      using insertI(1) \<open>S' \<noteq> {}\<close> s by auto
-    from this that show ?thesis by blast
-  qed (auto)
-qed (auto)
+  by (metis (mono_tags, lifting) Max_in assms finite_imageI imageE image_is_empty)
 
 subsection \<open>Definition of Monotonicity over a Carrier Set\<close>
 
 definition
-  "mono_on f R S = (\<forall>i\<in>S. \<forall>j\<in>S. i \<le> j \<longrightarrow> R (f i) (f j))"
+  "gen_mono_on f R S = (\<forall>i\<in>S. \<forall>j\<in>S. i \<le> j \<longrightarrow> R (f i) (f j))"
 
-lemma mono_on_empty [simp]: "mono_on f R {}"
-unfolding mono_on_def by auto
+lemma gen_mono_on_empty [simp]: "gen_mono_on f R {}"
+  unfolding gen_mono_on_def by auto
 
-lemma mono_on_singleton [simp]: "reflp R \<Longrightarrow> mono_on f R {x}"
-unfolding mono_on_def reflp_def by auto
+lemma gen_mono_on_singleton [simp]: "reflp R \<Longrightarrow> gen_mono_on f R {x}"
+  unfolding gen_mono_on_def reflp_def by auto
 
-lemma mono_on_subset: "T \<subseteq> S \<Longrightarrow> mono_on f R S \<Longrightarrow> mono_on f R T"
-unfolding mono_on_def by (simp add: subset_iff)
+lemma gen_mono_on_subset: "T \<subseteq> S \<Longrightarrow> gen_mono_on f R S \<Longrightarrow> gen_mono_on f R T"
+  unfolding gen_mono_on_def by (simp add: subset_iff)
 
-lemma not_mono_on_subset: "T \<subseteq> S \<Longrightarrow> \<not> mono_on f R T \<Longrightarrow> \<not> mono_on f R S"
-unfolding mono_on_def by blast
+lemma not_gen_mono_on_subset: "T \<subseteq> S \<Longrightarrow> \<not> gen_mono_on f R T \<Longrightarrow> \<not> gen_mono_on f R S"
+  unfolding gen_mono_on_def by blast
 
 lemma [simp]:
   "reflp ((\<le>) :: 'a::order \<Rightarrow> _ \<Rightarrow> bool)"
@@ -76,35 +41,35 @@
 
 lemma Erdoes_Szekeres:
   fixes f :: "_ \<Rightarrow> 'a::linorder"
-  shows "(\<exists>S. S \<subseteq> {0..m * n} \<and> card S = m + 1 \<and> mono_on f (\<le>) S) \<or>
-         (\<exists>S. S \<subseteq> {0..m * n} \<and> card S = n + 1 \<and> mono_on f (\<ge>) S)"
+  shows "(\<exists>S. S \<subseteq> {0..m * n} \<and> card S = m + 1 \<and> gen_mono_on f (\<le>) S) \<or>
+         (\<exists>S. S \<subseteq> {0..m * n} \<and> card S = n + 1 \<and> gen_mono_on f (\<ge>) S)"
 proof (rule ccontr)
-  let ?max_subseq = "\<lambda>R k. Max (card ` {S. S \<subseteq> {0..k} \<and> mono_on f R S \<and> k \<in> S})"
+  let ?max_subseq = "\<lambda>R k. Max (card ` {S. S \<subseteq> {0..k} \<and> gen_mono_on f R S \<and> k \<in> S})"
   define phi where "phi k = (?max_subseq (\<le>) k, ?max_subseq (\<ge>) k)" for k
 
-  have one_member: "\<And>R k. reflp R \<Longrightarrow> {k} \<in> {S. S \<subseteq> {0..k} \<and> mono_on f R S \<and> k \<in> S}" by auto
+  have one_member: "\<And>R k. reflp R \<Longrightarrow> {k} \<in> {S. S \<subseteq> {0..k} \<and> gen_mono_on f R S \<and> k \<in> S}" by auto
 
   {
     fix R
     assume reflp: "reflp (R :: 'a::linorder \<Rightarrow> _)"
-    from one_member[OF this] have non_empty: "\<And>k. {S. S \<subseteq> {0..k} \<and> mono_on f R S \<and> k \<in> S} \<noteq> {}" by force
-    from one_member[OF reflp] have "\<And>k. card {k} \<in> card ` {S. S \<subseteq> {0..k} \<and> mono_on f R S \<and> k \<in> S}" by blast
+    from one_member[OF this] have non_empty: "\<And>k. {S. S \<subseteq> {0..k} \<and> gen_mono_on f R S \<and> k \<in> S} \<noteq> {}" by force
+    from one_member[OF reflp] have "\<And>k. card {k} \<in> card ` {S. S \<subseteq> {0..k} \<and> gen_mono_on f R S \<and> k \<in> S}" by blast
     from this have lower_bound: "\<And>k. k \<le> m * n \<Longrightarrow> ?max_subseq R k \<ge> 1"
       by (auto intro!: Max_ge)
 
     fix b
-    assume not_mono_at: "\<forall>S. S \<subseteq> {0..m * n} \<and> card S = b + 1 \<longrightarrow> \<not> mono_on f R S"
+    assume not_mono_at: "\<forall>S. S \<subseteq> {0..m * n} \<and> card S = b + 1 \<longrightarrow> \<not> gen_mono_on f R S"
 
     {
       fix S
       assume "S \<subseteq> {0..m * n}" "card S \<ge> b + 1"
       moreover from \<open>card S \<ge> b + 1\<close> obtain T where "T \<subseteq> S \<and> card T = Suc b"
         using obtain_subset_with_card_n by (metis Suc_eq_plus1)
-      ultimately have "\<not> mono_on f R S" using not_mono_at by (auto dest: not_mono_on_subset)
+      ultimately have "\<not> gen_mono_on f R S" using not_mono_at by (auto dest: not_gen_mono_on_subset)
     }
-    from this have "\<forall>S. S \<subseteq> {0..m * n} \<and> mono_on f R S \<longrightarrow> card S \<le> b"
+    from this have "\<forall>S. S \<subseteq> {0..m * n} \<and> gen_mono_on f R S \<longrightarrow> card S \<le> b"
       by (metis Suc_eq_plus1 Suc_leI not_le)
-    from this have "\<And>k. k \<le> m * n \<Longrightarrow> \<forall>S. S \<subseteq> {0..k} \<and> mono_on f R S \<longrightarrow> card S \<le> b"
+    from this have "\<And>k. k \<le> m * n \<Longrightarrow> \<forall>S. S \<subseteq> {0..k} \<and> gen_mono_on f R S \<longrightarrow> card S \<le> b"
       using order_trans by force
     from this non_empty have upper_bound: "\<And>k. k \<le> m * n \<Longrightarrow> ?max_subseq R k \<le> b"
       by (auto intro: Max.boundedI)
@@ -132,19 +97,19 @@
       fix R
       assume R: "reflp (R :: 'a::linorder \<Rightarrow> _)" "transp R" "R (f i) (f j)"
       from one_member[OF \<open>reflp R\<close>, of "i"] have
-        "\<exists>S \<in> {S. S \<subseteq> {0..i} \<and> mono_on f R S \<and> i \<in> S}. card S = ?max_subseq R i"
+        "\<exists>S \<in> {S. S \<subseteq> {0..i} \<and> gen_mono_on f R S \<and> i \<in> S}. card S = ?max_subseq R i"
         by (intro exists_set_with_max_card) auto
-      from this obtain S where S: "S \<subseteq> {0..i} \<and> mono_on f R S \<and> i \<in> S" "card S = ?max_subseq R i" by auto
+      from this obtain S where S: "S \<subseteq> {0..i} \<and> gen_mono_on f R S \<and> i \<in> S" "card S = ?max_subseq R i" by auto
       from S \<open>i < j\<close> finite_subset have "j \<notin> S" "finite S" "insert j S \<subseteq> {0..j}" by auto
-      from S(1) R \<open>i < j\<close> this have "mono_on f R (insert j S)"
-        unfolding mono_on_def reflp_def transp_def
+      from S(1) R \<open>i < j\<close> this have "gen_mono_on f R (insert j S)"
+        unfolding gen_mono_on_def reflp_def transp_def
         by (metis atLeastAtMost_iff insert_iff le_antisym subsetCE)
-      from this have d: "insert j S \<in> {S. S \<subseteq> {0..j} \<and> mono_on f R S \<and> j \<in> S}"
+      from this have d: "insert j S \<in> {S. S \<subseteq> {0..j} \<and> gen_mono_on f R S \<and> j \<in> S}"
         using \<open>insert j S \<subseteq> {0..j}\<close> by blast
       from this \<open>j \<notin> S\<close> S(1) have "card (insert j S) \<in>
-        card ` {S. S \<subseteq> {0..j} \<and> mono_on f R S \<and> j \<in> S} \<and> card S < card (insert j S)"
+        card ` {S. S \<subseteq> {0..j} \<and> gen_mono_on f R S \<and> j \<in> S} \<and> card S < card (insert j S)"
         by (auto intro!: imageI) (auto simp add: \<open>finite S\<close>)
-      from this S(2) have "?max_subseq R i < ?max_subseq R j" by (auto intro: Max_gr)
+      from this S(2) have "?max_subseq R i < ?max_subseq R j" by (auto intro!: Max_gr_iff [THEN iffD2])
     } note max_subseq_increase = this
     have "?max_subseq (\<le>) i < ?max_subseq (\<le>) j \<or> ?max_subseq (\<ge>) i < ?max_subseq (\<ge>) j"
     proof (cases "f j \<ge> f i")