| author | wenzelm | 
| Fri, 20 Jul 2001 17:49:21 +0200 | |
| changeset 11431 | 2328d48eeba8 | 
| parent 11377 | 0f16ad464c62 | 
| child 11468 | 02cd5d5bc497 | 
| permissions | -rw-r--r-- | 
| 10751 | 1 | (* Title : HyperPow.ML | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 4 | Description : Natural Powers of hyperreals theory | |
| 5 | ||
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 6 | Exponentials on the hyperreals | 
| 10751 | 7 | *) | 
| 8 | ||
| 9 | Goal "(#0::hypreal) ^ (Suc n) = 0"; | |
| 10 | by (Auto_tac); | |
| 11 | qed "hrealpow_zero"; | |
| 12 | Addsimps [hrealpow_zero]; | |
| 13 | ||
| 14 | Goal "r ~= (#0::hypreal) --> r ^ n ~= 0"; | |
| 15 | by (induct_tac "n" 1); | |
| 16 | by Auto_tac; | |
| 17 | qed_spec_mp "hrealpow_not_zero"; | |
| 18 | ||
| 19 | Goal "r ~= (#0::hypreal) --> inverse(r ^ n) = (inverse r) ^ n"; | |
| 20 | by (induct_tac "n" 1); | |
| 21 | by (Auto_tac); | |
| 22 | by (forw_inst_tac [("n","n")] hrealpow_not_zero 1);
 | |
