1section \<open>The Reals as Dedekind Sections of Positive Rationals\<close>
2
3text \<open>Fundamentals of Abstract Analysis [Gleason- p. 121] provides some of the definitions.\<close>
4
5(*  Title:      HOL/ex/Dedekind_Real.thy
6    Author:     Jacques D. Fleuriot, University of Cambridge
7    Conversion to Isar and new proofs by Lawrence C Paulson, 2003/4; 2019
8*)
9
10theory Dedekind_Real
11imports Complex_Main 
12begin
13
14text\<open>Could be moved to \<open>Groups\<close>\<close>
15lemma add_eq_exists: "\<exists>x. a+x = (b::'a::ab_group_add)"
16  by (rule_tac x="b-a" in exI, simp)
17
18subsection \<open>Dedekind cuts or sections\<close>
19
20definition
21  cut :: "rat set \<Rightarrow> bool" where
22  "cut A \<equiv> {} \<subset> A \<and> A \<subset> {0<..} \<and>
23            (\<forall>y \<in> A. ((\<forall>z. 0<z \<and> z < y \<longrightarrow> z \<in> A) \<and> (\<exists>u \<in> A. y < u)))"
24
25lemma cut_of_rat: 
26  assumes q: "0 < q" shows "cut {r::rat. 0 < r \<and> r < q}" (is "cut ?A")
27proof -
28  from q have pos: "?A \<subset> {0<..}" by force
29  have nonempty: "{} \<subset> ?A"
30  proof
31    show "{} \<subseteq> ?A" by simp
32    show "{} \<noteq> ?A"
33      using field_lbound_gt_zero q by auto
34  qed
35  show ?thesis
36    by (simp add: cut_def pos nonempty,
37        blast dest: dense intro: order_less_trans)
38qed
39
40
41typedef preal = "Collect cut"
42  by (blast intro: cut_of_rat [OF zero_less_one])
43
44lemma Abs_preal_induct [induct type: preal]:
45  "(\<And>x. cut x \<Longrightarrow> P (Abs_preal x)) \<Longrightarrow> P x"
46  using Abs_preal_induct [of P x] by simp
47
48lemma cut_Rep_preal [simp]: "cut (Rep_preal x)"
49  using Rep_preal [of x] by simp
50
51definition
52  psup :: "preal set \<Rightarrow> preal" where
53  "psup P = Abs_preal (\<Union>X \<in> P. Rep_preal X)"
54
55definition
56  add_set :: "[rat set,rat set] \<Rightarrow> rat set" where
57  "add_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x + y}"
58
59definition
60  diff_set :: "[rat set,rat set] \<Rightarrow> rat set" where
61  "diff_set A B = {w. \<exists>x. 0 < w \<and> 0 < x \<and> x \<notin> B \<and> x + w \<in> A}"
62
63definition
64  mult_set :: "[rat set,rat set] \<Rightarrow> rat set" where
65  "mult_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x * y}"
66
67definition
68  inverse_set :: "rat set \<Rightarrow> rat set" where
69  "inverse_set A \<equiv> {x. \<exists>y. 0 < x \<and> x < y \<and> inverse y \<notin> A}"
70
71instantiation preal :: "{ord, plus, minus, times, inverse, one}"
72begin
73
74definition
75  preal_less_def:
76    "r < s \<equiv> Rep_preal r < Rep_preal s"
77
78definition
79  preal_le_def:
80    "r \<le> s \<equiv> Rep_preal r \<subseteq> Rep_preal s"
81
82definition
83  preal_add_def:
84    "r + s \<equiv> Abs_preal (add_set (Rep_preal r) (Rep_preal s))"
85
86definition
87  preal_diff_def:
88    "r - s \<equiv> Abs_preal (diff_set (Rep_preal r) (Rep_preal s))"
89
90definition
91  preal_mult_def:
92    "r * s \<equiv> Abs_preal (mult_set (Rep_preal r) (Rep_preal s))"
93
94definition
95  preal_inverse_def:
96    "inverse r \<equiv> Abs_preal (inverse_set (Rep_preal r))"
97
98definition "r div s = r * inverse (s::preal)"
99
100definition
101  preal_one_def:
102    "1 \<equiv> Abs_preal {x. 0 < x \<and> x < 1}"
103
104instance ..
105
106end
107
108
109text\<open>Reduces equality on abstractions to equality on representatives\<close>
110declare Abs_preal_inject [simp]
111declare Abs_preal_inverse [simp]
112
113lemma rat_mem_preal: "0 < q \<Longrightarrow> cut {r::rat. 0 < r \<and> r < q}"
114by (simp add: cut_of_rat)
115
116lemma preal_nonempty: "cut A \<Longrightarrow> \<exists>x\<in>A. 0 < x"
117  unfolding cut_def [abs_def] by blast
118
119lemma preal_Ex_mem: "cut A \<Longrightarrow> \<exists>x. x \<in> A"
120  using preal_nonempty by blast
121
122lemma preal_exists_bound: "cut A \<Longrightarrow> \<exists>x. 0 < x \<and> x \<notin> A"
123  using Dedekind_Real.cut_def by fastforce
124
125lemma preal_exists_greater: "\<lbrakk>cut A; y \<in> A\<rbrakk> \<Longrightarrow> \<exists>u \<in> A. y < u"
126  unfolding cut_def [abs_def] by blast
127
128lemma preal_downwards_closed: "\<lbrakk>cut A; y \<in> A; 0 < z; z < y\<rbrakk> \<Longrightarrow> z \<in> A"
129  unfolding cut_def [abs_def] by blast
130
131text\<open>Relaxing the final premise\<close>
132lemma preal_downwards_closed': "\<lbrakk>cut A; y \<in> A; 0 < z; z \<le> y\<rbrakk> \<Longrightarrow> z \<in> A"
133  using less_eq_rat_def preal_downwards_closed by blast
134
135text\<open>A positive fraction not in a positive real is an upper bound.
136 Gleason p. 122 - Remark (1)\<close>
137
138lemma not_in_preal_ub:
139  assumes A: "cut A"
140    and notx: "x \<notin> A"
141    and y: "y \<in> A"
142    and pos: "0 < x"
143  shows "y < x"
144proof (cases rule: linorder_cases)
145  assume "x<y"
146  with notx show ?thesis
147    by (simp add:  preal_downwards_closed [OF A y] pos)
148next
149  assume "x=y"
150  with notx and y show ?thesis by simp
151next
152  assume "y<x"
153  thus ?thesis .
154qed
155
156text \<open>preal lemmas instantiated to \<^term>\<open>Rep_preal X\<close>\<close>
157
158lemma mem_Rep_preal_Ex: "\<exists>x. x \<in> Rep_preal X"
159thm preal_Ex_mem
160by (rule preal_Ex_mem [OF cut_Rep_preal])
161
162lemma Rep_preal_exists_bound: "\<exists>x>0. x \<notin> Rep_preal X"
163by (rule preal_exists_bound [OF cut_Rep_preal])
164
165lemmas not_in_Rep_preal_ub = not_in_preal_ub [OF cut_Rep_preal]
166
167
168subsection\<open>Properties of Ordering\<close>
169
170instance preal :: order
171proof
172  fix w :: preal
173  show "w \<le> w" by (simp add: preal_le_def)
174next
175  fix i j k :: preal
176  assume "i \<le> j" and "j \<le> k"
177  then show "i \<le> k" by (simp add: preal_le_def)
178next
179  fix z w :: preal
180  assume "z \<le> w" and "w \<le> z"
181  then show "z = w" by (simp add: preal_le_def Rep_preal_inject)
182next
183  fix z w :: preal
184  show "z < w \<longleftrightarrow> z \<le> w \<and> \<not> w \<le> z"
185  by (auto simp: preal_le_def preal_less_def Rep_preal_inject)
186qed  
187
188lemma preal_imp_pos: "\<lbrakk>cut A; r \<in> A\<rbrakk> \<Longrightarrow> 0 < r"
189  by (auto simp: cut_def)
190
191instance preal :: linorder
192proof
193  fix x y :: preal
194  show "x \<le> y \<or> y \<le> x"
195    unfolding preal_le_def
196    by (meson cut_Rep_preal not_in_preal_ub preal_downwards_closed preal_imp_pos subsetI)
197qed
198
199instantiation preal :: distrib_lattice
200begin
201
202definition
203  "(inf :: preal \<Rightarrow> preal \<Rightarrow> preal) = min"
204
205definition
206  "(sup :: preal \<Rightarrow> preal \<Rightarrow> preal) = max"
207
208instance
209  by intro_classes
210    (auto simp: inf_preal_def sup_preal_def max_min_distrib2)
211
212end
213
214subsection\<open>Properties of Addition\<close>
215
216lemma preal_add_commute: "(x::preal) + y = y + x"
217  unfolding preal_add_def add_set_def
218  by (metis (no_types, hide_lams) add.commute)
219
220text\<open>Lemmas for proving that addition of two positive reals gives
221 a positive real\<close>
222
223lemma mem_add_set:
224  assumes "cut A" "cut B"
225  shows "cut (add_set A B)"
226proof -
227  have "{} \<subset> add_set A B"
228    using assms by (force simp: add_set_def dest: preal_nonempty)
229  moreover
230  obtain q where "q > 0" "q \<notin> add_set A B"
231  proof -
232    obtain a b where "a > 0" "a \<notin> A" "b > 0" "b \<notin> B" "\<And>x. x \<in> A \<Longrightarrow> x < a" "\<And>y. y \<in> B \<Longrightarrow> y < b"
233      by (meson assms preal_exists_bound not_in_preal_ub)
234    with assms have "a+b \<notin> add_set A B"
235      by (fastforce simp add: add_set_def)
236    then show thesis
237      using \<open>0 < a\<close> \<open>0 < b\<close> add_pos_pos that by blast
238  qed
239  then have "add_set A B \<subset> {0<..}"
240    unfolding add_set_def
241    using preal_imp_pos [OF \<open>cut A\<close>] preal_imp_pos [OF \<open>cut B\<close>]  by fastforce
242  moreover have "z \<in> add_set A B" 
243    if u: "u \<in> add_set A B" and "0 < z" "z < u" for u z
244    using u unfolding add_set_def
245  proof (clarify)
246    fix x::rat and y::rat
247    assume ueq: "u = x + y" and x: "x \<in> A" and y:"y \<in> B"
248    have xpos [simp]: "x > 0" and ypos [simp]: "y > 0"
249      using assms preal_imp_pos x y by blast+
250    have xypos [simp]: "x+y > 0" by (simp add: pos_add_strict)
251    let ?f = "z/(x+y)"
252    have fless: "?f < 1"
253      using divide_less_eq_1_pos \<open>z < u\<close> ueq xypos by blast
254    show "\<exists>x' \<in> A. \<exists>y'\<in>B. z = x' + y'"
255    proof (intro bexI)
256      show "z = x*?f + y*?f"
257        by (simp add: distrib_right [symmetric] divide_inverse ac_simps order_less_imp_not_eq2)
258    next
259      show "y * ?f \<in> B"
260      proof (rule preal_downwards_closed [OF \<open>cut B\<close> y])
261        show "0 < y * ?f"
262          by (simp add: \<open>0 < z\<close>)
263      next
264        show "y * ?f < y"
265          by (insert mult_strict_left_mono [OF fless ypos], simp)
266      qed
267    next
268      show "x * ?f \<in> A"
269      proof (rule preal_downwards_closed [OF \<open>cut A\<close> x])
270        show "0 < x * ?f"
271          by (simp add: \<open>0 < z\<close>)
272      next
273        show "x * ?f < x"
274          by (insert mult_strict_left_mono [OF fless xpos], simp)
275      qed
276    qed
277  qed
278  moreover
279  have "\<And>y. y \<in> add_set A B \<Longrightarrow> \<exists>u \<in> add_set A B. y < u"
280    unfolding add_set_def using preal_exists_greater assms by fastforce
281  ultimately show ?thesis
282    by (simp add: Dedekind_Real.cut_def)
283qed
284
285lemma preal_add_assoc: "((x::preal) + y) + z = x + (y + z)"
286  apply (simp add: preal_add_def mem_add_set)
287  apply (force simp: add_set_def ac_simps)
288  done
289
290instance preal :: ab_semigroup_add
291proof
292  fix a b c :: preal
293  show "(a + b) + c = a + (b + c)" by (rule preal_add_assoc)
294  show "a + b = b + a" by (rule preal_add_commute)
295qed
296
297
298subsection\<open>Properties of Multiplication\<close>
299
300text\<open>Proofs essentially same as for addition\<close>
301
302lemma preal_mult_commute: "(x::preal) * y = y * x"
303  unfolding preal_mult_def mult_set_def
304  by (metis (no_types, hide_lams) mult.commute)
305
306text\<open>Multiplication of two positive reals gives a positive real.\<close>
307
308lemma mem_mult_set:
309  assumes "cut A" "cut B"
310  shows "cut (mult_set A B)"
311proof -
312  have "{} \<subset> mult_set A B"
313    using assms
314      by (force simp: mult_set_def dest: preal_nonempty)
315    moreover
316    obtain q where "q > 0" "q \<notin> mult_set A B"
317    proof -
318      obtain x y where x [simp]: "0 < x" "x \<notin> A" and y [simp]: "0 < y" "y \<notin> B"
319        using preal_exists_bound assms by blast
320      show thesis
321      proof 
322        show "0 < x*y" by simp
323        show "x * y \<notin> mult_set A B"
324        proof -
325          {
326            fix u::rat and v::rat
327            assume u: "u \<in> A" and v: "v \<in> B" and xy: "x*y = u*v"
328            moreover have "u<x" and "v<y" using assms x y u v by (blast dest: not_in_preal_ub)+
329            moreover have "0\<le>v"
330              using less_imp_le preal_imp_pos assms x y u v by blast
331            moreover have "u*v < x*y"
332                using assms x \<open>u < x\<close> \<open>v < y\<close> \<open>0 \<le> v\<close> by (blast intro: mult_strict_mono)
333            ultimately have False by force
334          }
335          thus ?thesis by (auto simp: mult_set_def)
336        qed
337      qed
338    qed
339  then have "mult_set A B \<subset> {0<..}"
340    unfolding mult_set_def
341    using preal_imp_pos [OF \<open>cut A\<close>] preal_imp_pos [OF \<open>cut B\<close>]  by fastforce
342  moreover have "z \<in> mult_set A B"
343    if u: "u \<in> mult_set A B" and "0 < z" "z < u" for u z
344    using u unfolding mult_set_def
345  proof (clarify)
346    fix x::rat and y::rat
347    assume ueq: "u = x * y" and x: "x \<in> A" and y: "y \<in> B"  
348    have [simp]: "y > 0"
349      using \<open>cut B\<close> preal_imp_pos y by blast
350    show "\<exists>x' \<in> A. \<exists>y' \<in> B. z = x' * y'"
351    proof
352      have "z = (z/y)*y"
353          by (simp add: divide_inverse mult.commute [of y] mult.assoc order_less_imp_not_eq2)
354      then show "\<exists>y'\<in>B. z = (z/y) * y'"
355        using y by blast
356    next
357      show "z/y \<in> A"
358      proof (rule preal_downwards_closed [OF \<open>cut A\<close> x])
359        show "0 < z/y"
360          by (simp add: \<open>0 < z\<close>)
361        show "z/y < x"
362          using \<open>0 < y\<close> pos_divide_less_eq \<open>z < u\<close> ueq by blast  
363      qed
364    qed
365  qed
366  moreover have "\<And>y. y \<in> mult_set A B \<Longrightarrow> \<exists>u \<in> mult_set A B. y < u"
367    apply (simp add: mult_set_def)
368    by (metis preal_exists_greater mult_strict_right_mono preal_imp_pos assms)
369  ultimately show ?thesis
370    by (simp add: Dedekind_Real.cut_def)
371qed
372
373lemma preal_mult_assoc: "((x::preal) * y) * z = x * (y * z)"
374  apply (simp add: preal_mult_def mem_mult_set Rep_preal)
375  apply (force simp: mult_set_def ac_simps)
376  done
377
378instance preal :: ab_semigroup_mult
379proof
380  fix a b c :: preal
381  show "(a * b) * c = a * (b * c)" by (rule preal_mult_assoc)
382  show "a * b = b * a" by (rule preal_mult_commute)
383qed
384
385
386text\<open>Positive real 1 is the multiplicative identity element\<close>
387
388lemma preal_mult_1: "(1::preal) * z = z"
389proof (induct z)
390  fix A :: "rat set"
391  assume A: "cut A"
392  have "{w. \<exists>u. 0 < u \<and> u < 1 \<and> (\<exists>v \<in> A. w = u * v)} = A" (is "?lhs = A")
393  proof
394    show "?lhs \<subseteq> A"
395    proof clarify
396      fix x::rat and u::rat and v::rat
397      assume upos: "0<u" and "u<1" and v: "v \<in> A"
398      have vpos: "0<v" by (rule preal_imp_pos [OF A v])
399      hence "u*v < 1*v" by (simp only: mult_strict_right_mono upos \<open>u < 1\<close> v)
400      thus "u * v \<in> A"
401        by (force intro: preal_downwards_closed [OF A v] mult_pos_pos  upos vpos)
402    qed
403  next
404    show "A \<subseteq> ?lhs"
405    proof clarify
406      fix x::rat
407      assume x: "x \<in> A"
408      have xpos: "0<x" by (rule preal_imp_pos [OF A x])
409      from preal_exists_greater [OF A x]
410      obtain v where v: "v \<in> A" and xlessv: "x < v" ..
411      have vpos: "0<v" by (rule preal_imp_pos [OF A v])
412      show "\<exists>u. 0 < u \<and> u < 1 \<and> (\<exists>v\<in>A. x = u * v)"
413      proof (intro exI conjI)
414        show "0 < x/v"
415          by (simp add: zero_less_divide_iff xpos vpos)
416        show "x / v < 1"
417          by (simp add: pos_divide_less_eq vpos xlessv)
418        have "x = (x/v)*v"
419            by (simp add: divide_inverse mult.assoc vpos order_less_imp_not_eq2)
420        then show "\<exists>v'\<in>A. x = (x / v) * v'"
421          using v by blast
422      qed
423    qed
424  qed
425  thus "1 * Abs_preal A = Abs_preal A"
426    by (simp add: preal_one_def preal_mult_def mult_set_def rat_mem_preal A)
427qed
428
429instance preal :: comm_monoid_mult
430  by intro_classes (rule preal_mult_1)
431
432
433subsection\<open>Distribution of Multiplication across Addition\<close>
434
435lemma mem_Rep_preal_add_iff:
436  "(z \<in> Rep_preal(r+s)) = (\<exists>x \<in> Rep_preal r. \<exists>y \<in> Rep_preal s. z = x + y)"
437  apply (simp add: preal_add_def mem_add_set Rep_preal)
438  apply (simp add: add_set_def) 
439  done
440
441lemma mem_Rep_preal_mult_iff:
442  "(z \<in> Rep_preal(r*s)) = (\<exists>x \<in> Rep_preal r. \<exists>y \<in> Rep_preal s. z = x * y)"
443  apply (simp add: preal_mult_def mem_mult_set Rep_preal)
444  apply (simp add: mult_set_def) 
445  done
446
447lemma distrib_subset1:
448  "Rep_preal (w * (x + y)) \<subseteq> Rep_preal (w * x + w * y)"
449  by (force simp: Bex_def mem_Rep_preal_add_iff mem_Rep_preal_mult_iff distrib_left)
450
451lemma preal_add_mult_distrib_mean:
452  assumes a: "a \<in> Rep_preal w"
453    and b: "b \<in> Rep_preal w"
454    and d: "d \<in> Rep_preal x"
455    and e: "e \<in> Rep_preal y"
456  shows "\<exists>c \<in> Rep_preal w. a * d + b * e = c * (d + e)"
457proof
458  let ?c = "(a*d + b*e)/(d+e)"
459  have [simp]: "0<a" "0<b" "0<d" "0<e" "0<d+e"
460    by (blast intro: preal_imp_pos [OF cut_Rep_preal] a b d e pos_add_strict)+
461  have cpos: "0 < ?c"
462    by (simp add: zero_less_divide_iff zero_less_mult_iff pos_add_strict)
463  show "a * d + b * e = ?c * (d + e)"
464    by (simp add: divide_inverse mult.assoc order_less_imp_not_eq2)
465  show "?c \<in> Rep_preal w"
466  proof (cases rule: linorder_le_cases)
467    assume "a \<le> b"
468    hence "?c \<le> b"
469      by (simp add: pos_divide_le_eq distrib_left mult_right_mono
470                    order_less_imp_le)
471    thus ?thesis by (rule preal_downwards_closed' [OF cut_Rep_preal b cpos])
472  next
473    assume "b \<le> a"
474    hence "?c \<le> a"
475      by (simp add: pos_divide_le_eq distrib_left mult_right_mono
476                    order_less_imp_le)
477    thus ?thesis by (rule preal_downwards_closed' [OF cut_Rep_preal a cpos])
478  qed
479qed
480
481lemma distrib_subset2:
482  "Rep_preal (w * x + w * y) \<subseteq> Rep_preal (w * (x + y))"
483  apply (clarsimp simp: mem_Rep_preal_add_iff mem_Rep_preal_mult_iff)
484  using mem_Rep_preal_add_iff preal_add_mult_distrib_mean by blast
485
486lemma preal_add_mult_distrib2: "(w * ((x::preal) + y)) = (w * x) + (w * y)"
487  by (metis Rep_preal_inverse distrib_subset1 distrib_subset2 subset_antisym)
488
489lemma preal_add_mult_distrib: "(((x::preal) + y) * w) = (x * w) + (y * w)"
490  by (simp add: preal_mult_commute preal_add_mult_distrib2)
491
492instance preal :: comm_semiring
493  by intro_classes (rule preal_add_mult_distrib)
494
495
496subsection\<open>Existence of Inverse, a Positive Real\<close>
497
498lemma mem_inverse_set:
499  assumes "cut A" shows "cut (inverse_set A)"
500proof -
501  have "\<exists>x y. 0 < x \<and> x < y \<and> inverse y \<notin> A"
502  proof -
503    from preal_exists_bound [OF \<open>cut A\<close>]
504    obtain x where [simp]: "0<x" "x \<notin> A" by blast
505    show ?thesis
506    proof (intro exI conjI)
507      show "0 < inverse (x+1)"
508        by (simp add: order_less_trans [OF _ less_add_one]) 
509      show "inverse(x+1) < inverse x"
510        by (simp add: less_imp_inverse_less less_add_one)
511      show "inverse (inverse x) \<notin> A"
512        by (simp add: order_less_imp_not_eq2)
513    qed
514  qed
515  then have "{} \<subset> inverse_set A"
516    using inverse_set_def by fastforce
517  moreover obtain q where "q > 0" "q \<notin> inverse_set A"
518  proof -
519    from preal_nonempty [OF \<open>cut A\<close>]
520    obtain x where x: "x \<in> A" and  xpos [simp]: "0<x" ..
521    show ?thesis
522    proof 
523      show "0 < inverse x" by simp
524      show "inverse x \<notin> inverse_set A"
525      proof -
526        { fix y::rat 
527          assume ygt: "inverse x < y"
528          have [simp]: "0 < y" by (simp add: order_less_trans [OF _ ygt])
529          have iyless: "inverse y < x" 
530            by (simp add: inverse_less_imp_less [of x] ygt)
531          have "inverse y \<in> A"
532            by (simp add: preal_downwards_closed [OF \<open>cut A\<close> x] iyless)}
533        thus ?thesis by (auto simp: inverse_set_def)
534      qed
535    qed
536  qed
537  moreover have "inverse_set A \<subset> {0<..}"
538    using calculation inverse_set_def by blast
539  moreover have "z \<in> inverse_set A"
540    if u: "u \<in> inverse_set A" and "0 < z" "z < u" for u z
541    using u that less_trans unfolding inverse_set_def by auto
542  moreover have "\<And>y. y \<in> inverse_set A \<Longrightarrow> \<exists>u \<in> inverse_set A. y < u"
543    by (simp add: inverse_set_def) (meson dense less_trans)
544  ultimately show ?thesis
545    by (simp add: Dedekind_Real.cut_def)
546qed
547
548
549subsection\<open>Gleason's Lemma 9-3.4, page 122\<close>
550
551lemma Gleason9_34_exists:
552  assumes A: "cut A"
553    and "\<forall>x\<in>A. x + u \<in> A"
554    and "0 \<le> z"
555  shows "\<exists>b\<in>A. b + (of_int z) * u \<in> A"
556proof (cases z rule: int_cases)
557  case (nonneg n)
558  show ?thesis
559  proof (simp add: nonneg, induct n)
560    case 0
561    from preal_nonempty [OF A]
562    show ?case  by force 
563  next
564    case (Suc k)
565    then obtain b where b: "b \<in> A" "b + of_nat k * u \<in> A" ..
566    hence "b + of_int (int k)*u + u \<in> A" by (simp add: assms)
567    thus ?case by (force simp: algebra_simps b)
568  qed
569next
570  case (neg n)
571  with assms show ?thesis by simp
572qed
573
574lemma Gleason9_34_contra:
575  assumes A: "cut A"
576    shows "\<lbrakk>\<forall>x\<in>A. x + u \<in> A; 0 < u; 0 < y; y \<notin> A\<rbrakk> \<Longrightarrow> False"
577proof (induct u, induct y)
578  fix a::int and b::int
579  fix c::int and d::int
580  assume bpos [simp]: "0 < b"
581    and dpos [simp]: "0 < d"
582    and closed: "\<forall>x\<in>A. x + (Fract c d) \<in> A"
583    and upos: "0 < Fract c d"
584    and ypos: "0 < Fract a b"
585    and notin: "Fract a b \<notin> A"
586  have cpos [simp]: "0 < c" 
587    by (simp add: zero_less_Fract_iff [OF dpos, symmetric] upos) 
588  have apos [simp]: "0 < a" 
589    by (simp add: zero_less_Fract_iff [OF bpos, symmetric] ypos) 
590  let ?k = "a*d"
591  have frle: "Fract a b \<le> Fract ?k 1 * (Fract c d)" 
592  proof -
593    have "?thesis = ((a * d * b * d) \<le> c * b * (a * d * b * d))"
594      by (simp add: order_less_imp_not_eq2 ac_simps) 
595    moreover
596    have "(1 * (a * d * b * d)) \<le> c * b * (a * d * b * d)"
597      by (rule mult_mono, 
598          simp_all add: int_one_le_iff_zero_less zero_less_mult_iff 
599                        order_less_imp_le)
600    ultimately
601    show ?thesis by simp
602  qed
603  have k: "0 \<le> ?k" by (simp add: order_less_imp_le zero_less_mult_iff)  
604  from Gleason9_34_exists [OF A closed k]
605  obtain z where z: "z \<in> A" 
606             and mem: "z + of_int ?k * Fract c d \<in> A" ..
607  have less: "z + of_int ?k * Fract c d < Fract a b"
608    by (rule not_in_preal_ub [OF A notin mem ypos])
609  have "0<z" by (rule preal_imp_pos [OF A z])
610  with frle and less show False by (simp add: Fract_of_int_eq) 
611qed
612
613
614lemma Gleason9_34:
615  assumes "cut A" "0 < u"
616  shows "\<exists>r \<in> A. r + u \<notin> A"
617  using assms Gleason9_34_contra preal_exists_bound by blast
618
619
620
621subsection\<open>Gleason's Lemma 9-3.6\<close>
622
623lemma lemma_gleason9_36:
624  assumes A: "cut A"
625    and x: "1 < x"
626  shows "\<exists>r \<in> A. r*x \<notin> A"
627proof -
628  from preal_nonempty [OF A]
629  obtain y where y: "y \<in> A" and  ypos: "0<y" ..
630  show ?thesis 
631  proof (rule classical)
632    assume "~(\<exists>r\<in>A. r * x \<notin> A)"
633    with y have ymem: "y * x \<in> A" by blast 
634    from ypos mult_strict_left_mono [OF x]
635    have yless: "y < y*x" by simp 
636    let ?d = "y*x - y"
637    from yless have dpos: "0 < ?d" and eq: "y + ?d = y*x" by auto
638    from Gleason9_34 [OF A dpos]
639    obtain r where r: "r\<in>A" and notin: "r + ?d \<notin> A" ..
640    have rpos: "0<r" by (rule preal_imp_pos [OF A r])
641    with dpos have rdpos: "0 < r + ?d" by arith
642    have "~ (r + ?d \<le> y + ?d)"
643    proof
644      assume le: "r + ?d \<le> y + ?d" 
645      from ymem have yd: "y + ?d \<in> A" by (simp add: eq)
646      have "r + ?d \<in> A" by (rule preal_downwards_closed' [OF A yd rdpos le])
647      with notin show False by simp
648    qed
649    hence "y < r" by simp
650    with ypos have  dless: "?d < (r * ?d)/y"
651      using dpos less_divide_eq_1 by fastforce
652    have "r + ?d < r*x"
653    proof -
654      have "r + ?d < r + (r * ?d)/y" by (simp add: dless)
655      also from ypos have "\<dots> = (r/y) * (y + ?d)"
656        by (simp only: algebra_simps divide_inverse, simp)
657      also have "\<dots> = r*x" using ypos
658        by simp
659      finally show "r + ?d < r*x" .
660    qed
661    with r notin rdpos
662    show "\<exists>r\<in>A. r * x \<notin> A" by (blast dest:  preal_downwards_closed [OF A])
663  qed  
664qed
665
666subsection\<open>Existence of Inverse: Part 2\<close>
667
668lemma mem_Rep_preal_inverse_iff:
669  "(z \<in> Rep_preal(inverse r)) \<longleftrightarrow> (0 < z \<and> (\<exists>y. z < y \<and> inverse y \<notin> Rep_preal r))"
670  apply (simp add: preal_inverse_def mem_inverse_set Rep_preal)
671  apply (simp add: inverse_set_def) 
672  done
673
674lemma Rep_preal_one:
675     "Rep_preal 1 = {x. 0 < x \<and> x < 1}"
676by (simp add: preal_one_def rat_mem_preal)
677
678lemma subset_inverse_mult_lemma:
679  assumes xpos: "0 < x" and xless: "x < 1"
680  shows "\<exists>v u y. 0 < v \<and> v < y \<and> inverse y \<notin> Rep_preal R \<and> 
681    u \<in> Rep_preal R \<and> x = v * u"
682proof -
683  from xpos and xless have "1 < inverse x" by (simp add: one_less_inverse_iff)
684  from lemma_gleason9_36 [OF cut_Rep_preal this]
685  obtain t where t: "t \<in> Rep_preal R" 
686             and notin: "t * (inverse x) \<notin> Rep_preal R" ..
687  have rpos: "0<t" by (rule preal_imp_pos [OF cut_Rep_preal t])
688  from preal_exists_greater [OF cut_Rep_preal t]
689  obtain u where u: "u \<in> Rep_preal R" and rless: "t < u" ..
690  have upos: "0<u" by (rule preal_imp_pos [OF cut_Rep_preal u])
691  show ?thesis
692  proof (intro exI conjI)
693    show "0 < x/u" using xpos upos
694      by (simp add: zero_less_divide_iff)  
695    show "x/u < x/t" using xpos upos rpos
696      by (simp add: divide_inverse mult_less_cancel_left rless) 
697    show "inverse (x / t) \<notin> Rep_preal R" using notin
698      by (simp add: divide_inverse mult.commute) 
699    show "u \<in> Rep_preal R" by (rule u) 
700    show "x = x / u * u" using upos 
701      by (simp add: divide_inverse mult.commute) 
702  qed
703qed
704
705lemma subset_inverse_mult: 
706     "Rep_preal 1 \<subseteq> Rep_preal(inverse r * r)"
707  by (force simp: Rep_preal_one mem_Rep_preal_inverse_iff mem_Rep_preal_mult_iff dest: subset_inverse_mult_lemma)
708
709lemma inverse_mult_subset: "Rep_preal(inverse r * r) \<subseteq> Rep_preal 1"
710  proof -
711  have "0 < u * v" if "v \<in> Rep_preal r" "0 < u" "u < t" for u v t :: rat
712    using that by (simp add: zero_less_mult_iff preal_imp_pos [OF cut_Rep_preal]) 
713  moreover have "t * q < 1"
714    if "q \<in> Rep_preal r" "0 < t" "t < y" "inverse y \<notin> Rep_preal r"
715    for t q y :: rat
716  proof -
717    have "q < inverse y"
718      using not_in_Rep_preal_ub that by auto 
719    hence "t * q < t/y" 
720      using that by (simp add: divide_inverse mult_less_cancel_left)
721    also have "\<dots> \<le> 1" 
722      using that by (simp add: pos_divide_le_eq)
723    finally show ?thesis .
724  qed
725  ultimately show ?thesis
726    by (auto simp: Rep_preal_one mem_Rep_preal_inverse_iff mem_Rep_preal_mult_iff)
727qed 
728
729lemma preal_mult_inverse: "inverse r * r = (1::preal)"
730  by (meson Rep_preal_inject inverse_mult_subset subset_antisym subset_inverse_mult)
731
732lemma preal_mult_inverse_right: "r * inverse r = (1::preal)"
733  using preal_mult_commute preal_mult_inverse by auto
734
735
736text\<open>Theorems needing \<open>Gleason9_34\<close>\<close>
737
738lemma Rep_preal_self_subset: "Rep_preal (r) \<subseteq> Rep_preal(r + s)"
739proof 
740  fix x
741  assume x: "x \<in> Rep_preal r"
742  obtain y where y: "y \<in> Rep_preal s" and "y > 0"
743    using Rep_preal preal_nonempty by blast
744  have ry: "x+y \<in> Rep_preal(r + s)" using x y
745    by (auto simp: mem_Rep_preal_add_iff)
746  then show "x \<in> Rep_preal(r + s)"
747    by (meson \<open>0 < y\<close> add_less_same_cancel1 not_in_Rep_preal_ub order.asym preal_imp_pos [OF cut_Rep_preal x])
748qed
749
750lemma Rep_preal_sum_not_subset: "~ Rep_preal (r + s) \<subseteq> Rep_preal(r)"
751proof -
752  obtain y where y: "y \<in> Rep_preal s" and "y > 0"
753    using Rep_preal preal_nonempty by blast
754  obtain x where "x \<in> Rep_preal r" and notin: "x + y \<notin> Rep_preal r"
755    using Dedekind_Real.Rep_preal Gleason9_34 \<open>0 < y\<close> by blast 
756  then have "x + y \<in> Rep_preal (r + s)" using y
757    by (auto simp: mem_Rep_preal_add_iff)
758  thus ?thesis using notin by blast
759qed
760
761text\<open>at last, Gleason prop. 9-3.5(iii) page 123\<close>
762proposition preal_self_less_add_left: "(r::preal) < r + s"
763  by (meson Rep_preal_sum_not_subset not_less preal_le_def)
764
765
766subsection\<open>Subtraction for Positive Reals\<close>
767
768text\<open>gleason prop. 9-3.5(iv), page 123: proving \<^prop>\<open>a < b \<Longrightarrow> \<exists>d. a + d = b\<close>. 
769We define the claimed \<^term>\<open>D\<close> and show that it is a positive real\<close>
770
771lemma mem_diff_set:
772  assumes "r < s"
773  shows "cut (diff_set (Rep_preal s) (Rep_preal r))"
774proof -
775  obtain p where "Rep_preal r \<subseteq> Rep_preal s" "p \<in> Rep_preal s" "p \<notin> Rep_preal r"
776    using assms unfolding preal_less_def by auto
777  then have "{} \<subset> diff_set (Rep_preal s) (Rep_preal r)"
778    apply (simp add: diff_set_def psubset_eq)
779    by (metis cut_Rep_preal add_eq_exists less_add_same_cancel1 preal_exists_greater preal_imp_pos)
780  moreover
781  obtain q where "q > 0" "q \<notin> Rep_preal s"
782    using Rep_preal_exists_bound by blast
783  then have qnot: "q \<notin> diff_set (Rep_preal s) (Rep_preal r)"
784    by (auto simp: diff_set_def dest: cut_Rep_preal [THEN preal_downwards_closed])
785  moreover have "diff_set (Rep_preal s) (Rep_preal r) \<subset> {0<..}" (is "?lhs < ?rhs")
786    using \<open>0 < q\<close> diff_set_def qnot by blast
787  moreover have "z \<in> diff_set (Rep_preal s) (Rep_preal r)"
788    if u: "u \<in> diff_set (Rep_preal s) (Rep_preal r)" and "0 < z" "z < u" for u z
789    using u that less_trans Rep_preal unfolding diff_set_def Dedekind_Real.cut_def by auto
790  moreover have "\<exists>u \<in> diff_set (Rep_preal s) (Rep_preal r). y < u"
791    if y: "y \<in> diff_set (Rep_preal s) (Rep_preal r)" for y
792  proof -
793    obtain a b where "0 < a" "0 < b" "a \<notin> Rep_preal r" "a + y + b \<in> Rep_preal s"
794      using y
795      by (simp add: diff_set_def) (metis cut_Rep_preal add_eq_exists less_add_same_cancel1 preal_exists_greater) 
796    then have "a + (y + b) \<in> Rep_preal s"
797      by (simp add: add.assoc)
798    then have "y + b \<in> diff_set (Rep_preal s) (Rep_preal r)"
799      using \<open>0 < a\<close> \<open>0 < b\<close> \<open>a \<notin> Rep_preal r\<close> y
800      by (auto simp: diff_set_def)
801    then show ?thesis
802      using \<open>0 < b\<close> less_add_same_cancel1 by blast
803  qed
804  ultimately show ?thesis
805    by (simp add: Dedekind_Real.cut_def)
806qed
807
808lemma mem_Rep_preal_diff_iff:
809  "r < s \<Longrightarrow>
810       (z \<in> Rep_preal (s - r)) \<longleftrightarrow> 
811       (\<exists>x. 0 < x \<and> 0 < z \<and> x \<notin> Rep_preal r \<and> x + z \<in> Rep_preal s)"
812  apply (simp add: preal_diff_def mem_diff_set Rep_preal)
813  apply (force simp: diff_set_def) 
814  done
815
816proposition less_add_left:
817  fixes r::preal 
818  assumes "r < s"
819  shows "r + (s-r) = s"
820proof -
821  have "a + b \<in> Rep_preal s"
822    if "a \<in> Rep_preal r" "c + b \<in> Rep_preal s" "c \<notin> Rep_preal r"
823    and "0 < b" "0 < c" for a b c
824    by (meson cut_Rep_preal add_less_imp_less_right add_pos_pos not_in_Rep_preal_ub preal_downwards_closed preal_imp_pos that)
825  then have "r + (s-r) \<le> s"
826    using assms mem_Rep_preal_add_iff mem_Rep_preal_diff_iff preal_le_def by auto
827  have "x \<in> Rep_preal (r + (s - r))" if "x \<in> Rep_preal s" for x
828  proof (cases "x \<in> Rep_preal r")
829    case True
830    then show ?thesis
831      using Rep_preal_self_subset by blast
832  next
833    case False
834    have "\<exists>u v z. 0 < v \<and> 0 < z \<and> u \<in> Rep_preal r \<and> z \<notin> Rep_preal r \<and> z + v \<in> Rep_preal s \<and> x = u + v"
835      if x: "x \<in> Rep_preal s"
836    proof -
837      have xpos: "x > 0"
838        using Rep_preal preal_imp_pos that by blast 
839      obtain e where epos: "0 < e" and xe: "x + e \<in> Rep_preal s"
840        by (metis cut_Rep_preal x add_eq_exists less_add_same_cancel1 preal_exists_greater)
841      from  Gleason9_34 [OF cut_Rep_preal epos]
842      obtain u where r: "u \<in> Rep_preal r" and notin: "u + e \<notin> Rep_preal r" ..
843      with x False xpos have rless: "u < x" by (blast intro: not_in_Rep_preal_ub)
844      from add_eq_exists [of u x]
845      obtain y where eq: "x = u+y" by auto
846      show ?thesis 
847      proof (intro exI conjI)
848        show "u + e \<notin> Rep_preal r" by (rule notin)
849        show "u + e + y \<in> Rep_preal s" using xe eq by (simp add: ac_simps)
850        show "0 < u + e" 
851          using epos preal_imp_pos [OF cut_Rep_preal r] by simp
852      qed (use r rless eq in auto)
853    qed
854    then show ?thesis
855      using assms mem_Rep_preal_add_iff mem_Rep_preal_diff_iff that by blast
856  qed
857  then have "s \<le> r + (s-r)"
858    by (auto simp: preal_le_def)
859  then show ?thesis
860    by (simp add: \<open>r + (s - r) \<le> s\<close> antisym)
861qed
862
863lemma preal_add_less2_mono1: "r < (s::preal) \<Longrightarrow> r + t < s + t"
864  by (metis add.assoc add.commute less_add_left preal_self_less_add_left)
865
866lemma preal_add_less2_mono2: "r < (s::preal) \<Longrightarrow> t + r < t + s"
867  by (auto intro: preal_add_less2_mono1 simp add: preal_add_commute [of t])
868
869lemma preal_add_right_less_cancel: "r + t < s + t \<Longrightarrow> r < (s::preal)"
870  by (metis linorder_cases order.asym preal_add_less2_mono1)
871
872lemma preal_add_left_less_cancel: "t + r < t + s \<Longrightarrow> r < (s::preal)"
873  by (auto elim: preal_add_right_less_cancel simp add: preal_add_commute [of t])
874
875lemma preal_add_less_cancel_left [simp]: "(t + (r::preal) < t + s) \<longleftrightarrow> (r < s)"
876  by (blast intro: preal_add_less2_mono2 preal_add_left_less_cancel)
877
878lemma preal_add_less_cancel_right [simp]: "((r::preal) + t < s + t) = (r < s)"
879  using preal_add_less_cancel_left [symmetric, of r s t] by (simp add: ac_simps)
880
881lemma preal_add_le_cancel_left [simp]: "(t + (r::preal) \<le> t + s) = (r \<le> s)"
882  by (simp add: linorder_not_less [symmetric]) 
883
884lemma preal_add_le_cancel_right [simp]: "((r::preal) + t \<le> s + t) = (r \<le> s)"
885  using preal_add_le_cancel_left [symmetric, of r s t] by (simp add: ac_simps)
886
887lemma preal_add_right_cancel: "(r::preal) + t = s + t \<Longrightarrow> r = s"
888  by (metis less_irrefl linorder_cases preal_add_less_cancel_right)
889
890lemma preal_add_left_cancel: "c + a = c + b \<Longrightarrow> a = (b::preal)"
891  by (auto intro: preal_add_right_cancel simp add: preal_add_commute)
892
893instance preal :: linordered_ab_semigroup_add
894proof
895  fix a b c :: preal
896  show "a \<le> b \<Longrightarrow> c + a \<le> c + b" by (simp only: preal_add_le_cancel_left)
897qed
898
899
900subsection\<open>Completeness of type \<^typ>\<open>preal\<close>\<close>
901
902text\<open>Prove that supremum is a cut\<close>
903
904text\<open>Part 1 of Dedekind sections definition\<close>
905
906lemma preal_sup:
907  assumes le: "\<And>X. X \<in> P \<Longrightarrow> X \<le> Y" and "P \<noteq> {}" 
908  shows "cut (\<Union>X \<in> P. Rep_preal(X))"
909proof -
910  have "{} \<subset> (\<Union>X \<in> P. Rep_preal(X))"
911    using \<open>P \<noteq> {}\<close> mem_Rep_preal_Ex by fastforce
912  moreover
913  obtain q where "q > 0" and "q \<notin> (\<Union>X \<in> P. Rep_preal(X))"
914    using Rep_preal_exists_bound [of Y] le by (auto simp: preal_le_def)
915  then have "(\<Union>X \<in> P. Rep_preal(X)) \<subset> {0<..}"
916    using cut_Rep_preal preal_imp_pos by force
917  moreover
918  have "\<And>u z. \<lbrakk>u \<in> (\<Union>X \<in> P. Rep_preal(X)); 0 < z; z < u\<rbrakk> \<Longrightarrow> z \<in> (\<Union>X \<in> P. Rep_preal(X))"
919    by (auto elim: cut_Rep_preal [THEN preal_downwards_closed])
920  moreover
921  have "\<And>y. y \<in> (\<Union>X \<in> P. Rep_preal(X)) \<Longrightarrow> \<exists>u \<in> (\<Union>X \<in> P. Rep_preal(X)). y < u"
922    by (blast dest: cut_Rep_preal [THEN preal_exists_greater])
923  ultimately show ?thesis
924    by (simp add: Dedekind_Real.cut_def)
925qed
926
927lemma preal_psup_le:
928     "\<lbrakk>\<And>X. X \<in> P \<Longrightarrow> X \<le> Y;  x \<in> P\<rbrakk> \<Longrightarrow> x \<le> psup P"
929  using preal_sup [of P Y] unfolding preal_le_def psup_def by fastforce 
930
931lemma psup_le_ub: "\<lbrakk>\<And>X. X \<in> P \<Longrightarrow> X \<le> Y; P \<noteq> {}\<rbrakk> \<Longrightarrow> psup P \<le> Y"
932  using preal_sup [of P Y] by (simp add: SUP_least preal_le_def psup_def) 
933
934text\<open>Supremum property\<close>
935proposition preal_complete:
936  assumes le: "\<And>X. X \<in> P \<Longrightarrow> X \<le> Y" and "P \<noteq> {}" 
937  shows "(\<exists>X \<in> P. Z < X) \<longleftrightarrow> (Z < psup P)" (is "?lhs = ?rhs")
938proof
939  assume ?lhs
940  then show ?rhs
941    using preal_sup [OF assms] preal_less_def psup_def by auto
942next
943  assume ?rhs
944  then show ?lhs
945    by (meson \<open>P \<noteq> {}\<close> not_less psup_le_ub) 
946qed
947
948subsection \<open>Defining the Reals from the Positive Reals\<close>
949
950text \<open>Here we do quotients the old-fashioned way\<close>
951
952definition
953  realrel   ::  "((preal * preal) * (preal * preal)) set" where
954  "realrel = {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) \<and> x1+y2 = x2+y1}"
955
956definition "Real = UNIV//realrel"
957
958typedef real = Real
959  morphisms Rep_Real Abs_Real
960  unfolding Real_def by (auto simp: quotient_def)
961
962text \<open>This doesn't involve the overloaded "real" function: users don't see it\<close>
963definition
964  real_of_preal :: "preal \<Rightarrow> real" where
965  "real_of_preal m = Abs_Real (realrel `` {(m + 1, 1)})"
966
967instantiation real :: "{zero, one, plus, minus, uminus, times, inverse, ord, abs, sgn}"
968begin
969
970definition
971  real_zero_def: "0 = Abs_Real(realrel``{(1, 1)})"
972
973definition
974  real_one_def: "1 = Abs_Real(realrel``{(1 + 1, 1)})"
975
976definition
977  real_add_def: "z + w =
978       the_elem (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
979                 { Abs_Real(realrel``{(x+u, y+v)}) })"
980
981definition
982  real_minus_def: "- r =  the_elem (\<Union>(x,y) \<in> Rep_Real(r). { Abs_Real(realrel``{(y,x)}) })"
983
984definition
985  real_diff_def: "r - (s::real) = r + - s"
986
987definition
988  real_mult_def:
989    "z * w =
990       the_elem (\<Union>(x,y) \<in> Rep_Real(z). \<Union>(u,v) \<in> Rep_Real(w).
991                 { Abs_Real(realrel``{(x*u + y*v, x*v + y*u)}) })"
992
993definition
994  real_inverse_def: "inverse (r::real) = (THE s. (r = 0 \<and> s = 0) \<or> s * r = 1)"
995
996definition
997  real_divide_def: "r div (s::real) = r * inverse s"
998
999definition
1000  real_le_def: "z \<le> (w::real) \<longleftrightarrow>
1001    (\<exists>x y u v. x+v \<le> u+y \<and> (x,y) \<in> Rep_Real z \<and> (u,v) \<in> Rep_Real w)"
1002
1003definition
1004  real_less_def: "x < (y::real) \<longleftrightarrow> x \<le> y \<and> x \<noteq> y"
1005
1006definition
1007  real_abs_def: "\<bar>r::real\<bar> = (if r < 0 then - r else r)"
1008
1009definition
1010  real_sgn_def: "sgn (x::real) = (if x=0 then 0 else if 0<x then 1 else - 1)"
1011
1012instance ..
1013
1014end
1015
1016subsection \<open>Equivalence relation over positive reals\<close>
1017
1018lemma realrel_iff [simp]: "(((x1,y1),(x2,y2)) \<in> realrel) = (x1 + y2 = x2 + y1)"
1019  by (simp add: realrel_def)
1020
1021lemma preal_trans_lemma:
1022  assumes "x + y1 = x1 + y" and "x + y2 = x2 + y"
1023  shows "x1 + y2 = x2 + (y1::preal)"
1024  by (metis add.left_commute assms preal_add_left_cancel)
1025
1026lemma equiv_realrel: "equiv UNIV realrel"
1027  by (auto simp: equiv_def refl_on_def sym_def trans_def realrel_def intro: dest: preal_trans_lemma)
1028
1029text\<open>Reduces equality of equivalence classes to the \<^term>\<open>realrel\<close> relation:
1030  \<^term>\<open>(realrel `` {x} = realrel `` {y}) = ((x,y) \<in> realrel)\<close>\<close>
1031lemmas equiv_realrel_iff [simp] = 
1032       eq_equiv_class_iff [OF equiv_realrel UNIV_I UNIV_I]
1033
1034lemma realrel_in_real [simp]: "realrel``{(x,y)} \<in> Real"
1035  by (simp add: Real_def realrel_def quotient_def, blast)
1036
1037declare Abs_Real_inject [simp] Abs_Real_inverse [simp]
1038
1039
1040text\<open>Case analysis on the representation of a real number as an equivalence
1041      class of pairs of positive reals.\<close>
1042lemma eq_Abs_Real [case_names Abs_Real, cases type: real]: 
1043     "(\<And>x y. z = Abs_Real(realrel``{(x,y)}) \<Longrightarrow> P) \<Longrightarrow> P"
1044  by (metis Rep_Real_inverse prod.exhaust  Rep_Real [of z, unfolded Real_def, THEN quotientE])
1045
1046subsection \<open>Addition and Subtraction\<close>
1047
1048lemma real_add:
1049     "Abs_Real (realrel``{(x,y)}) + Abs_Real (realrel``{(u,v)}) =
1050      Abs_Real (realrel``{(x+u, y+v)})"
1051proof -
1052  have "(\<lambda>z w. (\<lambda>(x,y). (\<lambda>(u,v). {Abs_Real (realrel `` {(x+u, y+v)})}) w) z)
1053        respects2 realrel"
1054  by (clarsimp simp: congruent2_def) (metis add.left_commute preal_add_assoc)
1055  thus ?thesis
1056    by (simp add: real_add_def UN_UN_split_split_eq UN_equiv_class2 [OF equiv_realrel equiv_realrel])
1057qed
1058
1059lemma real_minus: "- Abs_Real(realrel``{(x,y)}) = Abs_Real(realrel `` {(y,x)})"
1060proof -
1061  have "(\<lambda>(x,y). {Abs_Real (realrel``{(y,x)})}) respects realrel"
1062    by (auto simp: congruent_def add.commute) 
1063  thus ?thesis
1064    by (simp add: real_minus_def UN_equiv_class [OF equiv_realrel])
1065qed
1066
1067instance real :: ab_group_add
1068proof
1069  fix x y z :: real
1070  show "(x + y) + z = x + (y + z)"
1071    by (cases x, cases y, cases z, simp add: real_add add.assoc)
1072  show "x + y = y + x"
1073    by (cases x, cases y, simp add: real_add add.commute)
1074  show "0 + x = x"
1075    by (cases x, simp add: real_add real_zero_def ac_simps)
1076  show "- x + x = 0"
1077    by (cases x, simp add: real_minus real_add real_zero_def add.commute)
1078  show "x - y = x + - y"
1079    by (simp add: real_diff_def)
1080qed
1081
1082
1083subsection \<open>Multiplication\<close>
1084
1085lemma real_mult_congruent2_lemma:
1086     "!!(x1::preal). \<lbrakk>x1 + y2 = x2 + y1\<rbrakk> \<Longrightarrow>
1087          x * x1 + y * y1 + (x * y2 + y * x2) =
1088          x * x2 + y * y2 + (x * y1 + y * x1)"
1089  by (metis (no_types, hide_lams) add.left_commute preal_add_commute preal_add_mult_distrib2)
1090
1091lemma real_mult_congruent2:
1092  "(\<lambda>p1 p2.
1093        (\<lambda>(x1,y1). (\<lambda>(x2,y2). 
1094          { Abs_Real (realrel``{(x1*x2 + y1*y2, x1*y2+y1*x2)}) }) p2) p1)
1095     respects2 realrel"
1096  apply (rule congruent2_commuteI [OF equiv_realrel])
1097  by (auto simp: mult.commute add.commute combine_common_factor preal_add_assoc preal_add_commute)
1098
1099lemma real_mult:
1100  "Abs_Real((realrel``{(x1,y1)})) * Abs_Real((realrel``{(x2,y2)})) =
1101   Abs_Real(realrel `` {(x1*x2+y1*y2,x1*y2+y1*x2)})"
1102  by (simp add: real_mult_def UN_UN_split_split_eq
1103      UN_equiv_class2 [OF equiv_realrel equiv_realrel real_mult_congruent2])
1104
1105lemma real_mult_commute: "(z::real) * w = w * z"
1106by (cases z, cases w, simp add: real_mult ac_simps)
1107
1108lemma real_mult_assoc: "((z1::real) * z2) * z3 = z1 * (z2 * z3)"
1109  by (cases z1, cases z2, cases z3) (simp add: real_mult algebra_simps)
1110
1111lemma real_mult_1: "(1::real) * z = z"
1112  by (cases z) (simp add: real_mult real_one_def algebra_simps)
1113
1114lemma real_add_mult_distrib: "((z1::real) + z2) * w = (z1 * w) + (z2 * w)"
1115  by (cases z1, cases z2, cases w) (simp add: real_add real_mult algebra_simps)
1116
1117text\<open>one and zero are distinct\<close>
1118lemma real_zero_not_eq_one: "0 \<noteq> (1::real)"
1119proof -
1120  have "(1::preal) < 1 + 1"
1121    by (simp add: preal_self_less_add_left)
1122  then show ?thesis
1123    by (simp add: real_zero_def real_one_def neq_iff)
1124qed
1125
1126instance real :: comm_ring_1
1127proof
1128  fix x y z :: real
1129  show "(x * y) * z = x * (y * z)" by (rule real_mult_assoc)
1130  show "x * y = y * x" by (rule real_mult_commute)
1131  show "1 * x = x" by (rule real_mult_1)
1132  show "(x + y) * z = x * z + y * z" by (rule real_add_mult_distrib)
1133  show "0 \<noteq> (1::real)" by (rule real_zero_not_eq_one)
1134qed
1135
1136subsection \<open>Inverse and Division\<close>
1137
1138lemma real_zero_iff: "Abs_Real (realrel `` {(x, x)}) = 0"
1139  by (simp add: real_zero_def add.commute)
1140
1141lemma real_mult_inverse_left_ex:
1142  assumes "x \<noteq> 0" obtains y::real where "y*x = 1"
1143proof (cases x)
1144  case (Abs_Real u v)
1145  show ?thesis
1146  proof (cases u v rule: linorder_cases)
1147    case less
1148    then have "v * inverse (v - u) = 1 + u * inverse (v - u)"
1149      using less_add_left [of u v]
1150      by (metis preal_add_commute preal_add_mult_distrib preal_mult_inverse_right)
1151    then have "Abs_Real (realrel``{(1, inverse (v-u) + 1)}) * x - 1 = 0"
1152      by (simp add: Abs_Real real_mult preal_mult_inverse_right real_one_def) (simp add: algebra_simps)
1153    with that show thesis by auto
1154  next
1155    case equal
1156    then show ?thesis
1157      using Abs_Real assms real_zero_iff by blast
1158  next
1159    case greater
1160    then have "u * inverse (u - v) = 1 + v * inverse (u - v)"
1161      using less_add_left [of v u] by (metis add.commute distrib_right preal_mult_inverse_right)
1162    then have "Abs_Real (realrel``{(inverse (u-v) + 1, 1)}) * x - 1 = 0"
1163      by (simp add: Abs_Real real_mult preal_mult_inverse_right real_one_def) (simp add: algebra_simps)
1164    with that show thesis by auto
1165  qed
1166qed
1167
1168
1169lemma real_mult_inverse_left:
1170  fixes x :: real
1171  assumes "x \<noteq> 0" shows "inverse x * x = 1"
1172proof -
1173  obtain y where "y*x = 1"
1174    using assms real_mult_inverse_left_ex by blast
1175  then have "(THE s. s * x = 1) * x = 1"
1176  proof (rule theI)
1177    show "y' = y" if "y' * x = 1" for y'
1178      by (metis \<open>y * x = 1\<close> mult.left_commute mult.right_neutral that) 
1179  qed
1180  then show ?thesis
1181    using assms real_inverse_def by auto
1182qed
1183
1184
1185subsection\<open>The Real Numbers form a Field\<close>
1186
1187instance real :: field
1188proof
1189  fix x y z :: real
1190  show "x \<noteq> 0 \<Longrightarrow> inverse x * x = 1" by (rule real_mult_inverse_left)
1191  show "x / y = x * inverse y" by (simp add: real_divide_def)
1192  show "inverse 0 = (0::real)" by (simp add: real_inverse_def)
1193qed
1194
1195
1196subsection\<open>The \<open>\<le>\<close> Ordering\<close>
1197
1198lemma real_le_refl: "w \<le> (w::real)"
1199  by (cases w, force simp: real_le_def)
1200
1201text\<open>The arithmetic decision procedure is not set up for type preal.
1202  This lemma is currently unused, but it could simplify the proofs of the
1203  following two lemmas.\<close>
1204lemma preal_eq_le_imp_le:
1205  assumes eq: "a+b = c+d" and le: "c \<le> a"
1206  shows "b \<le> (d::preal)"
1207proof -
1208  from le have "c+d \<le> a+d" by simp
1209  hence "a+b \<le> a+d" by (simp add: eq)
1210  thus "b \<le> d" by simp
1211qed
1212
1213lemma real_le_lemma:
1214  assumes l: "u1 + v2 \<le> u2 + v1"
1215    and "x1 + v1 = u1 + y1"
1216    and "x2 + v2 = u2 + y2"
1217  shows "x1 + y2 \<le> x2 + (y1::preal)"
1218proof -
1219  have "(x1+v1) + (u2+y2) = (u1+y1) + (x2+v2)" by (simp add: assms)
1220  hence "(x1+y2) + (u2+v1) = (x2+y1) + (u1+v2)" by (simp add: ac_simps)
1221  also have "\<dots> \<le> (x2+y1) + (u2+v1)" by (simp add: assms)
1222  finally show ?thesis by simp
1223qed
1224
1225lemma real_le: 
1226  "Abs_Real(realrel``{(x1,y1)}) \<le> Abs_Real(realrel``{(x2,y2)})  \<longleftrightarrow>  x1 + y2 \<le> x2 + y1"
1227  unfolding real_le_def by (auto intro: real_le_lemma)
1228
1229lemma real_le_antisym: "\<lbrakk>z \<le> w; w \<le> z\<rbrakk> \<Longrightarrow> z = (w::real)"
1230  by (cases z, cases w, simp add: real_le)
1231
1232lemma real_trans_lemma:
1233  assumes "x + v \<le> u + y"
1234    and "u + v' \<le> u' + v"
1235    and "x2 + v2 = u2 + y2"
1236  shows "x + v' \<le> u' + (y::preal)"
1237proof -
1238  have "(x+v') + (u+v) = (x+v) + (u+v')" by (simp add: ac_simps)
1239  also have "\<dots> \<le> (u+y) + (u+v')" by (simp add: assms)
1240  also have "\<dots> \<le> (u+y) + (u'+v)" by (simp add: assms)
1241  also have "\<dots> = (u'+y) + (u+v)"  by (simp add: ac_simps)
1242  finally show ?thesis by simp
1243qed
1244
1245lemma real_le_trans: "\<lbrakk>i \<le> j; j \<le> k\<rbrakk> \<Longrightarrow> i \<le> (k::real)"
1246  by (cases i, cases j, cases k) (auto simp: real_le intro: real_trans_lemma)
1247
1248instance real :: order
1249proof
1250  show "u < v \<longleftrightarrow> u \<le> v \<and> \<not> v \<le> u" for u v::real
1251    by (auto simp: real_less_def intro: real_le_antisym)
1252qed (auto intro: real_le_refl real_le_trans real_le_antisym)
1253
1254instance real :: linorder
1255proof
1256  show "x \<le> y \<or> y \<le> x" for x y :: real
1257    by (meson eq_refl le_cases real_le_def)
1258qed
1259
1260instantiation real :: distrib_lattice
1261begin
1262
1263definition
1264  "(inf :: real \<Rightarrow> real \<Rightarrow> real) = min"
1265
1266definition
1267  "(sup :: real \<Rightarrow> real \<Rightarrow> real) = max"
1268
1269instance
1270  by standard (auto simp: inf_real_def sup_real_def max_min_distrib2)
1271
1272end
1273
1274subsection\<open>The Reals Form an Ordered Field\<close>
1275
1276lemma real_le_eq_diff: "(x \<le> y) \<longleftrightarrow> (x-y \<le> (0::real))"
1277  by (cases x, cases y) (simp add: real_le real_zero_def real_diff_def real_add real_minus preal_add_commute)
1278
1279lemma real_add_left_mono: 
1280  assumes le: "x \<le> y" shows "z + x \<le> z + (y::real)"
1281proof -
1282  have "z + x - (z + y) = (z + -z) + (x - y)" 
1283    by (simp add: algebra_simps) 
1284  with le show ?thesis 
1285    by (simp add: real_le_eq_diff[of x] real_le_eq_diff[of "z+x"])
1286qed
1287
1288lemma real_sum_gt_zero_less: "(0 < s + (-w::real)) \<Longrightarrow> (w < s)"
1289  by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of s])
1290
1291lemma real_less_sum_gt_zero: "(w < s) \<Longrightarrow> (0 < s + (-w::real))"
1292  by (simp add: linorder_not_le [symmetric] real_le_eq_diff [of s])
1293
1294lemma real_mult_order: 
1295  fixes x y::real
1296  assumes "0 < x" "0 < y"
1297  shows "0 < x * y"
1298  proof (cases x, cases y)
1299  show "0 < x * y"
1300    if x: "x = Abs_Real (Dedekind_Real.realrel `` {(x1, x2)})"
1301      and y: "y = Abs_Real (Dedekind_Real.realrel `` {(y1, y2)})"
1302    for x1 x2 y1 y2
1303  proof -
1304    have "x2 < x1" "y2 < y1"
1305      using assms not_le real_zero_def real_le x y
1306      by (metis preal_add_le_cancel_left real_zero_iff)+
1307    then obtain xd yd where "x1 = x2 + xd" "y1 = y2 + yd"
1308      using less_add_left by metis
1309    then have "\<not> (x * y \<le> 0)"
1310      apply (simp add: x y real_mult real_zero_def real_le)
1311      apply (simp add: not_le algebra_simps preal_self_less_add_left)
1312      done
1313    then show ?thesis
1314      by auto
1315  qed
1316qed
1317
1318lemma real_mult_less_mono2: "\<lbrakk>(0::real) < z; x < y\<rbrakk> \<Longrightarrow> z * x < z * y"
1319  by (metis add_uminus_conv_diff real_less_sum_gt_zero real_mult_order real_sum_gt_zero_less right_diff_distrib')
1320
1321
1322instance real :: linordered_field
1323proof
1324  fix x y z :: real
1325  show "x \<le> y \<Longrightarrow> z + x \<le> z + y" by (rule real_add_left_mono)
1326  show "\<bar>x\<bar> = (if x < 0 then -x else x)" by (simp only: real_abs_def)
1327  show "sgn x = (if x=0 then 0 else if 0<x then 1 else - 1)"
1328    by (simp only: real_sgn_def)
1329  show "z * x < z * y" if "x < y" "0 < z"
1330    by (simp add: real_mult_less_mono2 that)
1331qed
1332
1333
1334subsection \<open>Completeness of the reals\<close>
1335
1336text\<open>The function \<^term>\<open>real_of_preal\<close> requires many proofs, but it seems
1337to be essential for proving completeness of the reals from that of the
1338positive reals.\<close>
1339
1340lemma real_of_preal_add:
1341  "real_of_preal ((x::preal) + y) = real_of_preal x + real_of_preal y"
1342  by (simp add: real_of_preal_def real_add algebra_simps)
1343
1344lemma real_of_preal_mult:
1345  "real_of_preal ((x::preal) * y) = real_of_preal x * real_of_preal y"
1346  by (simp add: real_of_preal_def real_mult algebra_simps)
1347
1348text\<open>Gleason prop 9-4.4 p 127\<close>
1349lemma real_of_preal_trichotomy:
1350  "\<exists>m. (x::real) = real_of_preal m \<or> x = 0 \<or> x = -(real_of_preal m)"
1351proof (cases x)
1352  case (Abs_Real u v)
1353  show ?thesis
1354  proof (cases u v rule: linorder_cases)
1355    case less
1356    then show ?thesis
1357      using less_add_left
1358      apply (simp add: Abs_Real real_of_preal_def real_minus real_zero_def)
1359      by (metis preal_add_assoc preal_add_commute)      
1360  next
1361    case equal
1362    then show ?thesis
1363      using Abs_Real real_zero_iff by blast
1364  next
1365    case greater
1366    then show ?thesis
1367      using less_add_left
1368      apply (simp add: Abs_Real real_of_preal_def real_minus real_zero_def)
1369      by (metis preal_add_assoc preal_add_commute)      
1370  qed
1371qed
1372
1373lemma real_of_preal_less_iff [simp]:
1374  "(real_of_preal m1 < real_of_preal m2) = (m1 < m2)"
1375  by (metis not_less preal_add_less_cancel_right real_le real_of_preal_def)
1376
1377lemma real_of_preal_le_iff [simp]:
1378  "(real_of_preal m1 \<le> real_of_preal m2) = (m1 \<le> m2)"
1379  by (simp add: linorder_not_less [symmetric])
1380
1381lemma real_of_preal_zero_less [simp]: "0 < real_of_preal m"
1382  by (metis less_add_same_cancel2 preal_self_less_add_left real_of_preal_add real_of_preal_less_iff)
1383
1384
1385subsection\<open>Theorems About the Ordering\<close>
1386
1387lemma real_gt_zero_preal_Ex: "(0 < x) \<longleftrightarrow> (\<exists>y. x = real_of_preal y)"
1388  using order.asym real_of_preal_trichotomy by fastforce
1389
1390subsection \<open>Completeness of Positive Reals\<close>
1391
1392text \<open>
1393  Supremum property for the set of positive reals
1394
1395  Let \<open>P\<close> be a non-empty set of positive reals, with an upper
1396  bound \<open>y\<close>.  Then \<open>P\<close> has a least upper bound
1397  (written \<open>S\<close>).
1398
1399  FIXME: Can the premise be weakened to \<open>\<forall>x \<in> P. x\<le> y\<close>?
1400\<close>
1401
1402lemma posreal_complete:
1403  assumes positive_P: "\<forall>x \<in> P. (0::real) < x"
1404    and not_empty_P: "\<exists>x. x \<in> P"
1405    and upper_bound_Ex: "\<exists>y. \<forall>x \<in> P. x<y"
1406  shows "\<exists>s. \<forall>y. (\<exists>x \<in> P. y < x) = (y < s)"
1407proof (rule exI, rule allI)
1408  fix y
1409  let ?pP = "{w. real_of_preal w \<in> P}"
1410
1411  show "(\<exists>x\<in>P. y < x) = (y < real_of_preal (psup ?pP))"
1412  proof (cases "0 < y")
1413    assume neg_y: "\<not> 0 < y"
1414    show ?thesis
1415    proof
1416      assume "\<exists>x\<in>P. y < x"
1417      thus "y < real_of_preal (psup ?pP)"
1418        by (metis dual_order.strict_trans neg_y not_less_iff_gr_or_eq real_of_preal_zero_less) 
1419    next
1420      assume "y < real_of_preal (psup ?pP)"
1421      obtain "x" where x_in_P: "x \<in> P" using not_empty_P ..
1422      thus "\<exists>x \<in> P. y < x" using x_in_P
1423        using neg_y not_less_iff_gr_or_eq positive_P by fastforce 
1424    qed
1425  next
1426    assume pos_y: "0 < y"
1427    then obtain py where y_is_py: "y = real_of_preal py"
1428      by (auto simp: real_gt_zero_preal_Ex)
1429
1430    obtain a where "a \<in> P" using not_empty_P ..
1431    with positive_P have a_pos: "0 < a" ..
1432    then obtain pa where "a = real_of_preal pa"
1433      by (auto simp: real_gt_zero_preal_Ex)
1434    hence "pa \<in> ?pP" using \<open>a \<in> P\<close> by auto
1435    hence pP_not_empty: "?pP \<noteq> {}" by auto
1436
1437    obtain sup where sup: "\<forall>x \<in> P. x < sup"
1438      using upper_bound_Ex ..
1439    from this and \<open>a \<in> P\<close> have "a < sup" ..
1440    hence "0 < sup" using a_pos by arith
1441    then obtain possup where "sup = real_of_preal possup"
1442      by (auto simp: real_gt_zero_preal_Ex)
1443    hence "\<forall>X \<in> ?pP. X \<le> possup"
1444      using sup by auto
1445    with pP_not_empty have psup: "\<And>Z. (\<exists>X \<in> ?pP. Z < X) = (Z < psup ?pP)"
1446      by (meson preal_complete)
1447    show ?thesis
1448    proof
1449      assume "\<exists>x \<in> P. y < x"
1450      then obtain x where x_in_P: "x \<in> P" and y_less_x: "y < x" ..
1451      hence "0 < x" using pos_y by arith
1452      then obtain px where x_is_px: "x = real_of_preal px"
1453        by (auto simp: real_gt_zero_preal_Ex)
1454
1455      have py_less_X: "\<exists>X \<in> ?pP. py < X"
1456      proof
1457        show "py < px" using y_is_py and x_is_px and y_less_x
1458          by simp
1459        show "px \<in> ?pP" using x_in_P and x_is_px by simp
1460      qed
1461
1462      have "(\<exists>X \<in> ?pP. py < X) \<Longrightarrow> (py < psup ?pP)"
1463        using psup by simp
1464      hence "py < psup ?pP" using py_less_X by simp
1465      thus "y < real_of_preal (psup {w. real_of_preal w \<in> P})"
1466        using y_is_py and pos_y by simp
1467    next
1468      assume y_less_psup: "y < real_of_preal (psup ?pP)"
1469
1470      hence "py < psup ?pP" using y_is_py
1471        by simp
1472      then obtain "X" where py_less_X: "py < X" and X_in_pP: "X \<in> ?pP"
1473        using psup by auto
1474      then obtain x where x_is_X: "x = real_of_preal X"
1475        by (simp add: real_gt_zero_preal_Ex)
1476      hence "y < x" using py_less_X and y_is_py
1477        by simp
1478      moreover have "x \<in> P" 
1479        using x_is_X and X_in_pP by simp
1480      ultimately show "\<exists> x \<in> P. y < x" ..
1481    qed
1482  qed
1483qed
1484
1485
1486subsection \<open>Completeness\<close>
1487
1488lemma reals_complete:
1489  fixes S :: "real set"
1490  assumes notempty_S: "\<exists>X. X \<in> S"
1491    and exists_Ub: "bdd_above S"
1492  shows "\<exists>x. (\<forall>s\<in>S. s \<le> x) \<and> (\<forall>y. (\<forall>s\<in>S. s \<le> y) \<longrightarrow> x \<le> y)"
1493proof -
1494  obtain X where X_in_S: "X \<in> S" using notempty_S ..
1495  obtain Y where Y_isUb: "\<forall>s\<in>S. s \<le> Y"
1496    using exists_Ub by (auto simp: bdd_above_def)
1497  let ?SHIFT = "{z. \<exists>x \<in>S. z = x + (-X) + 1} \<inter> {x. 0 < x}"
1498
1499  {
1500    fix x
1501    assume S_le_x: "\<forall>s\<in>S. s \<le> x"
1502    {
1503      fix s
1504      assume "s \<in> {z. \<exists>x\<in>S. z = x + - X + 1}"
1505      hence "\<exists> x \<in> S. s = x + -X + 1" ..
1506      then obtain x1 where x1: "x1 \<in> S" "s = x1 + (-X) + 1" ..
1507      then have "x1 \<le> x" using S_le_x by simp
1508      with x1 have "s \<le> x + - X + 1" by arith
1509    }
1510    then have "\<forall>s\<in>?SHIFT. s \<le> x + (-X) + 1"
1511      by auto
1512  } note S_Ub_is_SHIFT_Ub = this
1513
1514  have *: "\<forall>s\<in>?SHIFT. s \<le> Y + (-X) + 1" using Y_isUb by (rule S_Ub_is_SHIFT_Ub)
1515  have "\<forall>s\<in>?SHIFT. s < Y + (-X) + 2"
1516  proof
1517    fix s assume "s\<in>?SHIFT"
1518    with * have "s \<le> Y + (-X) + 1" by simp
1519    also have "\<dots> < Y + (-X) + 2" by simp
1520    finally show "s < Y + (-X) + 2" .
1521  qed
1522  moreover have "\<forall>y \<in> ?SHIFT. 0 < y" by auto
1523  moreover have shifted_not_empty: "\<exists>u. u \<in> ?SHIFT"
1524    using X_in_S and Y_isUb by auto
1525  ultimately obtain t where t_is_Lub: "\<forall>y. (\<exists>x\<in>?SHIFT. y < x) = (y < t)"
1526    using posreal_complete [of ?SHIFT] unfolding bdd_above_def by blast
1527
1528  show ?thesis
1529  proof
1530    show "(\<forall>s\<in>S. s \<le> (t + X + (-1))) \<and> (\<forall>y. (\<forall>s\<in>S. s \<le> y) \<longrightarrow> (t + X + (-1)) \<le> y)"
1531    proof safe
1532      fix x
1533      assume "\<forall>s\<in>S. s \<le> x"
1534      hence "\<forall>s\<in>?SHIFT. s \<le> x + (-X) + 1"
1535        using S_Ub_is_SHIFT_Ub by simp
1536      then have "\<not> x + (-X) + 1 < t"
1537        by (subst t_is_Lub[rule_format, symmetric]) (simp add: not_less)
1538      thus "t + X + -1 \<le> x" by arith
1539    next
1540      fix y
1541      assume y_in_S: "y \<in> S"
1542      obtain "u" where u_in_shift: "u \<in> ?SHIFT" using shifted_not_empty ..
1543      hence "\<exists> x \<in> S. u = x + - X + 1" by simp
1544      then obtain "x" where x_and_u: "u = x + - X + 1" ..
1545      have u_le_t: "u \<le> t"
1546      proof (rule dense_le)
1547        fix x assume "x < u" then have "x < t"
1548          using u_in_shift t_is_Lub by auto
1549        then show "x \<le> t"  by simp
1550      qed
1551
1552      show "y \<le> t + X + -1"
1553      proof cases
1554        assume "y \<le> x"
1555        moreover have "x = u + X + - 1" using x_and_u by arith
1556        moreover have "u + X + - 1  \<le> t + X + -1" using u_le_t by arith
1557        ultimately show "y  \<le> t + X + -1" by arith
1558      next
1559        assume "~(y \<le> x)"
1560        hence x_less_y: "x < y" by arith
1561
1562        have "x + (-X) + 1 \<in> ?SHIFT" using x_and_u and u_in_shift by simp
1563        hence "0 < x + (-X) + 1" by simp
1564        hence "0 < y + (-X) + 1" using x_less_y by arith
1565        hence *: "y + (-X) + 1 \<in> ?SHIFT" using y_in_S by simp
1566        have "y + (-X) + 1 \<le> t"
1567        proof (rule dense_le)
1568          fix x assume "x < y + (-X) + 1" then have "x < t"
1569            using * t_is_Lub by auto
1570          then show "x \<le> t"  by simp
1571        qed
1572        thus ?thesis by simp
1573      qed
1574    qed
1575  qed
1576qed
1577
1578subsection \<open>The Archimedean Property of the Reals\<close>
1579
1580theorem reals_Archimedean:
1581  fixes x :: real
1582  assumes x_pos: "0 < x"
1583  shows "\<exists>n. inverse (of_nat (Suc n)) < x"
1584proof (rule ccontr)
1585  assume contr: "\<not> ?thesis"
1586  have "\<forall>n. x * of_nat (Suc n) \<le> 1"
1587  proof
1588    fix n
1589    from contr have "x \<le> inverse (of_nat (Suc n))"
1590      by (simp add: linorder_not_less)
1591    hence "x \<le> (1 / (of_nat (Suc n)))"
1592      by (simp add: inverse_eq_divide)
1593    moreover have "(0::real) \<le> of_nat (Suc n)"
1594      by (rule of_nat_0_le_iff)
1595    ultimately have "x * of_nat (Suc n) \<le> (1 / of_nat (Suc n)) * of_nat (Suc n)"
1596      by (rule mult_right_mono)
1597    thus "x * of_nat (Suc n) \<le> 1" by (simp del: of_nat_Suc)
1598  qed
1599  hence 2: "bdd_above {z. \<exists>n. z = x * (of_nat (Suc n))}"
1600    by (auto intro!: bdd_aboveI[of _ 1])
1601  have 1: "\<exists>X. X \<in> {z. \<exists>n. z = x* (of_nat (Suc n))}" by auto
1602  obtain t where
1603    upper: "\<And>z. z \<in> {z. \<exists>n. z = x * of_nat (Suc n)} \<Longrightarrow> z \<le> t" and
1604    least: "\<And>y. (\<And>a. a \<in> {z. \<exists>n. z = x * of_nat (Suc n)} \<Longrightarrow> a \<le> y) \<Longrightarrow> t \<le> y"
1605    using reals_complete[OF 1 2] by auto
1606
1607  have "t \<le> t + - x"
1608  proof (rule least)
1609    fix a assume a: "a \<in> {z. \<exists>n. z = x * (of_nat (Suc n))}"
1610    have "\<forall>n::nat. x * of_nat n \<le> t + - x"
1611    proof
1612      fix n
1613      have "x * of_nat (Suc n) \<le> t"
1614        by (simp add: upper)
1615      hence  "x * (of_nat n) + x \<le> t"
1616        by (simp add: distrib_left)
1617      thus  "x * (of_nat n) \<le> t + - x" by arith
1618    qed    hence "\<forall>m. x * of_nat (Suc m) \<le> t + - x" by (simp del: of_nat_Suc)
1619    with a show "a \<le> t + - x"
1620      by auto
1621  qed
1622  thus False using x_pos by arith
1623qed
1624
1625text \<open>
1626  There must be other proofs, e.g. \<open>Suc\<close> of the largest
1627  integer in the cut representing \<open>x\<close>.
1628\<close>
1629
1630lemma reals_Archimedean2: "\<exists>n. (x::real) < of_nat (n::nat)"
1631proof cases
1632  assume "x \<le> 0"
1633  hence "x < of_nat (1::nat)" by simp
1634  thus ?thesis ..
1635next
1636  assume "\<not> x \<le> 0"
1637  hence x_greater_zero: "0 < x" by simp
1638  hence "0 < inverse x" by simp
1639  then obtain n where "inverse (of_nat (Suc n)) < inverse x"
1640    using reals_Archimedean by blast
1641  hence "inverse (of_nat (Suc n)) * x < inverse x * x"
1642    using x_greater_zero by (rule mult_strict_right_mono)
1643  hence "inverse (of_nat (Suc n)) * x < 1"
1644    using x_greater_zero by simp
1645  hence "of_nat (Suc n) * (inverse (of_nat (Suc n)) * x) < of_nat (Suc n) * 1"
1646    by (rule mult_strict_left_mono) (simp del: of_nat_Suc)
1647  hence "x < of_nat (Suc n)"
1648    by (simp add: algebra_simps del: of_nat_Suc)
1649  thus "\<exists>(n::nat). x < of_nat n" ..
1650qed
1651
1652instance real :: archimedean_field
1653proof
1654  fix r :: real
1655  obtain n :: nat where "r < of_nat n"
1656    using reals_Archimedean2 ..
1657  then have "r \<le> of_int (int n)"
1658    by simp
1659  then show "\<exists>z. r \<le> of_int z" ..
1660qed
1661
1662end
1663