| author | paulson | 
| Wed, 05 Mar 2003 16:03:33 +0100 | |
| changeset 13846 | b2c494d76012 | 
| parent 12018 | ec054019c910 | 
| child 14268 | 5cf13e80be0e | 
| permissions | -rw-r--r-- | 
| 9428 | 1 | (* Title : HOL/Real/RComplete.ML | 
| 7219 | 2 | ID : $Id$ | 
| 5078 | 3 | Author : Jacques D. Fleuriot | 
| 4 | Copyright : 1998 University of Cambridge | |
| 9428 | 5 | |
| 6 | Completeness theorems for positive reals and reals. | |
| 5078 | 7 | *) | 
| 8 | ||
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 9 | Goal "x/2 + x/2 = (x::real)"; | 
| 10752 
c4f1bf2acf4c
tidying, and separation of HOL-Hyperreal from HOL-Real
 paulson parents: 
10677diff
changeset | 10 | by (Simp_tac 1); | 
| 
c4f1bf2acf4c
tidying, and separation of HOL-Hyperreal from HOL-Real
 paulson parents: 
10677diff
changeset | 11 | qed "real_sum_of_halves"; | 
| 
c4f1bf2acf4c
tidying, and separation of HOL-Hyperreal from HOL-Real
 paulson parents: 