| 23 | by (auto_tac (claset(), simpset() addsimps [hypreal_inverse_distrib])); | |
| 24 | qed_spec_mp "hrealpow_inverse"; | |
| 25 | ||
| 26 | Goal "abs (r::hypreal) ^ n = abs (r ^ n)"; | |
| 27 | by (induct_tac "n" 1); | |
| 28 | by (auto_tac (claset(), simpset() addsimps [hrabs_mult])); | |
| 29 | qed "hrealpow_hrabs"; | |
| 30 | ||
| 31 | Goal "(r::hypreal) ^ (n + m) = (r ^ n) * (r ^ m)"; | |
| 32 | by (induct_tac "n" 1); | |
| 33 | by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac)); | |
| 34 | qed "hrealpow_add"; | |
| 35 | ||
| 36 | Goal "(r::hypreal) ^ 1 = r"; | |
| 37 | by (Simp_tac 1); | |
| 38 | qed "hrealpow_one"; | |
| 39 | Addsimps [hrealpow_one]; | |
| 40 | ||
| 41 | Goal "(r::hypreal) ^ 2 = r * r"; | |
| 42 | by (Simp_tac 1); | |
| 43 | qed "hrealpow_two"; | |
| 44 | ||
| 45 | Goal "(#0::hypreal) <= r --> #0 <= r ^ n"; | |
| 46 | by (induct_tac "n" 1); | |
| 47 | by (auto_tac (claset(), simpset() addsimps [hypreal_0_le_mult_iff])); | |
| 48 | qed_spec_mp "hrealpow_ge_zero"; | |
| 49 | ||
| 50 | Goal "(#0::hypreal) < r --> #0 < r ^ n"; | |
| 51 | by (induct_tac "n" 1); | |
| 52 | by (auto_tac (claset(), simpset() addsimps [hypreal_0_less_mult_iff])); | |
| 53 | qed_spec_mp "hrealpow_gt_zero"; | |
| 54 | ||
| 55 | Goal "x <= y & (#0::hypreal) < x --> x ^ n <= y ^ n"; | |
| 56 | by (induct_tac "n" 1); | |
| 57 | by (auto_tac (claset() addSIs [hypreal_mult_le_mono], simpset())); | |
| 58 | by (asm_simp_tac (simpset() addsimps [hrealpow_ge_zero]) 1); | |
| 59 | qed_spec_mp "hrealpow_le"; | |
| 60 | ||
| 61 | Goal "x < y & (#0::hypreal) < x & 0 < n --> x ^ n < y ^ n"; | |
| 62 | by (induct_tac "n" 1); | |
| 63 | by (auto_tac (claset() addIs [hypreal_mult_less_mono,gr0I], | |
| 64 | simpset() addsimps [hrealpow_gt_zero])); | |
| 65 | qed "hrealpow_less"; | |
| 66 | ||
| 67 | Goal "#1 ^ n = (#1::hypreal)"; | |
| 68 | by (induct_tac "n" 1); | |
| 69 | by (Auto_tac); | |
| 70 | qed "hrealpow_eq_one"; | |
| 71 | Addsimps [hrealpow_eq_one]; | |
| 72 | ||
| 73 | Goal "abs(-(#1 ^ n)) = (#1::hypreal)"; | |
| 74 | by Auto_tac; | |
| 75 | qed "hrabs_minus_hrealpow_one"; | |
| 76 | Addsimps [hrabs_minus_hrealpow_one]; | |
| 77 | ||
| 78 | Goal "abs(#-1 ^ n) = (#1::hypreal)"; | |
| 79 | by (induct_tac "n" 1); | |
| 80 | by Auto_tac; | |
| 81 | qed "hrabs_hrealpow_minus_one"; | |
| 82 | Addsimps [hrabs_hrealpow_minus_one]; | |
| 83 | ||
| 84 | Goal "((r::hypreal) * s) ^ n = (r ^ n) * (s ^ n)"; | |
| 85 | by (induct_tac "n" 1); | |
| 86 | by (auto_tac (claset(), simpset() addsimps hypreal_mult_ac)); | |
| 87 | qed "hrealpow_mult"; | |
| 88 | ||
| 89 | Goal "(#0::hypreal) <= r ^ 2"; | |
| 90 | by (auto_tac (claset(), simpset() addsimps [hypreal_0_le_mult_iff])); | |
| 91 | qed "hrealpow_two_le"; | |
| 92 | Addsimps [hrealpow_two_le]; | |
| 93 | ||
| 94 | Goal "(#0::hypreal) <= u ^ 2 + v ^ 2"; | |
| 95 | by (simp_tac (HOL_ss addsimps [hrealpow_two_le, | |
| 96 | rename_numerals hypreal_le_add_order]) 1); | |
| 97 | qed "hrealpow_two_le_add_order"; | |
| 98 | Addsimps [hrealpow_two_le_add_order]; | |
| 99 | ||
| 100 | Goal "(#0::hypreal) <= u ^ 2 + v ^ 2 + w ^ 2"; | |
| 101 | by (simp_tac (HOL_ss addsimps [hrealpow_two_le, | |
| 102 | rename_numerals hypreal_le_add_order]) 1); | |
| 103 | qed "hrealpow_two_le_add_order2"; | |
| 104 | Addsimps [hrealpow_two_le_add_order2]; | |
| 105 | ||
| 106 | Goal "(x ^ 2 + y ^ 2 + z ^ 2 = (#0::hypreal)) = (x = #0 & y = #0 & z = #0)"; | |
| 107 | by (simp_tac (HOL_ss addsimps | |
| 108 | [rename_numerals hypreal_three_squares_add_zero_iff, hrealpow_two]) 1); | |
| 109 | qed "hrealpow_three_squares_add_zero_iff"; | |
| 110 | Addsimps [hrealpow_three_squares_add_zero_iff]; | |
| 111 | ||
| 112 | Goal "abs(x ^ 2) = (x::hypreal) ^ 2"; | |
| 113 | by (auto_tac (claset(), | |
| 114 | simpset() addsimps [hrabs_def, hypreal_0_le_mult_iff])); | |
| 115 | qed "hrabs_hrealpow_two"; | |
| 116 | Addsimps [hrabs_hrealpow_two]; | |
| 117 | ||
| 118 | Goal "abs(x) ^ 2 = (x::hypreal) ^ 2"; | |
| 119 | by (simp_tac (simpset() addsimps [hrealpow_hrabs, hrabs_eqI1] | |
| 120 | delsimps [hpowr_Suc]) 1); | |
| 121 | qed "hrealpow_two_hrabs"; | |
| 122 | Addsimps [hrealpow_two_hrabs]; | |
| 123 | ||
| 124 | Goal "(#1::hypreal) < r ==> #1 < r ^ 2"; | |
| 125 | by (auto_tac (claset(), simpset() addsimps [hrealpow_two])); | |
| 126 | by (res_inst_tac [("y","#1*#1")] order_le_less_trans 1); 
 | |
