57 "\<lbrakk>\<forall>x k. P x = P (x - k*D); \<forall>x k. Q x = Q (x - k*D)\<rbrakk> |
57 "\<lbrakk>\<forall>x k. P x = P (x - k*D); \<forall>x k. Q x = Q (x - k*D)\<rbrakk> |
58 \<Longrightarrow> \<forall>x k. (P x \<or> Q x) = (P (x - k*D) \<or> Q (x - k*D))" |
58 \<Longrightarrow> \<forall>x k. (P x \<or> Q x) = (P (x - k*D) \<or> Q (x - k*D))" |
59 "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \<Longrightarrow> \<forall>x k. (d dvd x + t) = (d dvd (x - k*D) + t)" |
59 "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \<Longrightarrow> \<forall>x k. (d dvd x + t) = (d dvd (x - k*D) + t)" |
60 "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \<Longrightarrow> \<forall>x k. (\<not>d dvd x + t) = (\<not>d dvd (x - k*D) + t)" |
60 "(d::'a::{comm_ring,Ring_and_Field.dvd}) dvd D \<Longrightarrow> \<forall>x k. (\<not>d dvd x + t) = (\<not>d dvd (x - k*D) + t)" |
61 "\<forall>x k. F = F" |
61 "\<forall>x k. F = F" |
62 apply (auto elim!: dvdE simp add: ring_simps) |
62 apply (auto elim!: dvdE simp add: algebra_simps) |
63 unfolding mult_assoc [symmetric] left_distrib [symmetric] left_diff_distrib [symmetric] |
63 unfolding mult_assoc [symmetric] left_distrib [symmetric] left_diff_distrib [symmetric] |
64 unfolding dvd_def mult_commute [of d] |
64 unfolding dvd_def mult_commute [of d] |
65 by auto |
65 by auto |
66 |
66 |
67 subsection{* The A and B sets *} |
67 subsection{* The A and B sets *} |
99 next |
99 next |
100 assume dp: "D > 0" and tB:"t \<in> B" |
100 assume dp: "D > 0" and tB:"t \<in> B" |
101 {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j" and g: "x > t" and ng: "\<not> (x - D) > t" |
101 {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j" and g: "x > t" and ng: "\<not> (x - D) > t" |
102 hence "x -t \<le> D" and "1 \<le> x - t" by simp+ |
102 hence "x -t \<le> D" and "1 \<le> x - t" by simp+ |
103 hence "\<exists>j \<in> {1 .. D}. x - t = j" by auto |
103 hence "\<exists>j \<in> {1 .. D}. x - t = j" by auto |
104 hence "\<exists>j \<in> {1 .. D}. x = t + j" by (simp add: ring_simps) |
104 hence "\<exists>j \<in> {1 .. D}. x = t + j" by (simp add: algebra_simps) |
105 with nob tB have "False" by simp} |
105 with nob tB have "False" by simp} |
106 thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (x > t) \<longrightarrow> (x - D > t)" by blast |
106 thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (x > t) \<longrightarrow> (x - D > t)" by blast |
107 next |
107 next |
108 assume dp: "D > 0" and tB:"t - 1\<in> B" |
108 assume dp: "D > 0" and tB:"t - 1\<in> B" |
109 {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j" and g: "x \<ge> t" and ng: "\<not> (x - D) \<ge> t" |
109 {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j" and g: "x \<ge> t" and ng: "\<not> (x - D) \<ge> t" |
110 hence "x - (t - 1) \<le> D" and "1 \<le> x - (t - 1)" by simp+ |
110 hence "x - (t - 1) \<le> D" and "1 \<le> x - (t - 1)" by simp+ |
111 hence "\<exists>j \<in> {1 .. D}. x - (t - 1) = j" by auto |
111 hence "\<exists>j \<in> {1 .. D}. x - (t - 1) = j" by auto |
112 hence "\<exists>j \<in> {1 .. D}. x = (t - 1) + j" by (simp add: ring_simps) |
112 hence "\<exists>j \<in> {1 .. D}. x = (t - 1) + j" by (simp add: algebra_simps) |
113 with nob tB have "False" by simp} |
113 with nob tB have "False" by simp} |
114 thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (x \<ge> t) \<longrightarrow> (x - D \<ge> t)" by blast |
114 thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (x \<ge> t) \<longrightarrow> (x - D \<ge> t)" by blast |
115 next |
115 next |
116 assume d: "d dvd D" |
116 assume d: "d dvd D" |
117 {fix x assume H: "d dvd x + t" with d have "d dvd (x - D) + t" by algebra} |
117 {fix x assume H: "d dvd x + t" with d have "d dvd (x - D) + t" by algebra} |
118 thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (d dvd x+t) \<longrightarrow> (d dvd (x - D) + t)" by simp |
118 thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (d dvd x+t) \<longrightarrow> (d dvd (x - D) + t)" by simp |
119 next |
119 next |
120 assume d: "d dvd D" |
120 assume d: "d dvd D" |
121 {fix x assume H: "\<not>(d dvd x + t)" with d have "\<not> d dvd (x - D) + t" |
121 {fix x assume H: "\<not>(d dvd x + t)" with d have "\<not> d dvd (x - D) + t" |
122 by (clarsimp simp add: dvd_def,erule_tac x= "ka + k" in allE,simp add: ring_simps)} |
122 by (clarsimp simp add: dvd_def,erule_tac x= "ka + k" in allE,simp add: algebra_simps)} |
123 thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (\<not>d dvd x+t) \<longrightarrow> (\<not>d dvd (x - D) + t)" by auto |
123 thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>B. x \<noteq> b + j)\<longrightarrow> (\<not>d dvd x+t) \<longrightarrow> (\<not>d dvd (x - D) + t)" by auto |
124 qed blast |
124 qed blast |
125 |
125 |
126 lemma aset: |
126 lemma aset: |
127 "\<lbrakk>\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> P x \<longrightarrow> P(x + D) ; |
127 "\<lbrakk>\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> P x \<longrightarrow> P(x + D) ; |
156 next |
156 next |
157 assume dp: "D > 0" and tA:"t \<in> A" |
157 assume dp: "D > 0" and tA:"t \<in> A" |
158 {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j" and g: "x < t" and ng: "\<not> (x + D) < t" |
158 {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j" and g: "x < t" and ng: "\<not> (x + D) < t" |
159 hence "t - x \<le> D" and "1 \<le> t - x" by simp+ |
159 hence "t - x \<le> D" and "1 \<le> t - x" by simp+ |
160 hence "\<exists>j \<in> {1 .. D}. t - x = j" by auto |
160 hence "\<exists>j \<in> {1 .. D}. t - x = j" by auto |
161 hence "\<exists>j \<in> {1 .. D}. x = t - j" by (auto simp add: ring_simps) |
161 hence "\<exists>j \<in> {1 .. D}. x = t - j" by (auto simp add: algebra_simps) |
162 with nob tA have "False" by simp} |
162 with nob tA have "False" by simp} |
163 thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (x < t) \<longrightarrow> (x + D < t)" by blast |
163 thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (x < t) \<longrightarrow> (x + D < t)" by blast |
164 next |
164 next |
165 assume dp: "D > 0" and tA:"t + 1\<in> A" |
165 assume dp: "D > 0" and tA:"t + 1\<in> A" |
166 {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j" and g: "x \<le> t" and ng: "\<not> (x + D) \<le> t" |
166 {fix x assume nob: "\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j" and g: "x \<le> t" and ng: "\<not> (x + D) \<le> t" |
167 hence "(t + 1) - x \<le> D" and "1 \<le> (t + 1) - x" by (simp_all add: ring_simps) |
167 hence "(t + 1) - x \<le> D" and "1 \<le> (t + 1) - x" by (simp_all add: algebra_simps) |
168 hence "\<exists>j \<in> {1 .. D}. (t + 1) - x = j" by auto |
168 hence "\<exists>j \<in> {1 .. D}. (t + 1) - x = j" by auto |
169 hence "\<exists>j \<in> {1 .. D}. x = (t + 1) - j" by (auto simp add: ring_simps) |
169 hence "\<exists>j \<in> {1 .. D}. x = (t + 1) - j" by (auto simp add: algebra_simps) |
170 with nob tA have "False" by simp} |
170 with nob tA have "False" by simp} |
171 thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (x \<le> t) \<longrightarrow> (x + D \<le> t)" by blast |
171 thus "\<forall>x.(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (x \<le> t) \<longrightarrow> (x + D \<le> t)" by blast |
172 next |
172 next |
173 assume d: "d dvd D" |
173 assume d: "d dvd D" |
174 {fix x assume H: "d dvd x + t" with d have "d dvd (x + D) + t" |
174 {fix x assume H: "d dvd x + t" with d have "d dvd (x + D) + t" |
175 by (clarsimp simp add: dvd_def,rule_tac x= "ka + k" in exI,simp add: ring_simps)} |
175 by (clarsimp simp add: dvd_def,rule_tac x= "ka + k" in exI,simp add: algebra_simps)} |
176 thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (d dvd x+t) \<longrightarrow> (d dvd (x + D) + t)" by simp |
176 thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (d dvd x+t) \<longrightarrow> (d dvd (x + D) + t)" by simp |
177 next |
177 next |
178 assume d: "d dvd D" |
178 assume d: "d dvd D" |
179 {fix x assume H: "\<not>(d dvd x + t)" with d have "\<not>d dvd (x + D) + t" |
179 {fix x assume H: "\<not>(d dvd x + t)" with d have "\<not>d dvd (x + D) + t" |
180 by (clarsimp simp add: dvd_def,erule_tac x= "ka - k" in allE,simp add: ring_simps)} |
180 by (clarsimp simp add: dvd_def,erule_tac x= "ka - k" in allE,simp add: algebra_simps)} |
181 thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (\<not>d dvd x+t) \<longrightarrow> (\<not>d dvd (x + D) + t)" by auto |
181 thus "\<forall>(x::int).(\<forall>j\<in>{1 .. D}. \<forall>b\<in>A. x \<noteq> b - j)\<longrightarrow> (\<not>d dvd x+t) \<longrightarrow> (\<not>d dvd (x + D) + t)" by auto |
182 qed blast |
182 qed blast |
183 |
183 |
184 subsection{* Cooper's Theorem @{text "-\<infinity>"} and @{text "+\<infinity>"} Version *} |
184 subsection{* Cooper's Theorem @{text "-\<infinity>"} and @{text "+\<infinity>"} Version *} |
185 |
185 |
302 assume eP1: "EX x. P' x" |
302 assume eP1: "EX x. P' x" |
303 then obtain x where P1: "P' x" .. |
303 then obtain x where P1: "P' x" .. |
304 from ePeqP1 obtain z where P1eqP: "\<forall>x>z. P x = P' x" .. |
304 from ePeqP1 obtain z where P1eqP: "\<forall>x>z. P x = P' x" .. |
305 let ?w' = "x + (abs(x-z)+1) * d" |
305 let ?w' = "x + (abs(x-z)+1) * d" |
306 let ?w = "x - (-(abs(x-z) + 1))*d" |
306 let ?w = "x - (-(abs(x-z) + 1))*d" |
307 have ww'[simp]: "?w = ?w'" by (simp add: ring_simps) |
307 have ww'[simp]: "?w = ?w'" by (simp add: algebra_simps) |
308 from dpos have w: "?w > z" by(simp only: ww' incr_lemma) |
308 from dpos have w: "?w > z" by(simp only: ww' incr_lemma) |
309 hence "P' x = P' ?w" using P1eqP1 by blast |
309 hence "P' x = P' ?w" using P1eqP1 by blast |
310 also have "\<dots> = P(?w)" using w P1eqP by blast |
310 also have "\<dots> = P(?w)" using w P1eqP by blast |
311 finally have "P ?w" using P1 by blast |
311 finally have "P ?w" using P1 by blast |
312 thus "EX x. P x" .. |
312 thus "EX x. P x" .. |