10677diff
changeset | 12 | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 13 | (*--------------------------------------------------------- | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 14 | Completeness of reals: use supremum property of | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 15 | preal and theorems about real_preal. Theorems | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 16 | previously in Real.ML. | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 17 | ---------------------------------------------------------*) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 18 | (*a few lemmas*) | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 19 | Goal "ALL x:P. 0 < x ==> \ | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 20 | \ ((EX x:P. y < x) = (EX X. real_of_preal X : P & \ | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 21 | \ y < real_of_preal X))"; | 
| 9065 | 22 | by (blast_tac (claset() addSDs [bspec, | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 23 | real_gt_zero_preal_Ex RS iffD1]) 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 24 | qed "real_sup_lemma1"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 25 | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 26 | Goal "[| ALL x:P. 0 < x; a: P; ALL x: P. x < y |] \ | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 27 | \         ==> (EX X. X: {w. real_of_preal w : P}) & \
 | 
| 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 28 | \             (EX Y. ALL X: {w. real_of_preal w : P}. X < Y)";
 | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 29 | by (rtac conjI 1); | 
| 9065 | 30 | by (blast_tac (claset() addDs [bspec, | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 31 | real_gt_zero_preal_Ex RS iffD1]) 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 32 | by Auto_tac; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 33 | by (dtac bspec 1 THEN assume_tac 1); | 
| 7499 | 34 | by (ftac bspec 1 THEN assume_tac 1); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 35 | by (dtac order_less_trans 1 THEN assume_tac 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 36 | by (dtac (real_gt_zero_preal_Ex RS iffD1) 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 37 | by (Force_tac 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 38 | qed "real_sup_lemma2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 39 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 40 | (*------------------------------------------------------------- | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 41 | Completeness of Positive Reals | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 42 | -------------------------------------------------------------*) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 43 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 44 | (** | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 45 | Supremum property for the set of positive reals | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 46 | FIXME: long proof - should be improved | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 47 | **) | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 48 | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 49 | (*Let P be a non-empty set of positive reals, with an upper bound y. | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 50 | Then P has a least upper bound (written S).*) | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 51 | Goal "[| ALL x:P. (0::real) < x; EX x. x: P; EX y. ALL x: P. x<y |] \ | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 52 | \ ==> (EX S. ALL y. (EX x: P. y < x) = (y < S))"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 53 | by (res_inst_tac | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 54 |    [("x","real_of_preal (psup({w. real_of_preal w : P}))")] exI 1);
 | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 55 | by (Clarify_tac 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 56 | by (case_tac "0 < ya" 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 57 | by Auto_tac; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 58 | by (ftac real_sup_lemma2 1 THEN REPEAT (assume_tac 1)); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 59 | by (dtac (real_gt_zero_preal_Ex RS iffD1) 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 60 | by (dtac (real_less_all_real2) 3); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 61 | by Auto_tac; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 62 | by (rtac (preal_complete RS spec RS iffD1) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 63 | by Auto_tac; | 
| 7499 | 64 | by (ftac real_gt_preal_preal_Ex 1); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 65 | by (Force_tac 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 66 | (* second part *) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 67 | by (rtac (real_sup_lemma1 RS iffD2) 1 THEN assume_tac 1); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 68 | by (auto_tac (claset() addSDs [real_less_all_real2, | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 69 | real_gt_zero_preal_Ex RS iffD1], | 
| 9065 | 70 | simpset())); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 71 | by (ftac real_sup_lemma2 2 THEN REPEAT (assume_tac 1)); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 72 | by (ftac real_sup_lemma2 1 THEN REPEAT (assume_tac 1)); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 73 | by (rtac (preal_complete RS spec RS iffD2 RS bexE) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 74 | by (Blast_tac 3); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 75 | by (ALLGOALS(Blast_tac)); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 76 | qed "posreal_complete"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 77 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 78 | (*-------------------------------------------------------- | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 79 | Completeness properties using isUb, isLub etc. | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 80 | -------------------------------------------------------*) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 81 | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 82 | Goal "[| isLub R S x; isLub R S y |] ==> x = (y::real)"; | 
| 7499 | 83 | by (ftac isLub_isUb 1); | 
| 5078 | 84 | by (forw_inst_tac [("x","y")] isLub_isUb 1);
 | 
| 85 | by (blast_tac (claset() addSIs [real_le_anti_sym] | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 86 | addSDs [isLub_le_isUb]) 1); | 
| 5078 | 87 | qed "real_isLub_unique"; | 
| 88 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 89 | Goalw [setle_def,setge_def] "[| (x::real) <=* S'; S <= S' |] ==> x <=* S"; | 
| 5078 | 90 | by (Blast_tac 1); | 
| 91 | qed "real_order_restrict"; | |
| 92 | ||
| 93 | (*---------------------------------------------------------------- | |
| 94 | Completeness theorem for the positive reals(again) | |
| 95 | ----------------------------------------------------------------*) | |
| 96 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 97 | Goal "[| ALL x: S. 0 < x; \ | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 98 | \ EX x. x: S; \ | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 99 | \ EX u. isUb (UNIV::real set) S u \ | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 100 | \ |] ==> EX t. isLub (UNIV::real set) S t"; | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 101 | by (res_inst_tac | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 102 |     [("x","real_of_preal(psup({w. real_of_preal w : S}))")] exI 1);
 | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 103 | by (auto_tac (claset(), simpset() addsimps | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 104 | [isLub_def,leastP_def,isUb_def])); | 
| 5078 | 105 | by (auto_tac (claset() addSIs [setleI,setgeI] | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 106 | addSDs [(real_gt_zero_preal_Ex) RS iffD1], | 
| 9065 | 107 | simpset())); | 
| 5078 | 108 | by (forw_inst_tac [("x","y")] bspec 1 THEN assume_tac 1);
 | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 109 | by (dtac ((real_gt_zero_preal_Ex) RS iffD1) 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 110 | by (auto_tac (claset(), simpset() addsimps [real_of_preal_le_iff])); | 
| 5078 | 111 | by (rtac preal_psup_leI2a 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 112 | by (forw_inst_tac [("y","real_of_preal ya")] setleD 1 THEN assume_tac 1);
 | 
| 7499 | 113 | by (ftac real_ge_preal_preal_Ex 1); | 
| 5078 | 114 | by (Step_tac 1); | 
| 115 | by (res_inst_tac [("x","y")] exI 1);
 | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 116 | by (blast_tac (claset() addSDs [setleD] addIs [real_of_preal_le_iff RS iffD1]) 1); | 
| 5078 | 117 | by (forw_inst_tac [("x","x")] bspec 1 THEN assume_tac 1);
 | 
| 7499 | 118 | by (ftac isUbD2 1); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 119 | by (dtac ((real_gt_zero_preal_Ex) RS iffD1) 1); | 
| 5588 | 120 | by (auto_tac (claset() addSDs [isUbD, real_ge_preal_preal_Ex], | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 121 | simpset() addsimps [real_of_preal_le_iff])); | 
| 5588 | 122 | by (blast_tac (claset() addSDs [setleD] addSIs [psup_le_ub1] | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 123 | addIs [real_of_preal_le_iff RS iffD1]) 1); | 
| 5078 | 124 | qed "posreals_complete"; | 
| 125 | ||
| 126 | ||
| 127 | (*------------------------------- | |
| 128 | Lemmas | |
| 129 | -------------------------------*) | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 130 | Goal "ALL y : {z. EX x: P. z = x + (-xa) + 1} Int {x. 0 < x}. 0 < y";
 | 
| 5078 | 131 | by Auto_tac; | 
| 132 | qed "real_sup_lemma3"; | |
| 133 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 134 | Goal "(xa <= S + X + (-Z)) = (xa + (-X) + Z <= (S::real))"; | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 135 | by Auto_tac; | 
| 5078 | 136 | qed "lemma_le_swap2"; | 
| 137 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 138 | Goal "[| (x::real) + (-X) + 1 <= S; xa <= x |] ==> xa <= S + X + (- 1)"; | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 139 | by (arith_tac 1); | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 140 | by Auto_tac; | 
| 5078 | 141 | qed "lemma_real_complete2b"; | 
| 142 | ||
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 143 | (*---------------------------------------------------------- | 
| 5078 | 144 | reals Completeness (again!) | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 145 | ----------------------------------------------------------*) | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 146 | Goal "[| EX X. X: S; EX Y. isUb (UNIV::real set) S Y |] \ | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 147 | \ ==> EX t. isLub (UNIV :: real set) S t"; | 
| 5078 | 148 | by (Step_tac 1); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 149 | by (subgoal_tac "EX u. u: {z. EX x: S. z = x + (-X) + 1} \
 | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 150 | \                Int {x. 0 < x}" 1);
 | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 151 | by (subgoal_tac "isUb (UNIV::real set) ({z. EX x: S. z = x + (-X) + 1} \
 | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 152 | \                Int {x. 0 < x})  (Y + (-X) + 1)" 1); 
 | 
| 5078 | 153 | by (cut_inst_tac [("P","S"),("xa","X")] real_sup_lemma3 1);
 | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 154 | by (EVERY1[forward_tac [exI RSN (3,posreals_complete)], Blast_tac, Blast_tac, | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 155 | Step_tac]); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 156 | by (res_inst_tac [("x","t + X + (- 1)")] exI 1);
 | 
| 5078 | 157 | by (rtac isLubI2 1); | 
| 158 | by (rtac setgeI 2 THEN Step_tac 2); | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 159 | by (subgoal_tac "isUb (UNIV:: real set) ({z. EX x: S. z = x + (-X) + 1} \
 | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 160 | \                Int {x. 0 < x})  (y + (-X) + 1)" 2); 
 | 
| 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 161 | by (dres_inst_tac [("y","(y + (- X) + 1)")] isLub_le_isUb 2 
 | 
| 5078 | 162 | THEN assume_tac 2); | 
| 5588 | 163 | by (full_simp_tac | 
| 164 | (simpset() addsimps [real_diff_def, real_diff_le_eq RS sym] @ | |
| 165 | real_add_ac) 2); | |
| 5078 | 166 | by (rtac (setleI RS isUbI) 1); | 
| 167 | by (Step_tac 1); | |
| 168 | by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
 | |
| 169 | by (stac lemma_le_swap2 1); | |
| 7499 | 170 | by (ftac isLubD2 1 THEN assume_tac 2); | 
| 5078 | 171 | by (Step_tac 1); | 
| 172 | by (Blast_tac 1); | |
| 10677 | 173 | by (arith_tac 1); | 
| 5078 | 174 | by (stac lemma_le_swap2 1); | 
| 7499 | 175 | by (ftac isLubD2 1 THEN assume_tac 2); | 
| 5078 | 176 | by (Blast_tac 1); | 
| 177 | by (rtac lemma_real_complete2b 1); | |
| 10752 
c4f1bf2acf4c
tidying, and separation of HOL-Hyperreal from HOL-Real
 paulson parents: 
10677diff
changeset | 178 | by (etac order_less_imp_le 2); | 
| 5078 | 179 | by (blast_tac (claset() addSIs [isLubD2]) 1 THEN Step_tac 1); | 
| 5588 | 180 | by (full_simp_tac (simpset() addsimps [real_add_assoc]) 1); | 
| 181 | by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI] | |
| 182 | addIs [real_add_le_mono1]) 1); | |
| 183 | by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI] | |
| 184 | addIs [real_add_le_mono1]) 1); | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
7583diff
changeset | 185 | by (Auto_tac); | 
| 5078 | 186 | qed "reals_complete"; | 
| 187 | ||
| 188 | (*---------------------------------------------------------------- | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 189 | Related: Archimedean property of reals | 
| 5078 | 190 | ----------------------------------------------------------------*) | 
| 191 | ||
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 192 | Goal "0 < real (Suc n)"; | 
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10784diff
changeset | 193 | by (res_inst_tac [("y","real n")] order_le_less_trans 1); 
 | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 194 | by (rtac (real_of_nat_ge_zero) 1); | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 195 | by (simp_tac (simpset() addsimps [real_of_nat_Suc]) 1); | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 196 | qed "real_of_nat_Suc_gt_zero"; | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 197 | |
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 198 | Goal "0 < x ==> EX n. inverse (real(Suc n)) < x"; | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 199 | by (rtac ccontr 1); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 200 | by (subgoal_tac "ALL n. x * real (Suc n) <= 1" 1); | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 201 | by (asm_full_simp_tac | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 202 | (simpset() addsimps [linorder_not_less, real_inverse_eq_divide]) 2); | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 203 | by (Clarify_tac 2); | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 204 | by (dres_inst_tac [("x","n")] spec 2); 
 | 
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10784diff
changeset | 205 | by (dres_inst_tac [("k","real (Suc n)")] (real_mult_le_mono1) 2); 
 | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 206 | by (rtac real_of_nat_ge_zero 2); | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 207 | by (asm_full_simp_tac (simpset() | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 208 | addsimps [real_of_nat_Suc_gt_zero RS real_not_refl2 RS not_sym, | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 209 | real_mult_commute]) 2); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 210 | by (subgoal_tac "isUb (UNIV::real set) \ | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 211 | \                     {z. EX n. z = x*(real (Suc n))} 1" 1);
 | 
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10784diff
changeset | 212 | by (subgoal_tac "EX X. X : {z. EX n. z = x*(real (Suc n))}" 1);
 | 
| 5078 | 213 | by (dtac reals_complete 1); | 
| 214 | by (auto_tac (claset() addIs [isUbI,setleI],simpset())); | |
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10784diff
changeset | 215 | by (subgoal_tac "ALL m. x*(real(Suc m)) <= t" 1); | 
| 5078 | 216 | by (asm_full_simp_tac (simpset() addsimps | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 217 | [real_of_nat_Suc, real_add_mult_distrib2]) 1); | 
| 5078 | 218 | by (blast_tac (claset() addIs [isLubD2]) 2); | 
| 5588 | 219 | by (asm_full_simp_tac | 
| 220 | (simpset() addsimps [real_le_diff_eq RS sym, real_diff_def]) 1); | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 221 | by (subgoal_tac "isUb (UNIV::real set) \ | 
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10784diff
changeset | 222 | \                  {z. EX n. z = x*(real (Suc n))} (t + (-x))" 1);
 | 
| 5078 | 223 | by (blast_tac (claset() addSIs [isUbI,setleI]) 2); | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 224 | by (dres_inst_tac [("y","t+(-x)")] isLub_le_isUb 1);
 | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 225 | by (auto_tac (claset(), | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 226 | simpset() addsimps [real_of_nat_Suc,real_add_mult_distrib2])); | 
| 5078 | 227 | qed "reals_Archimedean"; | 
| 228 | ||
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 229 | (*There must be other proofs, e.g. Suc of the largest integer in the | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 230 | cut representing x*) | 
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10784diff
changeset | 231 | Goal "EX n. (x::real) < real (n::nat)"; | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 232 | by (res_inst_tac [("R1.0","x"),("R2.0","0")] real_linear_less2 1);
 | 
| 5078 | 233 | by (res_inst_tac [("x","0")] exI 1);
 | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 234 | by (res_inst_tac [("x","1")] exI 2);
 | 
| 10752 
c4f1bf2acf4c
tidying, and separation of HOL-Hyperreal from HOL-Real
 paulson parents: 
10677diff
changeset | 235 | by (auto_tac (claset() addEs [order_less_trans], | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 236 | simpset() addsimps [real_of_nat_one])); | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 237 | by (ftac (real_inverse_gt_0 RS reals_Archimedean) 1); | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 238 | by (Step_tac 1 THEN res_inst_tac [("x","Suc n")] exI 1);
 | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 239 | by (forw_inst_tac [("y","inverse x")] real_mult_less_mono1 1);
 | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 240 | by Auto_tac; | 
| 12018 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 paulson parents: 
11704diff
changeset | 241 | by (dres_inst_tac [("y","1"),("z","real (Suc n)")]
 | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 242 | (rotate_prems 1 real_mult_less_mono2) 1); | 
| 5588 | 243 | by (auto_tac (claset(), | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10752diff
changeset | 244 | simpset() addsimps [real_of_nat_Suc_gt_zero, | 
| 5588 | 245 | real_not_refl2 RS not_sym, | 
| 246 | real_mult_assoc RS sym])); | |
| 5078 | 247 | qed "reals_Archimedean2"; |