| 127 | by (rtac hypreal_mult_less_mono 2); | |
| 128 | by Auto_tac; | |
| 129 | qed "hrealpow_two_gt_one"; | |
| 130 | ||
| 131 | Goal "(#1::hypreal) <= r ==> #1 <= r ^ 2"; | |
| 132 | by (etac (order_le_imp_less_or_eq RS disjE) 1); | |
| 133 | by (etac (hrealpow_two_gt_one RS order_less_imp_le) 1); | |
| 134 | by Auto_tac; | |
| 135 | qed "hrealpow_two_ge_one"; | |
| 136 | ||
| 137 | Goal "(#1::hypreal) <= #2 ^ n"; | |
| 138 | by (res_inst_tac [("y","#1 ^ n")] order_trans 1);
 | |
| 139 | by (rtac hrealpow_le 2); | |
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 140 | by Auto_tac; | 
| 10751 | 141 | qed "two_hrealpow_ge_one"; | 
| 142 | ||
| 143 | Goal "hypreal_of_nat n < #2 ^ n"; | |
| 144 | by (induct_tac "n" 1); | |
| 145 | by (auto_tac (claset(), | |
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 146 | simpset() addsimps [hypreal_of_nat_Suc, hypreal_add_mult_distrib])); | 
| 10751 | 147 | by (cut_inst_tac [("n","n")] two_hrealpow_ge_one 1);
 | 
| 148 | by (arith_tac 1); | |
| 149 | qed "two_hrealpow_gt"; | |
| 150 | Addsimps [two_hrealpow_gt,two_hrealpow_ge_one]; | |
| 151 | ||
| 11377 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 152 | Goal "#-1 ^ (#2*n) = (#1::hypreal)"; | 
| 10751 | 153 | by (induct_tac "n" 1); | 
| 154 | by (Auto_tac); | |
| 155 | qed "hrealpow_minus_one"; | |
| 156 | ||
| 11377 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 157 | Goal "n+n = (#2*n::nat)"; | 
| 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 158 | by Auto_tac; | 
| 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 159 | qed "double_lemma"; | 
| 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 160 | |
| 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 161 | (*ugh: need to get rid fo the n+n*) | 
| 10751 | 162 | Goal "#-1 ^ (n + n) = (#1::hypreal)"; | 
| 11377 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 163 | by (auto_tac (claset(), | 
| 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 164 | simpset() addsimps [double_lemma, hrealpow_minus_one])); | 
| 10751 | 165 | qed "hrealpow_minus_one2"; | 
| 166 | Addsimps [hrealpow_minus_one2]; | |
| 167 | ||
| 168 | Goal "(-(x::hypreal)) ^ 2 = x ^ 2"; | |
| 169 | by (Auto_tac); | |
| 170 | qed "hrealpow_minus_two"; | |
| 171 | Addsimps [hrealpow_minus_two]; | |
| 172 | ||
| 173 | Goal "(#0::hypreal) < r & r < #1 --> r ^ Suc n < r ^ n"; | |
| 174 | by (induct_tac "n" 1); | |
| 175 | by (auto_tac (claset(), | |
| 176 | simpset() addsimps [hypreal_mult_less_mono2])); | |
| 177 | qed_spec_mp "hrealpow_Suc_less"; | |
| 178 | ||
| 179 | Goal "(#0::hypreal) <= r & r < #1 --> r ^ Suc n <= r ^ n"; | |
| 180 | by (induct_tac "n" 1); | |
| 181 | by (auto_tac (claset() addIs [order_less_imp_le] | |
| 182 | addSDs [order_le_imp_less_or_eq,hrealpow_Suc_less], | |
| 183 | simpset() addsimps [hypreal_mult_less_mono2])); | |
| 184 | qed_spec_mp "hrealpow_Suc_le"; | |
| 185 | ||
| 10834 | 186 | Goal "Abs_hypreal(hyprel``{%n. X n}) ^ m = Abs_hypreal(hyprel``{%n. (X n) ^ m})";
 | 
