src/HOL/Real/RComplete.thy
author paulson
Tue, 27 Jan 2004 15:39:51 +0100
changeset 14365 3d4df8c166ae
parent 9429 8ebc549e9326
child 14387 e96d5c42c4b0
permissions -rw-r--r--
replacing HOL/Real/PRat, PNat by the rational number development of Markus Wenzel
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     1
(*  Title       : RComplete.thy
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 5078
diff changeset
     2
    ID          : $Id$
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     5
    Description : Completeness theorems for positive
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     6
                  reals and reals 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     7
*) 
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
     8
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
     9
header{*Completeness Theorems for Positive Reals and Reals.*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    10
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    11
theory RComplete = Lubs + RealArith:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    12
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    13
lemma real_sum_of_halves: "x/2 + x/2 = (x::real)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    14
apply (simp)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    15
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    16
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    17
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    18
subsection{*Completeness of Reals by Supremum Property of type @{typ preal}*} 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    19
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    20
 (*a few lemmas*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    21
lemma real_sup_lemma1:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    22
     "\<forall>x \<in> P. 0 < x ==>   
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    23
      ((\<exists>x \<in> P. y < x) = (\<exists>X. real_of_preal X \<in> P & y < real_of_preal X))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    24
by (blast dest!: bspec real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    25
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    26
lemma real_sup_lemma2:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    27
     "[| \<forall>x \<in> P. 0 < x;  a \<in> P;   \<forall>x \<in> P. x < y |]  
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    28
      ==> (\<exists>X. X\<in> {w. real_of_preal w \<in> P}) &  
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    29
          (\<exists>Y. \<forall>X\<in> {w. real_of_preal w \<in> P}. X < Y)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    30
apply (rule conjI)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    31
apply (blast dest: bspec real_gt_zero_preal_Ex [THEN iffD1], auto)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    32
apply (drule bspec, assumption)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    33
apply (frule bspec, assumption)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    34
apply (drule order_less_trans, assumption)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    35
apply (drule real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    36
apply (force) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    37
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    38
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    39
(*-------------------------------------------------------------
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    40
            Completeness of Positive Reals
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    41
 -------------------------------------------------------------*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    42
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    43
(**
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    44
 Supremum property for the set of positive reals
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    45
 FIXME: long proof - should be improved
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    46
**)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    47
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    48
(*Let P be a non-empty set of positive reals, with an upper bound y.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    49
  Then P has a least upper bound (written S).  
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    50
FIXME: Can the premise be weakened to \<forall>x \<in> P. x\<le> y ??*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    51
lemma posreal_complete: "[| \<forall>x \<in> P. (0::real) < x;  \<exists>x. x \<in> P;  \<exists>y. \<forall>x \<in> P. x<y |]  
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    52
      ==> (\<exists>S. \<forall>y. (\<exists>x \<in> P. y < x) = (y < S))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    53
apply (rule_tac x = "real_of_preal (psup ({w. real_of_preal w \<in> P}))" in exI)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    54
apply clarify
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    55
apply (case_tac "0 < ya", auto)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    56
apply (frule real_sup_lemma2, assumption+)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    57
apply (drule real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    58
apply (drule_tac [3] real_less_all_real2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    59
apply (auto)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    60
apply (rule preal_complete [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    61
apply (auto intro: order_less_imp_le)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    62
apply (frule real_gt_preal_preal_Ex)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    63
apply (force)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    64
(* second part *)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    65
apply (rule real_sup_lemma1 [THEN iffD2], assumption)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    66
apply (auto dest!: real_less_all_real2 real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    67
apply (frule_tac [2] real_sup_lemma2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    68
apply (frule real_sup_lemma2, assumption+, clarify) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    69
apply (rule preal_complete [THEN iffD2, THEN bexE])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    70
prefer 3 apply blast
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    71
apply (blast intro!: order_less_imp_le)+
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    72
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    73
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    74
(*--------------------------------------------------------
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    75
   Completeness properties using isUb, isLub etc.
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    76
 -------------------------------------------------------*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    77
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    78
lemma real_isLub_unique: "[| isLub R S x; isLub R S y |] ==> x = (y::real)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    79
apply (frule isLub_isUb)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    80
apply (frule_tac x = y in isLub_isUb)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    81
apply (blast intro!: real_le_anti_sym dest!: isLub_le_isUb)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    82
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    83
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    84
lemma real_order_restrict: "[| (x::real) <=* S'; S <= S' |] ==> x <=* S"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    85
by (unfold setle_def setge_def, blast)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    86
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    87
(*----------------------------------------------------------------
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    88
           Completeness theorem for the positive reals(again)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    89
 ----------------------------------------------------------------*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    90
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    91
lemma posreals_complete:
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    92
     "[| \<forall>x \<in>S. 0 < x;  
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    93
         \<exists>x. x \<in>S;  
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    94
         \<exists>u. isUb (UNIV::real set) S u  
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    95
      |] ==> \<exists>t. isLub (UNIV::real set) S t"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    96
apply (rule_tac x = "real_of_preal (psup ({w. real_of_preal w \<in> S}))" in exI)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    97
apply (auto simp add: isLub_def leastP_def isUb_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    98
apply (auto intro!: setleI setgeI dest!: real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
    99
apply (frule_tac x = y in bspec, assumption)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   100
apply (drule real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   101
apply (auto simp add: real_of_preal_le_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   102
apply (frule_tac y = "real_of_preal ya" in setleD, assumption)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   103
apply (frule real_ge_preal_preal_Ex, safe)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   104
apply (blast intro!: preal_psup_le dest!: setleD intro: real_of_preal_le_iff [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   105
apply (frule_tac x = x in bspec, assumption)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   106
apply (frule isUbD2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   107
apply (drule real_gt_zero_preal_Ex [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   108
apply (auto dest!: isUbD real_ge_preal_preal_Ex simp add: real_of_preal_le_iff)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   109
apply (blast dest!: setleD intro!: psup_le_ub intro: real_of_preal_le_iff [THEN iffD1])
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   110
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   111
5078
7b5ea59c0275 Installation of target HOL-Real
paulson
parents:
diff changeset
   112
14365
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   113
(*-------------------------------
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   114
    Lemmas
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   115
 -------------------------------*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   116
lemma real_sup_lemma3: "\<forall>y \<in> {z. \<exists>x \<in> P. z = x + (-xa) + 1} Int {x. 0 < x}. 0 < y"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   117
by auto
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   118
 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   119
lemma lemma_le_swap2: "(xa <= S + X + (-Z)) = (xa + (-X) + Z <= (S::real))"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   120
by auto
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   121
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   122
lemma lemma_real_complete2b: "[| (x::real) + (-X) + 1 <= S; xa <= x |] ==> xa <= S + X + (- 1)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   123
by arith
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   124
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   125
(*----------------------------------------------------------
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   126
      reals Completeness (again!)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   127
 ----------------------------------------------------------*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   128
lemma reals_complete: "[| \<exists>X. X \<in>S;  \<exists>Y. isUb (UNIV::real set) S Y |]   
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   129
      ==> \<exists>t. isLub (UNIV :: real set) S t"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   130
apply safe
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   131
apply (subgoal_tac "\<exists>u. u\<in> {z. \<exists>x \<in>S. z = x + (-X) + 1} Int {x. 0 < x}")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   132
apply (subgoal_tac "isUb (UNIV::real set) ({z. \<exists>x \<in>S. z = x + (-X) + 1} Int {x. 0 < x}) (Y + (-X) + 1) ")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   133
apply (cut_tac P = S and xa = X in real_sup_lemma3)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   134
apply (frule posreals_complete [OF _ _ exI], blast, blast) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   135
apply safe
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   136
apply (rule_tac x = "t + X + (- 1) " in exI)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   137
apply (rule isLubI2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   138
apply (rule_tac [2] setgeI, safe)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   139
apply (subgoal_tac [2] "isUb (UNIV:: real set) ({z. \<exists>x \<in>S. z = x + (-X) + 1} Int {x. 0 < x}) (y + (-X) + 1) ")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   140
apply (drule_tac [2] y = " (y + (- X) + 1) " in isLub_le_isUb)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   141
 prefer 2 apply assumption
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   142
 prefer 2
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   143
apply arith
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   144
apply (rule setleI [THEN isUbI], safe)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   145
apply (rule_tac x = x and y = y in linorder_cases)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   146
apply (subst lemma_le_swap2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   147
apply (frule isLubD2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   148
 prefer 2 apply assumption
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   149
apply safe
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   150
apply blast
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   151
apply arith
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   152
apply (subst lemma_le_swap2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   153
apply (frule isLubD2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   154
 prefer 2 apply assumption
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   155
apply blast
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   156
apply (rule lemma_real_complete2b)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   157
apply (erule_tac [2] order_less_imp_le)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   158
apply (blast intro!: isLubD2, blast) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   159
apply (simp (no_asm_use) add: real_add_assoc)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   160
apply (blast dest: isUbD intro!: setleI [THEN isUbI] intro: add_right_mono)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   161
apply (blast dest: isUbD intro!: setleI [THEN isUbI] intro: add_right_mono, auto)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   162
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   163
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   164
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   165
subsection{*Corollary: the Archimedean Property of the Reals*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   166
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   167
lemma reals_Archimedean: "0 < x ==> \<exists>n. inverse (real(Suc n)) < x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   168
apply (rule ccontr)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   169
apply (subgoal_tac "\<forall>n. x * real (Suc n) <= 1")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   170
 prefer 2
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   171
apply (simp add: linorder_not_less inverse_eq_divide, clarify) 
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   172
apply (drule_tac x = n in spec)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   173
apply (drule_tac c = "real (Suc n)"  in mult_right_mono)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   174
apply (rule real_of_nat_ge_zero)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   175
apply (simp add: real_of_nat_Suc_gt_zero [THEN real_not_refl2, THEN not_sym] real_mult_commute)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   176
apply (subgoal_tac "isUb (UNIV::real set) {z. \<exists>n. z = x* (real (Suc n))} 1")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   177
apply (subgoal_tac "\<exists>X. X \<in> {z. \<exists>n. z = x* (real (Suc n))}")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   178
apply (drule reals_complete)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   179
apply (auto intro: isUbI setleI)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   180
apply (subgoal_tac "\<forall>m. x* (real (Suc m)) <= t")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   181
apply (simp add: real_of_nat_Suc right_distrib)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   182
prefer 2 apply (blast intro: isLubD2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   183
apply (simp add: le_diff_eq [symmetric] real_diff_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   184
apply (subgoal_tac "isUb (UNIV::real set) {z. \<exists>n. z = x* (real (Suc n))} (t + (-x))")
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   185
prefer 2 apply (blast intro!: isUbI setleI)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   186
apply (drule_tac y = "t+ (-x) " in isLub_le_isUb)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   187
apply (auto simp add: real_of_nat_Suc right_distrib)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   188
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   189
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   190
(*There must be other proofs, e.g. Suc of the largest integer in the
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   191
  cut representing x*)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   192
lemma reals_Archimedean2: "\<exists>n. (x::real) < real (n::nat)"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   193
apply (rule_tac x = x and y = 0 in linorder_cases)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   194
apply (rule_tac x = 0 in exI)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   195
apply (rule_tac [2] x = 1 in exI)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   196
apply (auto elim: order_less_trans simp add: real_of_nat_one)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   197
apply (frule positive_imp_inverse_positive [THEN reals_Archimedean], safe)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   198
apply (rule_tac x = "Suc n" in exI)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   199
apply (frule_tac b = "inverse x" in mult_strict_right_mono, auto)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   200
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   201
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   202
lemma reals_Archimedean3: "0 < x ==> \<forall>y. \<exists>(n::nat). y < real n * x"
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   203
apply safe
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   204
apply (cut_tac x = "y*inverse (x) " in reals_Archimedean2)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   205
apply safe
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   206
apply (frule_tac a = "y * inverse x" in mult_strict_right_mono)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   207
apply (auto simp add: mult_assoc real_of_nat_def)
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   208
done
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   209
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   210
ML
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   211
{*
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   212
val real_sum_of_halves = thm "real_sum_of_halves";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   213
val posreal_complete = thm "posreal_complete";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   214
val real_isLub_unique = thm "real_isLub_unique";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   215
val real_order_restrict = thm "real_order_restrict";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   216
val posreals_complete = thm "posreals_complete";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   217
val reals_complete = thm "reals_complete";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   218
val reals_Archimedean = thm "reals_Archimedean";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   219
val reals_Archimedean2 = thm "reals_Archimedean2";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   220
val reals_Archimedean3 = thm "reals_Archimedean3";
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   221
*}
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   222
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   223
end
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   224
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   225
3d4df8c166ae replacing HOL/Real/PRat, PNat by the rational number development
paulson
parents: 9429
diff changeset
   226