src/HOL/ex/Dedekind_Real.thy
author Thomas Sewell <thomas.sewell@nicta.com.au>
Wed, 11 Jun 2014 14:24:23 +1000
changeset 58834 74bf65a1910a
parent 57886 b60d5d119489
child 58854 cc97b347b301
permissions -rw-r--r--
Hypsubst preserves equality hypotheses

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