| 10751 | 187 | by (induct_tac "m" 1); | 
| 188 | by (auto_tac (claset(), | |
| 189 | simpset() delsimps [one_eq_numeral_1] | |
| 190 | addsimps [hypreal_one_def, hypreal_mult, | |
| 191 | one_eq_numeral_1 RS sym])); | |
| 192 | qed "hrealpow"; | |
| 193 | ||
| 194 | Goal "(x + (y::hypreal)) ^ 2 = \ | |
| 195 | \ x ^ 2 + y ^ 2 + (hypreal_of_nat 2)*x*y"; | |
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 196 | by (simp_tac (simpset() addsimps | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 197 | [hypreal_add_mult_distrib2, hypreal_add_mult_distrib, | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 198 | hypreal_of_nat_zero, hypreal_of_nat_Suc]) 1); | 
| 10751 | 199 | qed "hrealpow_sum_square_expand"; | 
| 200 | ||
| 201 | (*--------------------------------------------------------------- | |
| 202 | we'll prove the following theorem by going down to the | |
| 203 | level of the ultrafilter and relying on the analogous | |
| 204 | property for the real rather than prove it directly | |
| 205 | using induction: proof is much simpler this way! | |
| 206 | ---------------------------------------------------------------*) | |
| 207 | Goal "[|(#0::hypreal) <= x; #0 <= y;x ^ Suc n <= y ^ Suc n |] ==> x <= y"; | |
| 10784 | 208 | by (full_simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1); | 
| 10751 | 209 | by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
 | 
| 210 | by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
 | |
| 211 | by (auto_tac (claset(), | |
| 212 | simpset() addsimps [hrealpow,hypreal_le,hypreal_mult])); | |
| 213 | by (ultra_tac (claset() addIs [realpow_increasing], simpset()) 1); | |
| 214 | qed "hrealpow_increasing"; | |
| 215 | ||
| 216 | (*By antisymmetry with the above we conclude x=y, replacing the deleted | |
| 217 | theorem hrealpow_Suc_cancel_eq*) | |
| 218 | ||
| 219 | Goal "x : HFinite --> x ^ n : HFinite"; | |
| 220 | by (induct_tac "n" 1); | |
| 221 | by (auto_tac (claset() addIs [HFinite_mult], simpset())); | |
| 222 | qed_spec_mp "hrealpow_HFinite"; | |
| 223 | ||
| 224 | (*--------------------------------------------------------------- | |
| 225 | Hypernaturals Powers | |
| 226 | --------------------------------------------------------------*) | |
| 227 | Goalw [congruent_def] | |
| 228 | "congruent hyprel \ | |
| 10834 | 229 | \    (%X Y. hyprel``{%n. ((X::nat=>real) n ^ (Y::nat=>nat) n)})";
 | 
| 10751 | 230 | by (safe_tac (claset() addSIs [ext])); | 
| 231 | by (ALLGOALS(Fuf_tac)); | |
| 232 | qed "hyperpow_congruent"; | |
| 233 | ||
| 234 | Goalw [hyperpow_def] | |
| 10834 | 235 |   "Abs_hypreal(hyprel``{%n. X n}) pow Abs_hypnat(hypnatrel``{%n. Y n}) = \
 | 
| 236 | \  Abs_hypreal(hyprel``{%n. X n ^ Y n})";
 | |
| 10751 | 237 | by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
 | 
| 238 | by (auto_tac (claset() addSIs [lemma_hyprel_refl,bexI], | |
| 239 | simpset() addsimps [hyprel_in_hypreal RS | |
| 240 | Abs_hypreal_inverse,equiv_hyprel,hyperpow_congruent])); | |
| 241 | by (Fuf_tac 1); | |
| 242 | qed "hyperpow"; | |
| 243 | ||
| 244 | Goalw [hypnat_one_def] "(#0::hypreal) pow (n + 1hn) = #0"; | |
| 10784 | 245 | by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1); | 
| 10751 | 246 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | 
| 247 | by (auto_tac (claset(), simpset() addsimps [hyperpow,hypnat_add])); | |
| 248 | qed "hyperpow_zero"; | |
| 249 | Addsimps [hyperpow_zero]; | |
| 250 | ||
| 251 | Goal "r ~= (#0::hypreal) --> r pow n ~= #0"; | |
| 10784 | 252 | by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1); | 
| 10751 | 253 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | 
| 254 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 255 | by (auto_tac (claset(), simpset() addsimps [hyperpow])); | |
| 256 | by (dtac FreeUltrafilterNat_Compl_mem 1); | |
| 257 | by (fuf_empty_tac (claset() addIs [realpow_not_zero RS notE], | |
| 258 | simpset()) 1); | |
| 259 | qed_spec_mp "hyperpow_not_zero"; | |
| 260 | ||
| 261 | Goal "r ~= (#0::hypreal) --> inverse(r pow n) = (inverse r) pow n"; | |
| 10784 | 262 | by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1); | 
| 10751 | 263 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | 
| 264 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 265 | by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem], | |
| 266 | simpset() addsimps [hypreal_inverse,hyperpow])); | |
| 267 | by (rtac FreeUltrafilterNat_subset 1); | |
| 268 | by (auto_tac (claset() addDs [realpow_not_zero] | |
| 269 | addIs [realpow_inverse], | |
| 270 | simpset())); | |
| 271 | qed "hyperpow_inverse"; | |
| 272 | ||
| 273 | Goal "abs r pow n = abs (r pow n)"; | |
| 274 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 275 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 276 | by (auto_tac (claset(), | |
| 277 | simpset() addsimps [hypreal_hrabs, hyperpow,realpow_abs])); | |
| 278 | qed "hyperpow_hrabs"; | |
| 279 | ||
| 280 | Goal "r pow (n + m) = (r pow n) * (r pow m)"; | |
| 281 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 282 | by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
 | |
