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