| author | paulson | 
| Sat, 04 Mar 2000 12:02:41 +0100 | |
| changeset 8340 | c169cd21fe42 | 
| parent 7583 | d1b40e0464b1 | 
| child 9013 | 9dd0274f76af | 
| permissions | -rw-r--r-- | 
| 5078 | 1 | (* Title : RComplete.thy | 
| 7219 | 2 | ID : $Id$ | 
| 5078 | 3 | Author : Jacques D. Fleuriot | 
| 4 | Copyright : 1998 University of Cambridge | |
| 5 | Description : Completeness theorems for positive | |
| 6 | reals and reals | |
| 7 | *) | |
| 8 | ||
| 5521 | 9 | claset_ref() := claset() delWrapper "bspec"; | 
| 10 | ||
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 11 | (*--------------------------------------------------------- | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 12 | Completeness of reals: use supremum property of | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 13 | preal and theorems about real_preal. Theorems | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 14 | previously in Real.ML. | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 15 | ---------------------------------------------------------*) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 16 | (*a few lemmas*) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 17 | Goal "! x:P. 0r < x ==> \ | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 18 | \ ((? x:P. y < x) = (? X. real_of_preal X : P & \ | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 19 | \ y < real_of_preal X))"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 20 | by (blast_tac (claset() addSDs [bspec, | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 21 | real_gt_zero_preal_Ex RS iffD1]) 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 22 | qed "real_sup_lemma1"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 23 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 24 | Goal "[| ! x:P. 0r < x; ? x. x: P; ? y. !x: P. x < y |] \ | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 25 | \         ==> (? X. X: {w. real_of_preal w : P}) & \
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 26 | \             (? Y. !X: {w. real_of_preal w : P}. X < Y)";
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 27 | by (rtac conjI 1); | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 28 | by (blast_tac (claset() addDs [bspec, real_gt_zero_preal_Ex RS iffD1]) 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 29 | by Auto_tac; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 30 | by (dtac bspec 1 THEN assume_tac 1); | 
| 7499 | 31 | by (ftac bspec 1 THEN assume_tac 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 32 | by (dtac real_less_trans 1 THEN assume_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 33 | by (dtac (real_gt_zero_preal_Ex RS iffD1) 1 THEN etac exE 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 34 | by (res_inst_tac [("x","ya")] exI 1);
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 35 | by Auto_tac; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 36 | by (dres_inst_tac [("x","real_of_preal X")] bspec 1 THEN assume_tac 1);
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 37 | by (etac real_of_preal_lessD 1); | 
| 
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 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 44 | (* Supremum property for the set of positive reals *) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 45 | (* FIXME: long proof - can be improved - need only have | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 46 | one case split *) (* will do for now *) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 47 | Goal "[| ! x:P. 0r < x; ? x. x: P; ? y. !x: P. x < y |] \ | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 48 | \ ==> (? S. !y. (? x: P. y < x) = (y < S))"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 49 | by (res_inst_tac | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 50 |    [("x","real_of_preal (psup({w. real_of_preal w : P}))")] exI 1);
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 51 | by Auto_tac; | 
| 7499 | 52 | by (ftac real_sup_lemma2 1 THEN Auto_tac); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 53 | by (case_tac "0r < ya" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 54 | by (dtac (real_gt_zero_preal_Ex RS iffD1) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 55 | by (dtac real_less_all_real2 2); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 56 | by Auto_tac; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 57 | by (rtac (preal_complete RS spec RS iffD1) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 58 | by Auto_tac; | 
| 7499 | 59 | by (ftac real_gt_preal_preal_Ex 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 60 | by Auto_tac; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 61 | (* second part *) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 62 | by (rtac (real_sup_lemma1 RS iffD2) 1 THEN assume_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 63 | by (case_tac "0r < ya" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 64 | by (auto_tac (claset() addSDs [real_less_all_real2, | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 65 | real_gt_zero_preal_Ex RS iffD1],simpset())); | 
| 7499 | 66 | by (ftac real_sup_lemma2 2 THEN Auto_tac); | 
| 67 | by (ftac real_sup_lemma2 1 THEN Auto_tac); | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 68 | by (rtac (preal_complete RS spec RS iffD2 RS bexE) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 69 | by (Blast_tac 3); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 70 | by (Blast_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 71 | by (Blast_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 72 | by (Blast_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 73 | qed "posreal_complete"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 74 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 75 | (*-------------------------------------------------------- | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 76 | Completeness properties using isUb, isLub etc. | 
| 
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 | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 79 | Goal "[| isLub R S x; isLub R S y |] ==> x = (y::real)"; | 
| 7499 | 80 | by (ftac isLub_isUb 1); | 
| 5078 | 81 | by (forw_inst_tac [("x","y")] isLub_isUb 1);
 | 
| 82 | 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 | 83 | addSDs [isLub_le_isUb]) 1); | 
| 5078 | 84 | qed "real_isLub_unique"; | 
| 85 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 86 | Goalw [setle_def,setge_def] "[| (x::real) <=* S'; S <= S' |] ==> x <=* S"; | 
| 5078 | 87 | by (Blast_tac 1); | 
| 88 | qed "real_order_restrict"; | |
| 89 | ||
| 90 | (*---------------------------------------------------------------- | |
| 91 | Completeness theorem for the positive reals(again) | |
| 92 | ----------------------------------------------------------------*) | |
| 93 | ||
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5078diff
changeset | 94 | Goal "[| ALL x: S. 0r < x; \ | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 95 | \ EX x. x: S; \ | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 96 | \ 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 | 97 | \ |] ==> EX t. isLub (UNIV::real set) S t"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 98 | by (res_inst_tac [("x","real_of_preal(psup({w. real_of_preal w : S}))")] exI 1);
 | 
| 5588 | 99 | by (auto_tac (claset(), simpset() addsimps [isLub_def,leastP_def,isUb_def])); | 
| 5078 | 100 | by (auto_tac (claset() addSIs [setleI,setgeI] | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 101 | addSDs [real_gt_zero_preal_Ex RS iffD1], | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 102 | simpset())); | 
| 5078 | 103 | by (forw_inst_tac [("x","y")] bspec 1 THEN assume_tac 1);
 | 
| 104 | by (dtac (real_gt_zero_preal_Ex RS iffD1) 1); | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 105 | by (auto_tac (claset(), simpset() addsimps [real_of_preal_le_iff])); | 
| 5078 | 106 | by (rtac preal_psup_leI2a 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 107 | by (forw_inst_tac [("y","real_of_preal ya")] setleD 1 THEN assume_tac 1);
 | 
| 7499 | 108 | by (ftac real_ge_preal_preal_Ex 1); | 
| 5078 | 109 | by (Step_tac 1); | 
| 110 | by (res_inst_tac [("x","y")] exI 1);
 | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 111 | by (blast_tac (claset() addSDs [setleD] addIs [real_of_preal_le_iff RS iffD1]) 1); | 
| 5078 | 112 | by (forw_inst_tac [("x","x")] bspec 1 THEN assume_tac 1);
 | 
| 7499 | 113 | by (ftac isUbD2 1); | 
| 5078 | 114 | by (dtac (real_gt_zero_preal_Ex RS iffD1) 1); | 
| 5588 | 115 | 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 | 116 | simpset() addsimps [real_of_preal_le_iff])); | 
| 5588 | 117 | by (blast_tac (claset() addSDs [setleD] addSIs [psup_le_ub1] | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 118 | addIs [real_of_preal_le_iff RS iffD1]) 1); | 
| 5078 | 119 | qed "posreals_complete"; | 
| 120 | ||
| 121 | ||
| 122 | (*------------------------------- | |
| 123 | Lemmas | |
| 124 | -------------------------------*) | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 125 | Goal "! y : {z. ? x: P. z = x + (-xa) + 1r} Int {x. 0r < x}. 0r < y";
 | 
| 5078 | 126 | by Auto_tac; | 
| 127 | qed "real_sup_lemma3"; | |
| 128 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 129 | Goal "(xa <= S + X + (-Z)) = (xa + (-X) + Z <= (S::real))"; | 
| 5588 | 130 | by (simp_tac (simpset() addsimps [real_diff_def, real_diff_le_eq RS sym] @ | 
| 131 | real_add_ac) 1); | |
| 5078 | 132 | qed "lemma_le_swap2"; | 
| 133 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 134 | Goal "[| 0r < x + (-X) + 1r; x < xa |] ==> 0r < xa + (-X) + 1r"; | 
| 5078 | 135 | by (dtac real_add_less_mono 1); | 
| 136 | by (assume_tac 1); | |
| 5588 | 137 | by (dres_inst_tac [("C","-x"),("A","0r + x")] real_add_less_mono2 1);
 | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 138 | by (asm_full_simp_tac | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 139 | (simpset() addsimps [real_add_zero_right, real_add_assoc RS sym, | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 140 | real_add_minus_left,real_add_zero_left]) 1); | 
| 5078 | 141 | qed "lemma_real_complete1"; | 
| 142 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 143 | Goal "[| x + (-X) + 1r <= S; xa < x |] ==> xa + (-X) + 1r <= S"; | 
| 5078 | 144 | by (dtac real_less_imp_le 1); | 
| 145 | by (dtac real_add_le_mono 1); | |
| 146 | by (assume_tac 1); | |
| 147 | by (asm_full_simp_tac (simpset() addsimps real_add_ac) 1); | |
| 148 | qed "lemma_real_complete2"; | |
| 149 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 150 | Goal "[| x + (-X) + 1r <= S; xa < x |] ==> xa <= S + X + (-1r)"; (**) | 
| 5078 | 151 | by (rtac (lemma_le_swap2 RS iffD2) 1); | 
| 152 | by (etac lemma_real_complete2 1); | |
| 153 | by (assume_tac 1); | |
| 154 | qed "lemma_real_complete2a"; | |
| 155 | ||
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 156 | Goal "[| x + (-X) + 1r <= S; xa <= x |] ==> xa <= S + X + (-1r)"; | 
| 5078 | 157 | by (rotate_tac 1 1); | 
| 158 | by (etac (real_le_imp_less_or_eq RS disjE) 1); | |
| 159 | by (blast_tac (claset() addIs [lemma_real_complete2a]) 1); | |
| 160 | by (blast_tac (claset() addIs [(lemma_le_swap2 RS iffD2)]) 1); | |
| 161 | qed "lemma_real_complete2b"; | |
| 162 | ||
| 163 | (*------------------------------------ | |
| 164 | reals Completeness (again!) | |
| 165 | ------------------------------------*) | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 166 | 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 | 167 | \ ==> EX t. isLub (UNIV :: real set) S t"; | 
| 5078 | 168 | by (Step_tac 1); | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 169 | by (subgoal_tac "? u. u: {z. ? x: S. z = x + (-X) + 1r} \
 | 
| 5078 | 170 | \                Int {x. 0r < x}" 1);
 | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 171 | by (subgoal_tac "isUb (UNIV::real set) ({z. ? x: S. z = x + (-X) + 1r} \
 | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 172 | \                Int {x. 0r < x})  (Y + (-X) + 1r)" 1); 
 | 
| 5078 | 173 | 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 | 174 | 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 | 175 | Step_tac]); | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 176 | by (res_inst_tac [("x","t + X + (-1r)")] exI 1);
 | 
| 5078 | 177 | by (rtac isLubI2 1); | 
| 178 | by (rtac setgeI 2 THEN Step_tac 2); | |
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 179 | by (subgoal_tac "isUb (UNIV:: real set) ({z. ? x: S. z = x + (-X) + 1r} \
 | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 180 | \                Int {x. 0r < x})  (y + (-X) + 1r)" 2); 
 | 
| 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 181 | by (dres_inst_tac [("y","(y + (- X) + 1r)")] isLub_le_isUb 2 
 | 
| 5078 | 182 | THEN assume_tac 2); | 
| 5588 | 183 | by (full_simp_tac | 
| 184 | (simpset() addsimps [real_diff_def, real_diff_le_eq RS sym] @ | |
| 185 | real_add_ac) 2); | |
| 5078 | 186 | by (rtac (setleI RS isUbI) 1); | 
| 187 | by (Step_tac 1); | |
| 188 | by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
 | |
| 189 | by (stac lemma_le_swap2 1); | |
| 7499 | 190 | by (ftac isLubD2 1 THEN assume_tac 2); | 
| 5078 | 191 | by (Step_tac 1); | 
| 192 | by (Blast_tac 1); | |
| 193 | by (dtac lemma_real_complete1 1 THEN REPEAT(assume_tac 1)); | |
| 194 | by (stac lemma_le_swap2 1); | |
| 7499 | 195 | by (ftac isLubD2 1 THEN assume_tac 2); | 
| 5078 | 196 | by (Blast_tac 1); | 
| 197 | by (rtac lemma_real_complete2b 1); | |
| 198 | by (etac real_less_imp_le 2); | |
| 199 | by (blast_tac (claset() addSIs [isLubD2]) 1 THEN Step_tac 1); | |
| 5588 | 200 | by (full_simp_tac (simpset() addsimps [real_add_assoc]) 1); | 
| 201 | by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI] | |
| 202 | addIs [real_add_le_mono1]) 1); | |
| 203 | by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI] | |
| 204 | addIs [real_add_le_mono1]) 1); | |
| 205 | by (auto_tac (claset(), | |
| 206 | simpset() addsimps [real_add_assoc RS sym, | |
| 207 | real_zero_less_one])); | |
| 5078 | 208 | qed "reals_complete"; | 
| 209 | ||
| 210 | (*---------------------------------------------------------------- | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 211 | Related: Archimedean property of reals | 
| 5078 | 212 | ----------------------------------------------------------------*) | 
| 213 | ||
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 214 | Goal "0r < x ==> EX n. rinv(real_of_posnat n) < x"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 215 | by (stac real_of_posnat_rinv_Ex_iff 1); | 
| 5078 | 216 | by (EVERY1[rtac ccontr, Asm_full_simp_tac]); | 
| 217 | by (fold_tac [real_le_def]); | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 218 | by (subgoal_tac "isUb (UNIV::real set) \ | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 219 | \   {z. EX n. z = x*(real_of_posnat n)} 1r" 1);
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 220 | by (subgoal_tac "EX X. X : {z. EX n. z = x*(real_of_posnat n)}" 1);
 | 
| 5078 | 221 | by (dtac reals_complete 1); | 
| 222 | by (auto_tac (claset() addIs [isUbI,setleI],simpset())); | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 223 | by (subgoal_tac "ALL m. x*(real_of_posnat(Suc m)) <= t" 1); | 
| 5078 | 224 | by (asm_full_simp_tac (simpset() addsimps | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 225 | [real_of_posnat_Suc,real_add_mult_distrib2]) 1); | 
| 5078 | 226 | by (blast_tac (claset() addIs [isLubD2]) 2); | 
| 5588 | 227 | by (asm_full_simp_tac | 
| 228 | (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 | 229 | by (subgoal_tac "isUb (UNIV::real set) \ | 
| 7127 
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
 paulson parents: 
7077diff
changeset | 230 | \   {z. EX n. z = x*(real_of_posnat n)} (t + (-x))" 1);
 | 
| 5078 | 231 | 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 | 232 | by (dres_inst_tac [("y","t+(-x)")] isLub_le_isUb 1);
 | 
| 5588 | 233 | by (dres_inst_tac [("x","-t")] real_add_left_le_mono1 2);
 | 
| 5078 | 234 | by (auto_tac (claset() addDs [real_le_less_trans, | 
| 5588 | 235 | (real_minus_zero_less_iff2 RS iffD2)], | 
| 236 | simpset() addsimps [real_less_not_refl,real_add_assoc RS sym])); | |
| 5078 | 237 | qed "reals_Archimedean"; | 
| 238 | ||
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 239 | Goal "EX n. (x::real) < real_of_posnat n"; | 
| 5078 | 240 | by (res_inst_tac [("R1.0","x"),("R2.0","0r")] real_linear_less2 1);
 | 
| 241 | by (res_inst_tac [("x","0")] exI 1);
 | |
| 242 | by (res_inst_tac [("x","0")] exI 2);
 | |
| 243 | by (auto_tac (claset() addEs [real_less_trans], | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 244 | simpset() addsimps [real_of_posnat_one,real_zero_less_one])); | 
| 5078 | 245 | by (forward_tac [(real_rinv_gt_zero RS reals_Archimedean)] 1); | 
| 246 | by (Step_tac 1 THEN res_inst_tac [("x","n")] exI 1);
 | |
| 247 | by (forw_inst_tac [("y","rinv x")] real_mult_less_mono1 1);
 | |
| 248 | by (auto_tac (claset(),simpset() addsimps [real_not_refl2 RS not_sym])); | |
| 249 | by (dres_inst_tac [("n1","n"),("y","1r")] 
 | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 250 | (real_of_posnat_less_zero RS real_mult_less_mono2) 1); | 
| 5588 | 251 | by (auto_tac (claset(), | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: 
5588diff
changeset | 252 | simpset() addsimps [real_of_posnat_less_zero, | 
| 5588 | 253 | real_not_refl2 RS not_sym, | 
| 254 | real_mult_assoc RS sym])); | |
| 5078 | 255 | qed "reals_Archimedean2"; | 
| 256 | ||
| 257 | ||
| 258 | ||
| 259 | ||
| 260 |