| 283 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 284 | by (auto_tac (claset(), | |
| 285 | simpset() addsimps [hyperpow,hypnat_add, hypreal_mult,realpow_add])); | |
| 286 | qed "hyperpow_add"; | |
| 287 | ||
| 288 | Goalw [hypnat_one_def] "r pow 1hn = r"; | |
| 289 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 290 | by (auto_tac (claset(), simpset() addsimps [hyperpow])); | |
| 291 | qed "hyperpow_one"; | |
| 292 | Addsimps [hyperpow_one]; | |
| 293 | ||
| 294 | Goalw [hypnat_one_def] | |
| 10784 | 295 | "r pow (1hn + 1hn) = r * r"; | 
| 10751 | 296 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | 
| 297 | by (auto_tac (claset(), | |
| 10784 | 298 | simpset() addsimps [hyperpow,hypnat_add, hypreal_mult])); | 
| 10751 | 299 | qed "hyperpow_two"; | 
| 300 | ||
| 301 | Goal "(#0::hypreal) < r --> #0 < r pow n"; | |
| 10784 | 302 | by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1); | 
| 10751 | 303 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | 
| 304 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 305 | by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset, realpow_gt_zero], | |
| 306 | simpset() addsimps [hyperpow,hypreal_less, hypreal_le])); | |
| 307 | qed_spec_mp "hyperpow_gt_zero"; | |
| 308 | ||
| 309 | Goal "(#0::hypreal) <= r --> #0 <= r pow n"; | |
| 10784 | 310 | by (simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1); | 
| 10751 | 311 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | 
| 312 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 10784 | 313 | by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset, realpow_ge_zero], | 
| 10751 | 314 | simpset() addsimps [hyperpow,hypreal_le])); | 
| 10784 | 315 | qed "hyperpow_ge_zero"; | 
| 10751 | 316 | |
| 317 | Goal "(#0::hypreal) < x & x <= y --> x pow n <= y pow n"; | |
| 10784 | 318 | by (full_simp_tac (simpset() addsimps [rename_numerals hypreal_zero_def]) 1); | 
| 10751 | 319 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | 
| 320 | by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
 | |
| 321 | by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
 | |
