src/HOL/Real/Hyperreal/HyperPow.ML
author wenzelm
Fri, 06 Oct 2000 17:35:58 +0200
changeset 10168 50be659d4222
parent 10045 c76b73e16711
child 10607 352f6f209775
permissions -rw-r--r--
final tuning;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10045
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     1
(*  Title       : HyperPow.ML
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot  
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     4
    Description : Natural Powers of hyperreals theory
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     5
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     6
*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     7
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     8
Goal "(0::hypreal) ^ (Suc n) = 0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
     9
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    10
qed "hrealpow_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    11
Addsimps [hrealpow_zero];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    12
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    13
Goal "r ~= (0::hypreal) --> r ^ n ~= 0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    14
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    15
by (auto_tac (claset() addIs [hypreal_mult_not_0E],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    16
    simpset() addsimps [hypreal_zero_not_eq_one RS not_sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    17
qed_spec_mp "hrealpow_not_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    18
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    19
Goal "r ~= (0::hypreal) --> hrinv(r ^ n) = (hrinv r) ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    20
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    21
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    22
by (forw_inst_tac [("n","n")] hrealpow_not_zero 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    23
by (auto_tac (claset() addDs [hrinv_mult_eq],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    24
    simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    25
qed_spec_mp "hrealpow_hrinv";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    26
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    27
Goal "abs (r::hypreal) ^ n = abs (r ^ n)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    28
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    29
by (auto_tac (claset(),simpset() addsimps [hrabs_mult,hrabs_one]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    30
qed "hrealpow_hrabs";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    31
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    32
Goal "(r::hypreal) ^ (n + m) = (r ^ n) * (r ^ m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    33
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    34
by (auto_tac (claset(),simpset() addsimps hypreal_mult_ac));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    35
qed "hrealpow_add";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    36
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    37
Goal "(r::hypreal) ^ 1 = r";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    38
by (Simp_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    39
qed "hrealpow_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    40
Addsimps [hrealpow_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    41
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    42
Goal "(r::hypreal) ^ 2 = r * r";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    43
by (Simp_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    44
qed "hrealpow_two";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    45
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    46
Goal "(0::hypreal) < r --> 0 <= r ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    47
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    48
by (auto_tac (claset() addDs [hypreal_less_imp_le] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    49
    addIs [hypreal_le_mult_order],simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    50
    [hypreal_zero_less_one RS hypreal_less_imp_le]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    51
qed_spec_mp "hrealpow_ge_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    52
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    53
Goal "(0::hypreal) < r --> 0 < r ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    54
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    55
by (auto_tac (claset() addIs [hypreal_mult_order],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    56
    simpset() addsimps [hypreal_zero_less_one]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    57
qed_spec_mp "hrealpow_gt_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    58
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    59
Goal "(0::hypreal) <= r --> 0 <= r ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    60
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    61
by (auto_tac (claset() addIs [hypreal_le_mult_order],simpset() 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    62
    addsimps [hypreal_zero_less_one RS hypreal_less_imp_le]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    63
qed_spec_mp "hrealpow_ge_zero2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    64
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    65
Goal "(0::hypreal) < x & x <= y --> x ^ n <= y ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    66
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    67
by (auto_tac (claset() addSIs [hypreal_mult_le_mono],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    68
    simpset() addsimps [hypreal_le_refl]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    69
by (asm_simp_tac (simpset() addsimps [hrealpow_ge_zero]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    70
qed_spec_mp "hrealpow_le";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    71
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    72
Goal "(0::hypreal) < x & x < y & 0 < n --> x ^ n < y ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    73
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    74
by (auto_tac (claset() addIs [hypreal_mult_less_mono,gr0I] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    75
    addDs [hrealpow_gt_zero],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    76
qed "hrealpow_less";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    77
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    78
Goal "1hr ^ n = 1hr";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    79
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    80
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    81
qed "hrealpow_eq_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    82
Addsimps [hrealpow_eq_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    83
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    84
Goal "abs(-(1hr ^ n)) = 1hr";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    85
by (simp_tac (simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    86
    [hrabs_minus_cancel,hrabs_one]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    87
qed "hrabs_minus_hrealpow_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    88
Addsimps [hrabs_minus_hrealpow_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    89
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    90
Goal "abs((-1hr) ^ n) = 1hr";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    91
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    92
by (auto_tac (claset(),simpset() addsimps [hrabs_mult,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    93
         hrabs_minus_cancel,hrabs_one]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    94
qed "hrabs_hrealpow_minus_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    95
Addsimps [hrabs_hrealpow_minus_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    96
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    97
Goal "((r::hypreal) * s) ^ n = (r ^ n) * (s ^ n)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    98
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
    99
by (auto_tac (claset(),simpset() addsimps hypreal_mult_ac));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   100
qed "hrealpow_mult";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   101
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   102
Goal "(0::hypreal) <= r ^ 2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   103
by (simp_tac (simpset() addsimps [hrealpow_two]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   104
qed "hrealpow_two_le";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   105
Addsimps [hrealpow_two_le];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   106
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   107
Goal "(0::hypreal) <= u ^ 2 + v ^ 2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   108
by (auto_tac (claset() addIs [hypreal_le_add_order],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   109
qed "hrealpow_two_le_add_order";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   110
Addsimps [hrealpow_two_le_add_order];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   111
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   112
Goal "(0::hypreal) <= u ^ 2 + v ^ 2 + w ^ 2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   113
by (auto_tac (claset() addSIs [hypreal_le_add_order],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   114
qed "hrealpow_two_le_add_order2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   115
Addsimps [hrealpow_two_le_add_order2];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   116
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   117
(* See HYPER.ML *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   118
Goal "(x ^ 2 + y ^ 2 + z ^ 2 = (0::hypreal)) = \ 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   119
\               (x = 0 & y = 0 & z = 0)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   120
by (simp_tac (simpset() addsimps [hrealpow_two]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   121
qed "hrealpow_three_squares_add_zero_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   122
Addsimps [hrealpow_three_squares_add_zero_iff];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   123
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   124
Goal "abs(x ^ 2) = (x::hypreal) ^ 2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   125
by (simp_tac (simpset() addsimps [hrabs_eqI1]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   126
qed "hrabs_hrealpow_two";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   127
Addsimps [hrabs_hrealpow_two];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   128
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   129
Goal "abs(x) ^ 2 = (x::hypreal) ^ 2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   130
by (simp_tac (simpset() addsimps [hrealpow_hrabs,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   131
    hrabs_eqI1] delsimps [hpowr_Suc]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   132
qed "hrealpow_two_hrabs";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   133
Addsimps [hrealpow_two_hrabs];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   134
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   135
Goal "!!r. 1hr < r ==> 1hr < r ^ 2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   136
by (auto_tac (claset(),simpset() addsimps [hrealpow_two]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   137
by (cut_facts_tac [hypreal_zero_less_one] 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   138
by (forw_inst_tac [("R1.0","0")] hypreal_less_trans 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   139
by (assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   140
by (dres_inst_tac [("z","r"),("x","1hr")] hypreal_mult_less_mono1 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   141
by (auto_tac (claset() addIs [hypreal_less_trans],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   142
qed "hrealpow_two_gt_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   143
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   144
Goal "!!r. 1hr <= r ==> 1hr <= r ^ 2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   145
by (etac (hypreal_le_imp_less_or_eq RS disjE) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   146
by (etac (hrealpow_two_gt_one RS hypreal_less_imp_le) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   147
by (asm_simp_tac (simpset() addsimps [hypreal_le_refl]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   148
qed "hrealpow_two_ge_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   149
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   150
Goal "!!r. (0::hypreal) < r ==> 0 < r ^ Suc n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   151
by (forw_inst_tac [("n","n")] hrealpow_ge_zero 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   152
by (forw_inst_tac [("n1","n")]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   153
    ((hypreal_not_refl2 RS not_sym) RS hrealpow_not_zero RS not_sym) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   154
by (auto_tac (claset() addSDs [hypreal_le_imp_less_or_eq]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   155
     addIs [hypreal_mult_order],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   156
qed "hrealpow_Suc_gt_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   157
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   158
Goal "!!r. (0::hypreal) <= r ==> 0 <= r ^ Suc n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   159
by (etac (hypreal_le_imp_less_or_eq RS disjE) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   160
by (etac (hrealpow_ge_zero) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   161
by (asm_simp_tac (simpset() addsimps [hypreal_le_refl]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   162
qed "hrealpow_Suc_ge_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   163
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   164
Goal "1hr <= (1hr +1hr) ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   165
by (res_inst_tac [("j","1hr ^ n")] hypreal_le_trans 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   166
by (rtac hrealpow_le 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   167
by (auto_tac (claset() addIs [hypreal_less_imp_le],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   168
    simpset() addsimps [hypreal_zero_less_one,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   169
    hypreal_one_less_two,hypreal_le_refl]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   170
qed "two_hrealpow_ge_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   171
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   172
Goal "hypreal_of_nat n < (1hr + 1hr) ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   173
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   174
by (auto_tac (claset(),simpset() addsimps [hypreal_of_nat_Suc,hypreal_of_nat_zero,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   175
    hypreal_zero_less_one,hypreal_add_mult_distrib,hypreal_of_nat_one]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   176
by (blast_tac (claset() addSIs [hypreal_add_less_le_mono,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   177
    two_hrealpow_ge_one]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   178
qed "two_hrealpow_gt";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   179
Addsimps [two_hrealpow_gt,two_hrealpow_ge_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   180
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   181
Goal "(-1hr) ^ (2*n) = 1hr";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   182
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   183
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   184
qed "hrealpow_minus_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   185
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   186
Goal "(-1hr) ^ (n + n) = 1hr";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   187
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   188
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   189
qed "hrealpow_minus_one2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   190
Addsimps [hrealpow_minus_one2];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   191
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   192
Goal "(-(x::hypreal)) ^ 2 = x ^ 2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   193
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   194
qed "hrealpow_minus_two";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   195
Addsimps [hrealpow_minus_two];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   196
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   197
Goal "(0::hypreal) < r & r < 1hr --> r ^ Suc n < r ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   198
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   199
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   200
        [hypreal_mult_less_mono2]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   201
qed_spec_mp "hrealpow_Suc_less";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   202
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   203
Goal "(0::hypreal) <= r & r < 1hr --> r ^ Suc n <= r ^ n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   204
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   205
by (auto_tac (claset() addIs [hypreal_less_imp_le] addSDs
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   206
     [hypreal_le_imp_less_or_eq,hrealpow_Suc_less],simpset()
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   207
     addsimps [hypreal_le_refl,hypreal_mult_less_mono2]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   208
qed_spec_mp "hrealpow_Suc_le";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   209
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   210
(* a few more theorems needed here *)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   211
Goal "1hr <= r --> r ^ n <= r ^ Suc n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   212
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   213
by (auto_tac (claset() addSIs [hypreal_mult_le_le_mono1],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   214
by (rtac ccontr 1 THEN dtac not_hypreal_leE 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   215
by (dtac hypreal_le_less_trans 1 THEN assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   216
by (etac (hypreal_zero_less_one RS hypreal_less_asym) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   217
qed "hrealpow_less_Suc";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   218
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   219
Goal "Abs_hypreal(hyprel^^{%n. X n}) ^ m = Abs_hypreal(hyprel^^{%n. (X n) ^ m})";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   220
by (nat_ind_tac "m" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   221
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   222
    [hypreal_one_def,hypreal_mult]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   223
qed "hrealpow";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   224
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   225
Goal "(x + (y::hypreal)) ^ 2 = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   226
\     x ^ 2 + y ^ 2 + (hypreal_of_nat 2)*x*y";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   227
by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib2,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   228
    hypreal_add_mult_distrib,hypreal_of_nat_two] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   229
    @ hypreal_add_ac @ hypreal_mult_ac) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   230
qed "hrealpow_sum_square_expand";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   231
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   232
(*---------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   233
   we'll prove the following theorem by going down to the
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   234
   level of the ultrafilter and relying on the analogous
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   235
   property for the real rather than prove it directly 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   236
   using induction: proof is much simpler this way!
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   237
 ---------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   238
Goalw [hypreal_zero_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   239
  "[|(0::hypreal) <= x;0 <= y;x ^ Suc n <= y ^ Suc n |] ==> x <= y";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   240
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   241
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   242
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   243
    [hrealpow,hypreal_le,hypreal_mult]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   244
by (ultra_tac (claset() addIs [realpow_increasing],simpset()) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   245
qed "hrealpow_increasing";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   246
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   247
goalw HyperPow.thy [hypreal_zero_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   248
  "!!x. [|(0::hypreal) <= x;0 <= y;x ^ Suc n = y ^ Suc n |] ==> x = y";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   249
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   250
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   251
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   252
    [hrealpow,hypreal_mult,hypreal_le]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   253
by (ultra_tac (claset() addIs [realpow_Suc_cancel_eq],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   254
    simpset()) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   255
qed "hrealpow_Suc_cancel_eq";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   256
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   257
Goal "x : HFinite --> x ^ n : HFinite";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   258
by (induct_tac "n" 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   259
by (auto_tac (claset() addIs [HFinite_mult],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   260
qed_spec_mp "hrealpow_HFinite";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   261
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   262
(*---------------------------------------------------------------
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   263
                  Hypernaturals Powers
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   264
 --------------------------------------------------------------*)
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   265
Goalw [congruent_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   266
     "congruent hyprel \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   267
\    (%X Y. hyprel^^{%n. ((X::nat=>real) n ^ (Y::nat=>nat) n)})";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   268
by (safe_tac (claset() addSIs [ext]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   269
by (ALLGOALS(Fuf_tac));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   270
qed "hyperpow_congruent";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   271
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   272
Goalw [hyperpow_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   273
  "Abs_hypreal(hyprel^^{%n. X n}) pow Abs_hypnat(hypnatrel^^{%n. Y n}) = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   274
\  Abs_hypreal(hyprel^^{%n. X n ^ Y n})";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   275
by (res_inst_tac [("f","Abs_hypreal")] arg_cong 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   276
by (auto_tac (claset() addSIs [lemma_hyprel_refl,bexI],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   277
    simpset() addsimps [hyprel_in_hypreal RS 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   278
    Abs_hypreal_inverse,equiv_hyprel,hyperpow_congruent]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   279
by (Fuf_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   280
qed "hyperpow";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   281
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   282
Goalw [hypreal_zero_def,hypnat_one_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   283
      "(0::hypreal) pow (n + 1hn) = 0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   284
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   285
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   286
    [hyperpow,hypnat_add]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   287
qed "hyperpow_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   288
Addsimps [hyperpow_zero];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   289
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   290
Goalw [hypreal_zero_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   291
      "r ~= (0::hypreal) --> r pow n ~= 0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   292
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   293
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   294
by (auto_tac (claset(),simpset() addsimps [hyperpow]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   295
by (dtac FreeUltrafilterNat_Compl_mem 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   296
by (fuf_empty_tac (claset() addIs [realpow_not_zero RS notE],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   297
    simpset()) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   298
qed_spec_mp "hyperpow_not_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   299
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   300
Goalw [hypreal_zero_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   301
      "r ~= (0::hypreal) --> hrinv(r pow n) = (hrinv r) pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   302
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   303
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   304
by (auto_tac (claset() addSDs [FreeUltrafilterNat_Compl_mem],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   305
    simpset() addsimps [hypreal_hrinv,hyperpow]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   306
by (rtac FreeUltrafilterNat_subset 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   307
by (auto_tac (claset() addDs [realpow_not_zero] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   308
    addIs [realpow_rinv],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   309
qed "hyperpow_hrinv";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   310
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   311
Goal "abs r pow n = abs (r pow n)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   312
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   313
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   314
by (auto_tac (claset(),simpset() addsimps [hypreal_hrabs,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   315
    hyperpow,realpow_abs]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   316
qed "hyperpow_hrabs";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   317
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   318
Goal "r pow (n + m) = (r pow n) * (r pow m)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   319
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   320
by (res_inst_tac [("z","m")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   321
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   322
by (auto_tac (claset(),simpset() addsimps [hyperpow,hypnat_add,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   323
     hypreal_mult,realpow_add]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   324
qed "hyperpow_add";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   325
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   326
Goalw [hypnat_one_def] "r pow 1hn = r";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   327
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   328
by (auto_tac (claset(),simpset() addsimps [hyperpow]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   329
qed "hyperpow_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   330
Addsimps [hyperpow_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   331
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   332
Goalw [hypnat_one_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   333
      "r pow (1hn + 1hn) = r * r";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   334
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   335
by (auto_tac (claset(),simpset() addsimps [hyperpow,hypnat_add,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   336
     hypreal_mult,realpow_two]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   337
qed "hyperpow_two";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   338
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   339
Goalw [hypreal_zero_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   340
      "(0::hypreal) < r --> 0 < r pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   341
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   342
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   343
by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   344
    realpow_gt_zero],simpset() addsimps [hyperpow,hypreal_less,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   345
    hypreal_le]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   346
qed_spec_mp "hyperpow_gt_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   347
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   348
Goal "(0::hypreal) < r --> 0 <= r pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   349
by (blast_tac (claset() addSIs [hyperpow_gt_zero,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   350
    hypreal_less_imp_le]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   351
qed_spec_mp "hyperpow_ge_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   352
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   353
Goalw [hypreal_zero_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   354
      "(0::hypreal) <= r --> 0 <= r pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   355
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   356
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   357
by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   358
    realpow_ge_zero2],simpset() addsimps [hyperpow,hypreal_le]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   359
qed "hyperpow_ge_zero2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   360
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   361
Goalw [hypreal_zero_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   362
      "(0::hypreal) < x & x <= y --> x pow n <= y pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   363
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   364
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   365
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   366
by (auto_tac (claset() addIs [realpow_le,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   367
    (FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset)],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   368
    simpset() addsimps [hyperpow,hypreal_le,hypreal_less]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   369
qed_spec_mp "hyperpow_le";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   370
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   371
Goalw [hypreal_one_def] "1hr pow n = 1hr";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   372
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   373
by (auto_tac (claset(),simpset() addsimps [hyperpow]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   374
qed "hyperpow_eq_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   375
Addsimps [hyperpow_eq_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   376
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   377
Goalw [hypreal_one_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   378
      "abs(-(1hr pow n)) = 1hr";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   379
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   380
by (auto_tac (claset(),simpset() addsimps [abs_one,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   381
    hrabs_minus_cancel,hyperpow,hypreal_hrabs]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   382
qed "hrabs_minus_hyperpow_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   383
Addsimps [hrabs_minus_hyperpow_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   384
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   385
Goalw [hypreal_one_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   386
      "abs((-1hr) pow n) = 1hr";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   387
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   388
by (auto_tac (claset(),simpset() addsimps 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   389
    [hyperpow,hypreal_minus,hypreal_hrabs]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   390
qed "hrabs_hyperpow_minus_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   391
Addsimps [hrabs_hyperpow_minus_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   392
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   393
Goal "(r * s) pow n = (r pow n) * (s pow n)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   394
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   395
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   396
by (res_inst_tac [("z","s")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   397
by (auto_tac (claset(),simpset() addsimps [hyperpow,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   398
    hypreal_mult,realpow_mult]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   399
qed "hyperpow_mult";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   400
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   401
Goal "(0::hypreal) <= r pow (1hn + 1hn)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   402
by (simp_tac (simpset() addsimps [hyperpow_two]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   403
qed "hyperpow_two_le";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   404
Addsimps [hyperpow_two_le];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   405
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   406
Goal "abs(x pow (1hn + 1hn)) = x pow (1hn + 1hn)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   407
by (simp_tac (simpset() addsimps [hrabs_eqI1]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   408
qed "hrabs_hyperpow_two";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   409
Addsimps [hrabs_hyperpow_two];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   410
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   411
Goal "abs(x) pow (1hn + 1hn)  = x pow (1hn + 1hn)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   412
by (simp_tac (simpset() addsimps [hyperpow_hrabs,hrabs_eqI1]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   413
qed "hyperpow_two_hrabs";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   414
Addsimps [hyperpow_two_hrabs];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   415
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   416
Goal "!!r. 1hr < r ==> 1hr < r pow (1hn + 1hn)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   417
by (auto_tac (claset(),simpset() addsimps [hyperpow_two]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   418
by (cut_facts_tac [hypreal_zero_less_one] 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   419
by (forw_inst_tac [("R1.0","0")] hypreal_less_trans 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   420
by (assume_tac 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   421
by (dres_inst_tac [("z","r"),("x","1hr")] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   422
    hypreal_mult_less_mono1 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   423
by (auto_tac (claset() addIs [hypreal_less_trans],simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   424
qed "hyperpow_two_gt_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   425
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   426
Goal "!!r. 1hr <= r ==> 1hr <= r pow (1hn + 1hn)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   427
by (auto_tac (claset() addSDs [hypreal_le_imp_less_or_eq] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   428
     addIs [hyperpow_two_gt_one,hypreal_less_imp_le],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   429
     simpset() addsimps [hypreal_le_refl]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   430
qed "hyperpow_two_ge_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   431
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   432
Goalw [hypnat_one_def,hypreal_zero_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   433
      "!!r. (0::hypreal) < r ==> 0 < r pow (n + 1hn)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   434
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   435
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   436
by (auto_tac (claset() addSEs [FreeUltrafilterNat_subset]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   437
    addDs [realpow_Suc_gt_zero],simpset() addsimps [hyperpow,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   438
    hypreal_less,hypnat_add]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   439
qed "hyperpow_Suc_gt_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   440
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   441
Goal "!!r. (0::hypreal) <= r ==> 0 <= r pow (n + 1hn)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   442
by (auto_tac (claset() addSDs [hypreal_le_imp_less_or_eq] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   443
    addIs [hyperpow_ge_zero,hypreal_less_imp_le], 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   444
    simpset() addsimps [hypreal_le_refl]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   445
qed "hyperpow_Suc_ge_zero";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   446
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   447
Goal "1hr <= (1hr +1hr) pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   448
by (res_inst_tac [("j","1hr pow n")] hypreal_le_trans 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   449
by (rtac hyperpow_le 2);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   450
by (auto_tac (claset() addIs [hypreal_less_imp_le],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   451
    simpset() addsimps [hypreal_zero_less_one,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   452
    hypreal_one_less_two,hypreal_le_refl]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   453
qed "two_hyperpow_ge_one";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   454
Addsimps [two_hyperpow_ge_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   455
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   456
Addsimps [simplify (simpset()) realpow_minus_one];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   457
Goalw [hypreal_one_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   458
      "(-1hr) pow ((1hn + 1hn)*n) = 1hr";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   459
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   460
by (auto_tac (claset(),simpset() addsimps [hyperpow,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   461
    hypnat_add,hypreal_minus]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   462
qed "hyperpow_minus_one2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   463
Addsimps [hyperpow_minus_one2];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   464
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   465
Goalw [hypreal_zero_def,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   466
      hypreal_one_def,hypnat_one_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   467
     "(0::hypreal) < r & r < 1hr --> r pow (n + 1hn) < r pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   468
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   469
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   470
by (auto_tac (claset() addSDs [conjI RS realpow_Suc_less] addEs
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   471
    [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset ],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   472
    simpset() addsimps [hyperpow,hypreal_less,hypnat_add]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   473
qed_spec_mp "hyperpow_Suc_less";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   474
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   475
Goalw [hypreal_zero_def,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   476
      hypreal_one_def,hypnat_one_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   477
     "0 <= r & r < 1hr --> r pow (n + 1hn) <= r pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   478
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   479
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   480
by (auto_tac (claset() addSDs [conjI RS realpow_Suc_le] addEs
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   481
    [FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset ],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   482
    simpset() addsimps [hyperpow,hypreal_le,hypnat_add,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   483
    hypreal_less]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   484
qed_spec_mp "hyperpow_Suc_le";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   485
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   486
Goalw [hypreal_zero_def,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   487
      hypreal_one_def,hypnat_one_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   488
     "(0::hypreal) <= r & r < 1hr & n < N --> r pow N <= r pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   489
by (res_inst_tac [("z","n")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   490
by (res_inst_tac [("z","N")] eq_Abs_hypnat 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   491
by (res_inst_tac [("z","r")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   492
by (auto_tac (claset(),simpset() addsimps [hyperpow,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   493
    hypreal_le,hypreal_less,hypnat_less]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   494
by (etac (FreeUltrafilterNat_Int RS FreeUltrafilterNat_subset) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   495
by (etac FreeUltrafilterNat_Int 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   496
by (auto_tac (claset() addSDs [conjI RS realpow_less_le],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   497
    simpset()));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   498
qed_spec_mp "hyperpow_less_le";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   499
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   500
Goal "!!r. [| (0::hypreal) <= r; r < 1hr |] ==> \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   501
\          ALL N n. n < N --> r pow N <= r pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   502
by (blast_tac (claset() addSIs [hyperpow_less_le]) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   503
qed "hyperpow_less_le2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   504
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   505
Goal "!!r. [| 0 <= r; r < 1hr; \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   506
\             N : HNatInfinite \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   507
\          |] ==> ALL n:SHNat. r pow N <= r pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   508
by (auto_tac (claset() addSIs [hyperpow_less_le],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   509
              simpset() addsimps [HNatInfinite_iff]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   510
qed "hyperpow_SHNat_le";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   511
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   512
Goalw [hypreal_of_real_def,hypnat_of_nat_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   513
      "(hypreal_of_real r) pow (hypnat_of_nat n) = hypreal_of_real (r ^ n)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   514
by (auto_tac (claset(),simpset() addsimps [hyperpow]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   515
qed "hyperpow_realpow";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   516
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   517
Goalw [SReal_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   518
     "(hypreal_of_real r) pow (hypnat_of_nat n) : SReal";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   519
by (auto_tac (claset(),simpset() addsimps [hyperpow_realpow]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   520
qed "hyperpow_SReal";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   521
Addsimps [hyperpow_SReal];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   522
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   523
Goal "!!N. N : HNatInfinite ==> (0::hypreal) pow N = 0";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   524
by (dtac HNatInfinite_is_Suc 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   525
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   526
qed "hyperpow_zero_HNatInfinite";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   527
Addsimps [hyperpow_zero_HNatInfinite];
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   528
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   529
Goal "!!r. [| (0::hypreal) <= r; r < 1hr; n <= N |] ==> r pow N <= r pow n";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   530
by (dres_inst_tac [("y","N")] hypnat_le_imp_less_or_eq 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   531
by (auto_tac (claset() addIs [hyperpow_less_le],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   532
    simpset() addsimps [hypreal_le_refl]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   533
qed "hyperpow_le_le";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   534
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   535
Goal "!!r. [| (0::hypreal) < r; r < 1hr |] ==> r pow (n + 1hn) <= r";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   536
by (dres_inst_tac [("n","1hn")] (hypreal_less_imp_le RS 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   537
    hyperpow_le_le) 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   538
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   539
qed "hyperpow_Suc_le_self";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   540
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   541
Goal "!!r. [| (0::hypreal) <= r; r < 1hr |] ==> r pow (n + 1hn) <= r";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   542
by (dres_inst_tac [("n","1hn")] hyperpow_le_le 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   543
by (Auto_tac);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   544
qed "hyperpow_Suc_le_self2";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   545
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   546
Goalw [Infinitesimal_def]
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   547
     "!!x. [| x : Infinitesimal; 0 < N |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   548
\          ==> (abs x) pow N <= abs x";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   549
by (auto_tac (claset() addSIs [hyperpow_Suc_le_self2],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   550
    simpset() addsimps [hyperpow_hrabs RS sym,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   551
    hypnat_gt_zero_iff2,hrabs_ge_zero,SReal_one,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   552
    hypreal_zero_less_one]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   553
qed "lemma_Infinitesimal_hyperpow";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   554
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   555
Goal "!!x. [| x : Infinitesimal; 0 < N |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   556
\           ==> x pow N : Infinitesimal";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   557
by (rtac hrabs_le_Infinitesimal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   558
by (dtac Infinitesimal_hrabs 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   559
by (auto_tac (claset() addSIs [lemma_Infinitesimal_hyperpow],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   560
    simpset() addsimps [hyperpow_hrabs RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   561
qed "Infinitesimal_hyperpow";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   562
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   563
goalw HyperPow.thy [hypnat_of_nat_def] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   564
     "(x ^ n : Infinitesimal) = \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   565
\     (x pow (hypnat_of_nat n) : Infinitesimal)";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   566
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   567
by (auto_tac (claset(),simpset() addsimps [hrealpow,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   568
    hyperpow]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   569
qed "hrealpow_hyperpow_Infinitesimal_iff";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   570
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   571
goal HyperPow.thy 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   572
     "!!x. [| x : Infinitesimal; 0 < n |] \
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   573
\           ==> x ^ n : Infinitesimal";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   574
by (auto_tac (claset() addSIs [Infinitesimal_hyperpow],
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   575
    simpset() addsimps [hrealpow_hyperpow_Infinitesimal_iff,
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   576
    hypnat_of_nat_less_iff,hypnat_of_nat_zero] 
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   577
    delsimps [hypnat_of_nat_less_iff RS sym]));
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   578
qed "Infinitesimal_hrealpow";
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   579
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   580
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   581
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   582
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   583
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   584
c76b73e16711 New theories: construction of hypernaturals, nonstandard extensions,
fleuriot
parents:
diff changeset
   585