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