| 10784 | 322 | by (auto_tac (claset(), | 
| 323 | simpset() addsimps [hyperpow,hypreal_le,hypreal_less])); | |
| 324 | by (etac (FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset) 1 | |
| 325 | THEN assume_tac 1); | |
| 326 | by (auto_tac (claset() addIs [realpow_le], simpset())); | |
| 10751 | 327 | qed_spec_mp "hyperpow_le"; | 
| 328 | ||
| 329 | Goal "#1 pow n = (#1::hypreal)"; | |
| 330 | by (simp_tac (HOL_ss addsimps [one_eq_numeral_1 RS sym, hypreal_one_def]) 1); | |
| 331 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 332 | by (auto_tac (claset(), simpset() addsimps [hyperpow])); | |
| 333 | qed "hyperpow_eq_one"; | |
| 334 | Addsimps [hyperpow_eq_one]; | |
| 335 | ||
| 336 | Goal "abs(-(#1 pow n)) = (#1::hypreal)"; | |
| 337 | by (simp_tac (HOL_ss addsimps [one_eq_numeral_1 RS sym, hypreal_one_def]) 1); | |
| 338 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 339 | by (auto_tac (claset(), simpset() addsimps [hyperpow,hypreal_hrabs])); | |
| 340 | qed "hrabs_minus_hyperpow_one"; | |
| 341 | Addsimps [hrabs_minus_hyperpow_one]; | |
| 342 | ||
| 343 | Goal "abs(#-1 pow n) = (#1::hypreal)"; | |
| 344 | by (subgoal_tac "abs((-1hr) pow n) = 1hr" 1); | |
| 345 | by (Asm_full_simp_tac 1); | |
| 346 | by (simp_tac (HOL_ss addsimps [hypreal_one_def]) 1); | |
| 347 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 348 | by (auto_tac (claset(), | |
| 349 | simpset() addsimps [hyperpow,hypreal_minus,hypreal_hrabs])); | |
| 350 | qed "hrabs_hyperpow_minus_one"; | |
| 351 | Addsimps [hrabs_hyperpow_minus_one]; | |
| 352 | ||
| 353 | Goal "(r * s) pow n = (r pow n) * (s pow n)"; | |
| 354 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 355 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 356 | by (res_inst_tac [("z","s")] eq_Abs_hypreal 1);
 | |
| 357 | by (auto_tac (claset(), | |
| 358 | simpset() addsimps [hyperpow, hypreal_mult,realpow_mult])); | |
| 359 | qed "hyperpow_mult"; | |
| 360 | ||
| 361 | Goal "(#0::hypreal) <= r pow (1hn + 1hn)"; | |
| 362 | by (auto_tac (claset(), | |
| 363 | simpset() addsimps [hyperpow_two, hypreal_0_le_mult_iff])); | |
| 364 | qed "hyperpow_two_le"; | |
| 365 | Addsimps [hyperpow_two_le]; | |
| 366 | ||
| 367 | Goal "abs(x pow (1hn + 1hn)) = x pow (1hn + 1hn)"; | |
| 368 | by (simp_tac (simpset() addsimps [hrabs_eqI1]) 1); | |
| 369 | qed "hrabs_hyperpow_two"; | |
| 370 | Addsimps [hrabs_hyperpow_two]; | |
| 371 | ||
| 372 | Goal "abs(x) pow (1hn + 1hn) = x pow (1hn + 1hn)"; | |
| 373 | by (simp_tac (simpset() addsimps [hyperpow_hrabs,hrabs_eqI1]) 1); | |
| 374 | qed "hyperpow_two_hrabs"; | |
| 375 | Addsimps [hyperpow_two_hrabs]; | |
| 376 | ||
| 377 | (*? very similar to hrealpow_two_gt_one *) | |
| 378 | Goal "(#1::hypreal) < r ==> #1 < r pow (1hn + 1hn)"; | |
| 379 | by (auto_tac (claset(), simpset() addsimps [hyperpow_two])); | |
| 380 | by (res_inst_tac [("y","#1*#1")] order_le_less_trans 1); 
 | |
| 381 | by (rtac hypreal_mult_less_mono 2); | |
| 382 | by Auto_tac; | |
| 383 | qed "hyperpow_two_gt_one"; | |
| 384 | ||
| 385 | Goal "(#1::hypreal) <= r ==> #1 <= r pow (1hn + 1hn)"; | |
| 386 | by (auto_tac (claset() addSDs [order_le_imp_less_or_eq] | |
| 387 | addIs [hyperpow_two_gt_one,order_less_imp_le], | |
| 388 | simpset())); | |
| 389 | qed "hyperpow_two_ge_one"; | |
| 390 | ||
| 391 | Goal "(#1::hypreal) <= #2 pow n"; | |
| 392 | by (res_inst_tac [("y","#1 pow n")] order_trans 1);
 | |
