author | wenzelm |
Fri, 21 Apr 2017 21:36:49 +0200 | |
changeset 65548 | b7caa2b8bdbf |
parent 63040 | eb4ddd18d635 |
child 67399 | eab6ce8368fa |
permissions | -rw-r--r-- |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
1 |
(* Title: HOL/ex/Erdoes_Szekeres.thy |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
2 |
Author: Lukas Bulwahn <lukas.bulwahn-at-gmail.com> |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
3 |
*) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
4 |
|
61343 | 5 |
section \<open>The Erdoes-Szekeres Theorem\<close> |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
6 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
7 |
theory Erdoes_Szekeres |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
8 |
imports Main |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
9 |
begin |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
10 |
|
61343 | 11 |
subsection \<open>Addition to @{theory Lattices_Big} Theory\<close> |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
12 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
13 |
lemma Max_gr: |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
14 |
assumes "finite A" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
15 |
assumes "a \<in> A" "a > x" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
16 |
shows "x < Max A" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
17 |
using assms Max_ge less_le_trans by blast |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
18 |
|
61343 | 19 |
subsection \<open>Additions to @{theory Finite_Set} Theory\<close> |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
20 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
21 |
lemma obtain_subset_with_card_n: |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
22 |
assumes "n \<le> card S" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
23 |
obtains T where "T \<subseteq> S" "card T = n" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
24 |
proof - |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
25 |
from assms obtain n' where "card S = n + n'" by (metis le_add_diff_inverse) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
26 |
from this that show ?thesis |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
27 |
proof (induct n' arbitrary: S) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
28 |
case 0 from this show ?case by auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
29 |
next |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
30 |
case Suc from this show ?case by (simp add: card_Suc_eq) (metis subset_insertI2) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
31 |
qed |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
32 |
qed |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
33 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
34 |
lemma exists_set_with_max_card: |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
35 |
assumes "finite S" "S \<noteq> {}" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
36 |
shows "\<exists>s \<in> S. card s = Max (card ` S)" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
37 |
using assms |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
38 |
proof (induct S rule: finite.induct) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
39 |
case (insertI S' s') |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
40 |
show ?case |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
41 |
proof (cases "S' \<noteq> {}") |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
42 |
case True |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
43 |
from this insertI.hyps(2) obtain s where s: "s \<in> S'" "card s = Max (card ` S')" by auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
44 |
from this(1) have that: "(if card s \<ge> card s' then s else s') \<in> insert s' S'" by auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
45 |
have "card (if card s \<ge> card s' then s else s') = Max (card ` insert s' S')" |
61343 | 46 |
using insertI(1) \<open>S' \<noteq> {}\<close> s by auto |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
47 |
from this that show ?thesis by blast |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
48 |
qed (auto) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
49 |
qed (auto) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
50 |
|
61343 | 51 |
subsection \<open>Definition of Monotonicity over a Carrier Set\<close> |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
52 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
53 |
definition |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
54 |
"mono_on f R S = (\<forall>i\<in>S. \<forall>j\<in>S. i \<le> j \<longrightarrow> R (f i) (f j))" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
55 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
56 |
lemma mono_on_empty [simp]: "mono_on f R {}" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
57 |
unfolding mono_on_def by auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
58 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
59 |
lemma mono_on_singleton [simp]: "reflp R \<Longrightarrow> mono_on f R {x}" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
60 |
unfolding mono_on_def reflp_def by auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
61 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
62 |
lemma mono_on_subset: "T \<subseteq> S \<Longrightarrow> mono_on f R S \<Longrightarrow> mono_on f R T" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
63 |
unfolding mono_on_def by (simp add: subset_iff) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
64 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
65 |
lemma not_mono_on_subset: "T \<subseteq> S \<Longrightarrow> \<not> mono_on f R T \<Longrightarrow> \<not> mono_on f R S" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
66 |
unfolding mono_on_def by blast |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
67 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
68 |
lemma [simp]: |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
69 |
"reflp (op \<le> :: 'a::order \<Rightarrow> _ \<Rightarrow> bool)" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
70 |
"reflp (op \<ge> :: 'a::order \<Rightarrow> _ \<Rightarrow> bool)" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
71 |
"transp (op \<le> :: 'a::order \<Rightarrow> _ \<Rightarrow> bool)" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
72 |
"transp (op \<ge> :: 'a::order \<Rightarrow> _ \<Rightarrow> bool)" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
73 |
unfolding reflp_def transp_def by auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
74 |
|
61343 | 75 |
subsection \<open>The Erdoes-Szekeres Theorem following Seidenberg's (1959) argument\<close> |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
76 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
77 |
lemma Erdoes_Szekeres: |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
78 |
fixes f :: "_ \<Rightarrow> 'a::linorder" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
79 |
shows "(\<exists>S. S \<subseteq> {0..m * n} \<and> card S = m + 1 \<and> mono_on f (op \<le>) S) \<or> |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
80 |
(\<exists>S. S \<subseteq> {0..m * n} \<and> card S = n + 1 \<and> mono_on f (op \<ge>) S)" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
81 |
proof (rule ccontr) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
82 |
let ?max_subseq = "\<lambda>R k. Max (card ` {S. S \<subseteq> {0..k} \<and> mono_on f R S \<and> k \<in> S})" |
63040 | 83 |
define phi where "phi k = (?max_subseq (op \<le>) k, ?max_subseq (op \<ge>) k)" for k |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
84 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
85 |
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 |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
86 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
87 |
{ |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
88 |
fix R |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
89 |
assume reflp: "reflp (R :: 'a::linorder \<Rightarrow> _)" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
90 |
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 |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
91 |
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 |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
92 |
from this have lower_bound: "\<And>k. k \<le> m * n \<Longrightarrow> ?max_subseq R k \<ge> 1" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
93 |
by (auto intro!: Max_ge) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
94 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
95 |
fix b |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
96 |
assume not_mono_at: "\<forall>S. S \<subseteq> {0..m * n} \<and> card S = b + 1 \<longrightarrow> \<not> mono_on f R S" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
97 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
98 |
{ |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
99 |
fix S |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
100 |
assume "S \<subseteq> {0..m * n}" "card S \<ge> b + 1" |
61343 | 101 |
moreover from \<open>card S \<ge> b + 1\<close> obtain T where "T \<subseteq> S \<and> card T = Suc b" |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
102 |
using obtain_subset_with_card_n by (metis Suc_eq_plus1) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
103 |
ultimately have "\<not> mono_on f R S" using not_mono_at by (auto dest: not_mono_on_subset) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
104 |
} |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
105 |
from this have "\<forall>S. S \<subseteq> {0..m * n} \<and> mono_on f R S \<longrightarrow> card S \<le> b" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
106 |
by (metis Suc_eq_plus1 Suc_leI not_le) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
107 |
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" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
108 |
using order_trans by force |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
109 |
from this non_empty have upper_bound: "\<And>k. k \<le> m * n \<Longrightarrow> ?max_subseq R k \<le> b" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
110 |
by (auto intro: Max.boundedI) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
111 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
112 |
from upper_bound lower_bound have "\<And>k. k \<le> m * n \<Longrightarrow> 1 \<le> ?max_subseq R k \<and> ?max_subseq R k \<le> b" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
113 |
by auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
114 |
} note bounds = this |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
115 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
116 |
assume contraposition: "\<not> ?thesis" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
117 |
from contraposition bounds[of "op \<le>" "m"] bounds[of "op \<ge>" "n"] |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
118 |
have "\<And>k. k \<le> m * n \<Longrightarrow> 1 \<le> ?max_subseq (op \<le>) k \<and> ?max_subseq (op \<le>) k \<le> m" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
119 |
and "\<And>k. k \<le> m * n \<Longrightarrow> 1 \<le> ?max_subseq (op \<ge>) k \<and> ?max_subseq (op \<ge>) k \<le> n" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
120 |
using reflp_def by simp+ |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
121 |
from this have "\<forall>i \<in> {0..m * n}. phi i \<in> {1..m} \<times> {1..n}" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
122 |
unfolding phi_def by auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
123 |
from this have subseteq: "phi ` {0..m * n} \<subseteq> {1..m} \<times> {1..n}" by blast |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
124 |
have card_product: "card ({1..m} \<times> {1..n}) = m * n" by (simp add: card_cartesian_product) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
125 |
have "finite ({1..m} \<times> {1..n})" by blast |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
126 |
from subseteq card_product this have card_le: "card (phi ` {0..m * n}) \<le> m * n" by (metis card_mono) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
127 |
|
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
128 |
{ |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
129 |
fix i j |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
130 |
assume "i < (j :: nat)" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
131 |
{ |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
132 |
fix R |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
133 |
assume R: "reflp (R :: 'a::linorder \<Rightarrow> _)" "transp R" "R (f i) (f j)" |
61343 | 134 |
from one_member[OF \<open>reflp R\<close>, of "i"] have |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
135 |
"\<exists>S \<in> {S. S \<subseteq> {0..i} \<and> mono_on f R S \<and> i \<in> S}. card S = ?max_subseq R i" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
136 |
by (intro exists_set_with_max_card) auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
137 |
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 |
61343 | 138 |
from S \<open>i < j\<close> finite_subset have "j \<notin> S" "finite S" "insert j S \<subseteq> {0..j}" by auto |
139 |
from S(1) R \<open>i < j\<close> this have "mono_on f R (insert j S)" |
|
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
140 |
unfolding mono_on_def reflp_def transp_def |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
141 |
by (metis atLeastAtMost_iff insert_iff le_antisym subsetCE) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
142 |
from this have d: "insert j S \<in> {S. S \<subseteq> {0..j} \<and> mono_on f R S \<and> j \<in> S}" |
61343 | 143 |
using \<open>insert j S \<subseteq> {0..j}\<close> by blast |
144 |
from this \<open>j \<notin> S\<close> S(1) have "card (insert j S) \<in> |
|
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
145 |
card ` {S. S \<subseteq> {0..j} \<and> mono_on f R S \<and> j \<in> S} \<and> card S < card (insert j S)" |
61343 | 146 |
by (auto intro!: imageI) (auto simp add: \<open>finite S\<close>) |
60603
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
147 |
from this S(2) have "?max_subseq R i < ?max_subseq R j" by (auto intro: Max_gr) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
148 |
} note max_subseq_increase = this |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
149 |
have "?max_subseq (op \<le>) i < ?max_subseq (op \<le>) j \<or> ?max_subseq (op \<ge>) i < ?max_subseq (op \<ge>) j" |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
150 |
proof (cases "f j \<ge> f i") |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
151 |
case True |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
152 |
from this max_subseq_increase[of "op \<le>", simplified] show ?thesis by simp |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
153 |
next |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
154 |
case False |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
155 |
from this max_subseq_increase[of "op \<ge>", simplified] show ?thesis by simp |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
156 |
qed |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
157 |
from this have "phi i \<noteq> phi j" using phi_def by auto |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
158 |
} |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
159 |
from this have "inj phi" unfolding inj_on_def by (metis less_linear) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
160 |
from this have card_eq: "card (phi ` {0..m * n}) = m * n + 1" by (simp add: card_image inj_on_def) |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
161 |
from card_le card_eq show False by simp |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
162 |
qed |
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff
changeset
|
163 |
|
62390 | 164 |
end |