author | paulson |
Thu, 01 Oct 1998 18:18:01 +0200 | |
changeset 5588 | a3ab526bb891 |
parent 5521 | 7970832271cc |
child 7077 | 60b098bb8b8a |
permissions | -rw-r--r-- |
5078 | 1 |
(* Title : RComplete.thy |
2 |
Author : Jacques D. Fleuriot |
|
3 |
Copyright : 1998 University of Cambridge |
|
4 |
Description : Completeness theorems for positive |
|
5 |
reals and reals |
|
6 |
*) |
|
7 |
||
8 |
||
9 |
open RComplete; |
|
10 |
||
5521 | 11 |
claset_ref() := claset() delWrapper "bspec"; |
12 |
||
5078 | 13 |
Goal "!!(x::real). [| isLub R S x; isLub R S y |] ==> x = y"; |
14 |
by (forward_tac [isLub_isUb] 1); |
|
15 |
by (forw_inst_tac [("x","y")] isLub_isUb 1); |
|
16 |
by (blast_tac (claset() addSIs [real_le_anti_sym] |
|
17 |
addSDs [isLub_le_isUb]) 1); |
|
18 |
qed "real_isLub_unique"; |
|
19 |
||
20 |
Goalw [setle_def,setge_def] |
|
21 |
"!!x::real. [| x <=* S'; S <= S' |] ==> x <=* S"; |
|
22 |
by (Blast_tac 1); |
|
23 |
qed "real_order_restrict"; |
|
24 |
||
25 |
(*---------------------------------------------------------------- |
|
26 |
Completeness theorem for the positive reals(again) |
|
27 |
----------------------------------------------------------------*) |
|
28 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
29 |
Goal "[| ALL x: S. 0r < x; \ |
5078 | 30 |
\ EX x. x: S; \ |
31 |
\ EX u. isUb (UNIV::real set) S u \ |
|
32 |
\ |] ==> EX t. isLub (UNIV::real set) S t"; |
|
33 |
by (res_inst_tac [("x","%#psup({w. %#w : S})")] exI 1); |
|
5588 | 34 |
by (auto_tac (claset(), simpset() addsimps [isLub_def,leastP_def,isUb_def])); |
5078 | 35 |
by (auto_tac (claset() addSIs [setleI,setgeI] |
5588 | 36 |
addSDs [real_gt_zero_preal_Ex RS iffD1],simpset())); |
5078 | 37 |
by (forw_inst_tac [("x","y")] bspec 1 THEN assume_tac 1); |
38 |
by (dtac (real_gt_zero_preal_Ex RS iffD1) 1); |
|
5588 | 39 |
by (auto_tac (claset(), simpset() addsimps [real_preal_le_iff])); |
5078 | 40 |
by (rtac preal_psup_leI2a 1); |
41 |
by (forw_inst_tac [("y","%#ya")] setleD 1 THEN assume_tac 1); |
|
42 |
by (forward_tac [real_ge_preal_preal_Ex] 1); |
|
43 |
by (Step_tac 1); |
|
44 |
by (res_inst_tac [("x","y")] exI 1); |
|
45 |
by (blast_tac (claset() addSDs [setleD] addIs [real_preal_le_iff RS iffD1]) 1); |
|
46 |
by (forw_inst_tac [("x","x")] bspec 1 THEN assume_tac 1); |
|
47 |
by (forward_tac [isUbD2] 1); |
|
48 |
by (dtac (real_gt_zero_preal_Ex RS iffD1) 1); |
|
5588 | 49 |
by (auto_tac (claset() addSDs [isUbD, real_ge_preal_preal_Ex], |
50 |
simpset() addsimps [real_preal_le_iff])); |
|
51 |
by (blast_tac (claset() addSDs [setleD] addSIs [psup_le_ub1] |
|
52 |
addIs [real_preal_le_iff RS iffD1]) 1); |
|
5078 | 53 |
qed "posreals_complete"; |
54 |
||
55 |
||
56 |
(*------------------------------- |
|
57 |
Lemmas |
|
58 |
-------------------------------*) |
|
5588 | 59 |
Goal "! y : {z. ? x: P. z = x + -xa + 1r} Int {x. 0r < x}. 0r < y"; |
5078 | 60 |
by Auto_tac; |
61 |
qed "real_sup_lemma3"; |
|
62 |
||
5588 | 63 |
Goal "(xa <= S + X + -Z) = (xa + -X + Z <= (S::real))"; |
64 |
by (simp_tac (simpset() addsimps [real_diff_def, real_diff_le_eq RS sym] @ |
|
65 |
real_add_ac) 1); |
|
5078 | 66 |
qed "lemma_le_swap2"; |
67 |
||
5588 | 68 |
Goal "[| 0r < x + -X + 1r; x < xa |] ==> 0r < xa + -X + 1r"; |
5078 | 69 |
by (dtac real_add_less_mono 1); |
70 |
by (assume_tac 1); |
|
5588 | 71 |
by (dres_inst_tac [("C","-x"),("A","0r + x")] real_add_less_mono2 1); |
5078 | 72 |
by (asm_full_simp_tac (simpset() addsimps [real_add_zero_right, |
73 |
real_add_assoc RS sym,real_add_minus_left,real_add_zero_left]) 1); |
|
74 |
by (asm_full_simp_tac (simpset() addsimps real_add_ac) 1); |
|
75 |
qed "lemma_real_complete1"; |
|
76 |
||
5588 | 77 |
Goal "!!x. [| x + -X + 1r <= S; xa < x |] ==> xa + -X + 1r <= S"; |
5078 | 78 |
by (dtac real_less_imp_le 1); |
79 |
by (dtac real_add_le_mono 1); |
|
80 |
by (assume_tac 1); |
|
81 |
by (asm_full_simp_tac (simpset() addsimps real_add_ac) 1); |
|
82 |
qed "lemma_real_complete2"; |
|
83 |
||
5588 | 84 |
Goal "[| x + -X + 1r <= S; xa < x |] ==> xa <= S + X + -1r"; (**) |
5078 | 85 |
by (rtac (lemma_le_swap2 RS iffD2) 1); |
86 |
by (etac lemma_real_complete2 1); |
|
87 |
by (assume_tac 1); |
|
88 |
qed "lemma_real_complete2a"; |
|
89 |
||
5588 | 90 |
Goal "[| x + -X + 1r <= S; xa <= x |] ==> xa <= S + X + -1r"; |
5078 | 91 |
by (rotate_tac 1 1); |
92 |
by (etac (real_le_imp_less_or_eq RS disjE) 1); |
|
93 |
by (blast_tac (claset() addIs [lemma_real_complete2a]) 1); |
|
94 |
by (blast_tac (claset() addIs [(lemma_le_swap2 RS iffD2)]) 1); |
|
95 |
qed "lemma_real_complete2b"; |
|
96 |
||
97 |
(*------------------------------------ |
|
98 |
reals Completeness (again!) |
|
99 |
------------------------------------*) |
|
100 |
Goal "!!(S::real set). [| EX X. X: S; \ |
|
101 |
\ EX Y. isUb (UNIV::real set) S Y \ |
|
102 |
\ |] ==> EX t. isLub (UNIV :: real set) S t"; |
|
103 |
by (Step_tac 1); |
|
5588 | 104 |
by (subgoal_tac "? u. u: {z. ? x: S. z = x + -X + 1r} \ |
5078 | 105 |
\ Int {x. 0r < x}" 1); |
5588 | 106 |
by (subgoal_tac "isUb (UNIV::real set) ({z. ? x: S. z = x + -X + 1r} \ |
107 |
\ Int {x. 0r < x}) (Y + -X + 1r)" 1); |
|
5078 | 108 |
by (cut_inst_tac [("P","S"),("xa","X")] real_sup_lemma3 1); |
109 |
by (EVERY1[forward_tac [exI RSN (3,posreals_complete)], Blast_tac, Blast_tac, Step_tac]); |
|
5588 | 110 |
by (res_inst_tac [("x","t + X + -1r")] exI 1); |
5078 | 111 |
by (rtac isLubI2 1); |
112 |
by (rtac setgeI 2 THEN Step_tac 2); |
|
5588 | 113 |
by (subgoal_tac "isUb (UNIV:: real set) ({z. ? x: S. z = x + -X + 1r} \ |
114 |
\ Int {x. 0r < x}) (y + -X + 1r)" 2); |
|
115 |
by (dres_inst_tac [("y","(y + - X + 1r)")] isLub_le_isUb 2 |
|
5078 | 116 |
THEN assume_tac 2); |
5588 | 117 |
by (full_simp_tac |
118 |
(simpset() addsimps [real_diff_def, real_diff_le_eq RS sym] @ |
|
119 |
real_add_ac) 2); |
|
5078 | 120 |
by (rtac (setleI RS isUbI) 1); |
121 |
by (Step_tac 1); |
|
122 |
by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1); |
|
123 |
by (stac lemma_le_swap2 1); |
|
124 |
by (forward_tac [isLubD2] 1 THEN assume_tac 2); |
|
125 |
by (Step_tac 1); |
|
126 |
by (Blast_tac 1); |
|
127 |
by (dtac lemma_real_complete1 1 THEN REPEAT(assume_tac 1)); |
|
128 |
by (stac lemma_le_swap2 1); |
|
129 |
by (forward_tac [isLubD2] 1 THEN assume_tac 2); |
|
130 |
by (Blast_tac 1); |
|
131 |
by (rtac lemma_real_complete2b 1); |
|
132 |
by (etac real_less_imp_le 2); |
|
133 |
by (blast_tac (claset() addSIs [isLubD2]) 1 THEN Step_tac 1); |
|
5588 | 134 |
by (full_simp_tac (simpset() addsimps [real_add_assoc]) 1); |
135 |
by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI] |
|
136 |
addIs [real_add_le_mono1]) 1); |
|
137 |
by (blast_tac (claset() addDs [isUbD] addSIs [setleI RS isUbI] |
|
138 |
addIs [real_add_le_mono1]) 1); |
|
139 |
by (auto_tac (claset(), |
|
140 |
simpset() addsimps [real_add_assoc RS sym, |
|
141 |
real_zero_less_one])); |
|
5078 | 142 |
qed "reals_complete"; |
143 |
||
144 |
(*---------------------------------------------------------------- |
|
145 |
Related property: Archimedean property of reals |
|
146 |
----------------------------------------------------------------*) |
|
147 |
||
5143
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
paulson
parents:
5078
diff
changeset
|
148 |
Goal "0r < x ==> EX n. rinv(%%#n) < x"; |
5078 | 149 |
by (stac real_nat_rinv_Ex_iff 1); |
150 |
by (EVERY1[rtac ccontr, Asm_full_simp_tac]); |
|
151 |
by (fold_tac [real_le_def]); |
|
152 |
by (subgoal_tac "isUb (UNIV::real set) {z. EX n. z = x*%%#n} 1r" 1); |
|
153 |
by (subgoal_tac "EX X. X : {z. EX n. z = x*%%#n}" 1); |
|
154 |
by (dtac reals_complete 1); |
|
155 |
by (auto_tac (claset() addIs [isUbI,setleI],simpset())); |
|
156 |
by (subgoal_tac "ALL m. x*(%%#Suc m) <= t" 1); |
|
157 |
by (asm_full_simp_tac (simpset() addsimps |
|
158 |
[real_nat_Suc,real_add_mult_distrib2]) 1); |
|
159 |
by (blast_tac (claset() addIs [isLubD2]) 2); |
|
5588 | 160 |
by (asm_full_simp_tac |
161 |
(simpset() addsimps [real_le_diff_eq RS sym, real_diff_def]) 1); |
|
162 |
by (subgoal_tac "isUb (UNIV::real set) {z. EX n. z = x*%%#n} (t + -x)" 1); |
|
5078 | 163 |
by (blast_tac (claset() addSIs [isUbI,setleI]) 2); |
5588 | 164 |
by (dres_inst_tac [("y","t+-x")] isLub_le_isUb 1); |
165 |
by (dres_inst_tac [("x","-t")] real_add_left_le_mono1 2); |
|
5078 | 166 |
by (auto_tac (claset() addDs [real_le_less_trans, |
5588 | 167 |
(real_minus_zero_less_iff2 RS iffD2)], |
168 |
simpset() addsimps [real_less_not_refl,real_add_assoc RS sym])); |
|
5078 | 169 |
qed "reals_Archimedean"; |
170 |
||
171 |
Goal "EX n. (x::real) < %%#n"; |
|
172 |
by (res_inst_tac [("R1.0","x"),("R2.0","0r")] real_linear_less2 1); |
|
173 |
by (res_inst_tac [("x","0")] exI 1); |
|
174 |
by (res_inst_tac [("x","0")] exI 2); |
|
175 |
by (auto_tac (claset() addEs [real_less_trans], |
|
5588 | 176 |
simpset() addsimps [real_nat_one,real_zero_less_one])); |
5078 | 177 |
by (forward_tac [(real_rinv_gt_zero RS reals_Archimedean)] 1); |
178 |
by (Step_tac 1 THEN res_inst_tac [("x","n")] exI 1); |
|
179 |
by (forw_inst_tac [("y","rinv x")] real_mult_less_mono1 1); |
|
180 |
by (auto_tac (claset(),simpset() addsimps [real_not_refl2 RS not_sym])); |
|
181 |
by (dres_inst_tac [("n1","n"),("y","1r")] |
|
182 |
(real_nat_less_zero RS real_mult_less_mono2) 1); |
|
5588 | 183 |
by (auto_tac (claset(), |
184 |
simpset() addsimps [real_nat_less_zero, |
|
185 |
real_not_refl2 RS not_sym, |
|
186 |
real_mult_assoc RS sym])); |
|
5078 | 187 |
qed "reals_Archimedean2"; |
188 |
||
189 |
||
190 |
||
191 |
||
192 |