| 393 | by (rtac hyperpow_le 2); | |
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 394 | by Auto_tac; | 
| 10751 | 395 | qed "two_hyperpow_ge_one"; | 
| 396 | Addsimps [two_hyperpow_ge_one]; | |
| 397 | ||
| 398 | Addsimps [simplify (simpset()) realpow_minus_one]; | |
| 399 | ||
| 400 | Goal "#-1 pow ((1hn + 1hn)*n) = (#1::hypreal)"; | |
| 401 | by (subgoal_tac "(-(1hr)) pow ((1hn + 1hn)*n) = 1hr" 1); | |
| 402 | by (Asm_full_simp_tac 1); | |
| 403 | by (simp_tac (HOL_ss addsimps [hypreal_one_def]) 1); | |
| 404 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 405 | by (auto_tac (claset(), | |
| 11377 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 406 | simpset() addsimps [double_lemma, hyperpow, hypnat_add, | 
| 
0f16ad464c62
Simprocs for type "nat" no longer introduce numerals unless they are already
 paulson parents: 
10919diff
changeset | 407 | hypreal_minus])); | 
| 10751 | 408 | qed "hyperpow_minus_one2"; | 
| 409 | Addsimps [hyperpow_minus_one2]; | |
| 410 | ||
| 411 | Goalw [hypnat_one_def] | |
| 412 | "(#0::hypreal) < r & r < #1 --> r pow (n + 1hn) < r pow n"; | |
| 413 | by (full_simp_tac | |
| 414 | (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def, | |
| 415 | one_eq_numeral_1 RS sym, hypreal_one_def]) 1); | |
| 416 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 417 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 418 | by (auto_tac (claset() addSDs [conjI RS realpow_Suc_less] | |
| 419 | addEs [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset], | |
| 420 | simpset() addsimps [hyperpow,hypreal_less,hypnat_add])); | |
| 421 | qed_spec_mp "hyperpow_Suc_less"; | |
| 422 | ||
| 423 | Goalw [hypnat_one_def] | |
| 424 | "#0 <= r & r < (#1::hypreal) --> r pow (n + 1hn) <= r pow n"; | |
| 425 | by (full_simp_tac | |
| 426 | (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def, | |
| 427 | one_eq_numeral_1 RS sym, hypreal_one_def]) 1); | |
| 428 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 429 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 430 | by (auto_tac (claset() addSDs [conjI RS realpow_Suc_le] addEs | |
| 431 | [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset ], | |
| 432 | simpset() addsimps [hyperpow,hypreal_le,hypnat_add, | |
| 433 | hypreal_less])); | |
| 434 | qed_spec_mp "hyperpow_Suc_le"; | |
| 435 | ||
| 436 | Goalw [hypnat_one_def] | |
| 437 | "(#0::hypreal) <= r & r < #1 & n < N --> r pow N <= r pow n"; | |
| 438 | by (full_simp_tac | |
| 439 | (HOL_ss addsimps [zero_eq_numeral_0 RS sym, hypreal_zero_def, | |
| 440 | one_eq_numeral_1 RS sym, hypreal_one_def]) 1); | |
| 441 | by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
 | |
| 442 | by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
 | |
| 443 | by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
 | |
| 444 | by (auto_tac (claset(), | |
| 445 | simpset() addsimps [hyperpow, hypreal_le,hypreal_less,hypnat_less])); | |
| 446 | by (etac (FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset) 1); | |
| 447 | by (etac FreeUltrafilterNat_Int 1); | |
| 448 | by (auto_tac (claset() addSDs [conjI RS realpow_less_le], | |
| 449 | simpset())); | |
| 450 | qed_spec_mp "hyperpow_less_le"; | |
| 451 | ||
| 452 | Goal "[| (#0::hypreal) <= r; r < #1 |] \ | |
| 453 | \ ==> ALL N n. n < N --> r pow N <= r pow n"; | |
| 454 | by (blast_tac (claset() addSIs [hyperpow_less_le]) 1); | |
| 455 | qed "hyperpow_less_le2"; | |
| 456 | ||
| 457 | Goal "[| #0 <= r; r < (#1::hypreal); N : HNatInfinite |] \ | |
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10834diff
changeset | 458 | \ ==> ALL n: Nats. r pow N <= r pow n"; | 
| 10751 | 459 | by (auto_tac (claset() addSIs [hyperpow_less_le], | 
| 460 | simpset() addsimps [HNatInfinite_iff])); | |
| 461 | qed "hyperpow_SHNat_le"; | |
| 462 | ||
| 463 | Goalw [hypreal_of_real_def,hypnat_of_nat_def] | |
| 464 | "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)"; | |
| 465 | by (auto_tac (claset(), simpset() addsimps [hyperpow])); | |
| 466 | qed "hyperpow_realpow"; | |
| 467 | ||
| 468 | Goalw [SReal_def] | |
| 10919 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 paulson parents: 
10834diff
changeset | 469 | "(hypreal_of_real r) pow (hypnat_of_nat n) : Reals"; | 
| 10751 | 470 | by (auto_tac (claset(), simpset() addsimps [hyperpow_realpow])); | 
| 471 | qed "hyperpow_SReal"; | |
| 472 | Addsimps [hyperpow_SReal]; | |
| 473 | ||
| 474 | Goal "N : HNatInfinite ==> (#0::hypreal) pow N = 0"; | |
| 475 | by (dtac HNatInfinite_is_Suc 1); | |
| 476 | by (Auto_tac); | |
| 477 | qed "hyperpow_zero_HNatInfinite"; | |
| 478 | Addsimps [hyperpow_zero_HNatInfinite]; | |
| 479 | ||
| 480 | Goal "[| (#0::hypreal) <= r; r < #1; n <= N |] ==> r pow N <= r pow n"; | |
| 481 | by (dres_inst_tac [("y","N")] hypnat_le_imp_less_or_eq 1);
 | |
| 482 | by (auto_tac (claset() addIs [hyperpow_less_le], simpset())); | |
| 483 | qed "hyperpow_le_le"; | |
| 484 | ||
| 485 | Goal "[| (#0::hypreal) < r; r < #1 |] ==> r pow (n + 1hn) <= r"; | |
| 486 | by (dres_inst_tac [("n","1hn")] (order_less_imp_le RS hyperpow_le_le) 1);
 | |
| 487 | by (Auto_tac); | |
| 488 | qed "hyperpow_Suc_le_self"; | |
| 489 | ||
| 490 | Goal "[| (#0::hypreal) <= r; r < #1 |] ==> r pow (n + 1hn) <= r"; | |
| 491 | by (dres_inst_tac [("n","1hn")] hyperpow_le_le 1);
 | |
| 492 | by (Auto_tac); | |
| 493 | qed "hyperpow_Suc_le_self2"; | |
| 494 | ||
| 495 | Goalw [Infinitesimal_def] | |
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 496 | "[| x : Infinitesimal; 0 < N |] ==> abs (x pow N) <= abs x"; | 
| 10751 | 497 | by (auto_tac (claset() addSIs [hyperpow_Suc_le_self2], | 
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 498 | simpset() addsimps [hyperpow_hrabs RS sym, | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 499 | hypnat_gt_zero_iff2, hrabs_ge_zero])); | 
| 10751 | 500 | qed "lemma_Infinitesimal_hyperpow"; | 
| 501 | ||
| 502 | Goal "[| x : Infinitesimal; 0 < N |] ==> x pow N : Infinitesimal"; | |
| 503 | by (rtac hrabs_le_Infinitesimal 1); | |
| 10778 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 504 | by (rtac lemma_Infinitesimal_hyperpow 2); | 
| 
2c6605049646
more tidying, especially to remove real_of_posnat
 paulson parents: 
10751diff
changeset | 505 | by Auto_tac; | 
| 10751 | 506 | qed "Infinitesimal_hyperpow"; | 
| 507 | ||
| 508 | Goalw [hypnat_of_nat_def] | |
| 509 | "(x ^ n : Infinitesimal) = (x pow (hypnat_of_nat n) : Infinitesimal)"; | |
| 510 | by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
 | |
| 511 | by (auto_tac (claset(), simpset() addsimps [hrealpow, hyperpow])); | |
| 512 | qed "hrealpow_hyperpow_Infinitesimal_iff"; | |
| 513 | ||
| 514 | Goal "[| x : Infinitesimal; 0 < n |] ==> x ^ n : Infinitesimal"; | |
| 515 | by (auto_tac (claset() addSIs [Infinitesimal_hyperpow], | |
| 516 | simpset() addsimps [hrealpow_hyperpow_Infinitesimal_iff, | |
| 517 | hypnat_of_nat_less_iff,hypnat_of_nat_zero] | |
| 518 | delsimps [hypnat_of_nat_less_iff RS sym])); | |
| 519 | qed "Infinitesimal_hrealpow"; |