src/HOL/Real/RealPow.ML
author kleing
Tue, 13 May 2003 08:59:21 +0200
changeset 14024 213dcc39358f
parent 12613 279facb4253a
permissions -rw-r--r--
HOL-Real -> HOL-Complex
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     1
(*  Title       : RealPow.ML
7219
4e3f386c2e37 inserted Id: lines
paulson
parents: 7077
diff changeset
     2
    ID          : $Id$
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     3
    Author      : Jacques D. Fleuriot  
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     4
    Copyright   : 1998  University of Cambridge
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     5
    Description : Natural Powers of reals theory
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
     6
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
     7
FIXME: most of the theorems for Suc (Suc 0) are redundant!
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     8
*)
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
     9
10690
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
    10
bind_thm ("realpow_Suc", thm "realpow_Suc");
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
    11
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    12
Goal "(0::real) ^ (Suc n) = 0";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
    13
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    14
qed "realpow_zero";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    15
Addsimps [realpow_zero];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    16
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    17
Goal "r ~= (0::real) --> r ^ n ~= 0";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    18
by (induct_tac "n" 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    19
by Auto_tac; 
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    20
qed_spec_mp "realpow_not_zero";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    21
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    22
Goal "r ^ n = (0::real) ==> r = 0";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
    23
by (rtac ccontr 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
    24
by (auto_tac (claset() addDs [realpow_not_zero], simpset()));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    25
qed "realpow_zero_zero";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    26
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10606
diff changeset
    27
Goal "inverse ((r::real) ^ n) = (inverse r) ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    28
by (induct_tac "n" 1);
10648
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10606
diff changeset
    29
by (auto_tac (claset(), simpset() addsimps [real_inverse_distrib]));
a8c647cfa31f first stage in tidying up Real and Hyperreal.
paulson
parents: 10606
diff changeset
    30
qed "realpow_inverse";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    31
12330
c69bee072501 *** empty log message ***
nipkow
parents: 12196
diff changeset
    32
Goal "abs(r ^ n) = abs(r::real) ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    33
by (induct_tac "n" 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
    34
by (auto_tac (claset(), simpset() addsimps [abs_mult]));
8838
4eaa99f0d223 replaced rabs by overloaded abs;
wenzelm
parents: 8442
diff changeset
    35
qed "realpow_abs";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    36
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    37
Goal "(r::real) ^ (n + m) = (r ^ n) * (r ^ m)";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    38
by (induct_tac "n" 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    39
by (auto_tac (claset(),simpset() addsimps real_mult_ac));
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    40
qed "realpow_add";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    41
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    42
Goal "(r::real) ^ 1 = r";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    43
by (Simp_tac 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    44
qed "realpow_one";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    45
Addsimps [realpow_one];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    46
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10919
diff changeset
    47
Goal "(r::real)^ (Suc (Suc 0)) = r * r";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    48
by (Simp_tac 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    49
qed "realpow_two";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    50
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    51
Goal "(0::real) < r --> 0 < r ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    52
by (induct_tac "n" 1);
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    53
by (auto_tac (claset() addIs [real_mult_order],
9070
99d93349914b full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
    54
	      simpset() addsimps [real_zero_less_one]));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    55
qed_spec_mp "realpow_gt_zero";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    56
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    57
Goal "(0::real) <= r --> 0 <= r ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    58
by (induct_tac "n" 1);
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
    59
by (auto_tac (claset(), simpset() addsimps [real_0_le_mult_iff]));
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
    60
qed_spec_mp "realpow_ge_zero";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    61
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    62
Goal "(0::real) <= x & x <= y --> x ^ n <= y ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    63
by (induct_tac "n" 1);
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
    64
by (auto_tac (claset() addSIs [real_mult_le_mono], simpset()));
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
    65
by (asm_simp_tac (simpset() addsimps [realpow_ge_zero]) 1);
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
    66
qed_spec_mp "realpow_le";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    67
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    68
Goal "(0::real) < x & x < y & 0 < n --> x ^ n < y ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    69
by (induct_tac "n" 1);
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    70
by (auto_tac (claset() addIs [real_mult_less_mono, gr0I] 
9070
99d93349914b full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
    71
                       addDs [realpow_gt_zero],
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
    72
    simpset()));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    73
qed_spec_mp "realpow_less";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    74
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    75
Goal "1 ^ n = (1::real)";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    76
by (induct_tac "n" 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
    77
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    78
qed "realpow_eq_one";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    79
Addsimps [realpow_eq_one];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    80
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    81
Goal "abs((-1) ^ n) = (1::real)";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    82
by (induct_tac "n" 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
    83
by (auto_tac (claset(), simpset() addsimps [abs_mult]));
8838
4eaa99f0d223 replaced rabs by overloaded abs;
wenzelm
parents: 8442
diff changeset
    84
qed "abs_realpow_minus_one";
4eaa99f0d223 replaced rabs by overloaded abs;
wenzelm
parents: 8442
diff changeset
    85
Addsimps [abs_realpow_minus_one];
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    86
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    87
Goal "((r::real) * s) ^ n = (r ^ n) * (s ^ n)";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    88
by (induct_tac "n" 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    89
by (auto_tac (claset(),simpset() addsimps real_mult_ac));
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    90
qed "realpow_mult";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    91
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    92
Goal "(0::real) <= r^ Suc (Suc 0)";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    93
by (simp_tac (simpset() addsimps [real_le_square]) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    94
qed "realpow_two_le";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    95
Addsimps [realpow_two_le];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
    96
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10919
diff changeset
    97
Goal "abs((x::real)^Suc (Suc 0)) = x^Suc (Suc 0)";
9070
99d93349914b full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
    98
by (simp_tac (simpset() addsimps [abs_eqI1, 
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
    99
				  real_le_square]) 1);
8838
4eaa99f0d223 replaced rabs by overloaded abs;
wenzelm
parents: 8442
diff changeset
   100
qed "abs_realpow_two";
4eaa99f0d223 replaced rabs by overloaded abs;
wenzelm
parents: 8442
diff changeset
   101
Addsimps [abs_realpow_two];
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   102
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10919
diff changeset
   103
Goal "abs(x::real)^Suc (Suc 0) = x^Suc (Suc 0)";
12330
c69bee072501 *** empty log message ***
nipkow
parents: 12196
diff changeset
   104
by (simp_tac (simpset() addsimps [realpow_abs RS sym,abs_eqI1]
10690
cd80241125b0 tidying and adding new proofs
paulson
parents: 10677
diff changeset
   105
                        delsimps [realpow_Suc]) 1);
8838
4eaa99f0d223 replaced rabs by overloaded abs;
wenzelm
parents: 8442
diff changeset
   106
qed "realpow_two_abs";
4eaa99f0d223 replaced rabs by overloaded abs;
wenzelm
parents: 8442
diff changeset
   107
Addsimps [realpow_two_abs];
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   108
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   109
Goal "(1::real) < r ==> 1 < r^ (Suc (Suc 0))";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   110
by Auto_tac;
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   111
by (cut_facts_tac [real_zero_less_one] 1);
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   112
by (forw_inst_tac [("x","0")] order_less_trans 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   113
by (assume_tac 1);
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   114
by (dres_inst_tac [("z","r"),("x","1")] 
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   115
    (real_mult_less_mono1) 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   116
by (auto_tac (claset() addIs [order_less_trans], simpset()));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   117
qed "realpow_two_gt_one";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   118
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   119
Goal "(1::real) < r --> 1 <= r ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   120
by (induct_tac "n" 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   121
by Auto_tac;  
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   122
by (subgoal_tac "1*1 <= r * r^n" 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   123
by (rtac real_mult_le_mono 2); 
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   124
by Auto_tac;  
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   125
qed_spec_mp "realpow_ge_one";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   126
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   127
Goal "(1::real) <= r ==> 1 <= r ^ n";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   128
by (dtac order_le_imp_less_or_eq 1); 
7588
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7292
diff changeset
   129
by (auto_tac (claset() addDs [realpow_ge_one], simpset()));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   130
qed "realpow_ge_one2";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   131
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   132
Goal "(1::real) <= 2 ^ n";
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   133
by (res_inst_tac [("y","1 ^ n")] order_trans 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   134
by (rtac realpow_le 2);
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   135
by (auto_tac (claset() addIs [order_less_imp_le], simpset()));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   136
qed "two_realpow_ge_one";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   137
11704
3c50a2cd6f00 * sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents: 11701
diff changeset
   138
Goal "real (n::nat) < 2 ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   139
by (induct_tac "n" 1);
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   140
by (auto_tac (claset(), simpset() addsimps [real_of_nat_Suc]));
9070
99d93349914b full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
   141
by (stac real_mult_2 1);
99d93349914b full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
   142
by (rtac real_add_less_le_mono 1);
10778
2c6605049646 more tidying, especially to remove real_of_posnat
paulson
parents: 10752
diff changeset
   143
by (auto_tac (claset(), simpset() addsimps [two_realpow_ge_one]));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   144
qed "two_realpow_gt";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   145
Addsimps [two_realpow_gt,two_realpow_ge_one];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   146
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   147
Goal "(-1) ^ (2*n) = (1::real)";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   148
by (induct_tac "n" 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   149
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   150
qed "realpow_minus_one";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   151
Addsimps [realpow_minus_one];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   152
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   153
Goal "(-1) ^ Suc (2*n) = -(1::real)";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   154
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   155
qed "realpow_minus_one_odd";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   156
Addsimps [realpow_minus_one_odd];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   157
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   158
Goal "(-1) ^ Suc (Suc (2*n)) = (1::real)";
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   159
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   160
qed "realpow_minus_one_even";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   161
Addsimps [realpow_minus_one_even];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   162
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   163
Goal "(0::real) < r & r < (1::real) --> r ^ Suc n < r ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   164
by (induct_tac "n" 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   165
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   166
qed_spec_mp "realpow_Suc_less";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   167
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   168
Goal "0 <= r & r < (1::real) --> r ^ Suc n <= r ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   169
by (induct_tac "n" 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   170
by (auto_tac (claset() addIs [order_less_imp_le] 
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   171
                       addSDs [order_le_imp_less_or_eq],
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   172
              simpset()));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   173
qed_spec_mp "realpow_Suc_le";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   174
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   175
Goal "(0::real) <= 0 ^ n";
8442
96023903c2df case_tac now subsumes both boolean and datatype cases;
wenzelm
parents: 8423
diff changeset
   176
by (case_tac "n" 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   177
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   178
qed "realpow_zero_le";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   179
Addsimps [realpow_zero_le];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   180
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   181
Goal "0 < r & r < (1::real) --> r ^ Suc n <= r ^ n";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   182
by (blast_tac (claset() addSIs [order_less_imp_le,
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   183
    realpow_Suc_less]) 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   184
qed_spec_mp "realpow_Suc_le2";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   185
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   186
Goal "[| 0 <= r; r < (1::real) |] ==> r ^ Suc n <= r ^ n";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   187
by (etac (order_le_imp_less_or_eq RS disjE) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   188
by (rtac realpow_Suc_le2 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   189
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   190
qed "realpow_Suc_le3";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   191
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   192
Goal "0 <= r & r < (1::real) & n < N --> r ^ N <= r ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   193
by (induct_tac "N" 1);
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
   194
by (ALLGOALS Asm_simp_tac); 
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
   195
by (Clarify_tac 1);
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   196
by (subgoal_tac "r * r ^ na <= 1 * r ^ n" 1); 
10784
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
   197
by (Asm_full_simp_tac 1); 
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
   198
by (rtac real_mult_le_mono 1); 
27e4d90b35b5 more removal of obsolete rules
paulson
parents: 10778
diff changeset
   199
by (auto_tac (claset(), simpset() addsimps [realpow_ge_zero, less_Suc_eq]));  
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   200
qed_spec_mp "realpow_less_le";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   201
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   202
Goal "[| 0 <= r; r < (1::real); n <= N |] ==> r ^ N <= r ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   203
by (dres_inst_tac [("n","N")] le_imp_less_or_eq 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   204
by (auto_tac (claset() addIs [realpow_less_le],
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   205
    simpset()));
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   206
qed "realpow_le_le";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   207
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   208
Goal "[| 0 < r; r < (1::real) |] ==> r ^ Suc n <= r";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   209
by (dres_inst_tac [("n","1"),("N","Suc n")] 
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   210
    (order_less_imp_le RS realpow_le_le) 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   211
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   212
qed "realpow_Suc_le_self";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   213
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   214
Goal "[| 0 < r; r < (1::real) |] ==> r ^ Suc n < 1";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   215
by (blast_tac (claset() addIs [realpow_Suc_le_self, order_le_less_trans]) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   216
qed "realpow_Suc_less_one";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   217
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   218
Goal "(1::real) <= r --> r ^ n <= r ^ Suc n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   219
by (induct_tac "n" 1);
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   220
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   221
qed_spec_mp "realpow_le_Suc";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   222
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   223
Goal "(1::real) < r --> r ^ n < r ^ Suc n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   224
by (induct_tac "n" 1);
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   225
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   226
qed_spec_mp "realpow_less_Suc";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   227
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   228
Goal "(1::real) < r --> r ^ n <= r ^ Suc n";
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   229
by (blast_tac (claset() addSIs [order_less_imp_le, realpow_less_Suc]) 1);
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   230
qed_spec_mp "realpow_le_Suc2";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   231
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   232
Goal "(1::real) < r & n < N --> r ^ n <= r ^ N";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   233
by (induct_tac "N" 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   234
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   235
by (ALLGOALS(forw_inst_tac [("n","na")] realpow_ge_one));
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   236
by (ALLGOALS(dtac (real_mult_self_le)));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   237
by (assume_tac 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   238
by (assume_tac 2);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   239
by (auto_tac (claset() addIs [order_trans],
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   240
              simpset() addsimps [less_Suc_eq]));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   241
qed_spec_mp "realpow_gt_ge";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   242
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   243
Goal "(1::real) <= r & n < N --> r ^ n <= r ^ N";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   244
by (induct_tac "N" 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   245
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   246
by (ALLGOALS(forw_inst_tac [("n","na")] realpow_ge_one2));
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   247
by (ALLGOALS(dtac (real_mult_self_le2)));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   248
by (assume_tac 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   249
by (assume_tac 2);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   250
by (auto_tac (claset() addIs [order_trans],
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   251
              simpset() addsimps [less_Suc_eq]));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   252
qed_spec_mp "realpow_gt_ge2";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   253
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   254
Goal "[| (1::real) < r; n <= N |] ==> r ^ n <= r ^ N";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   255
by (dres_inst_tac [("n","N")] le_imp_less_or_eq 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   256
by (auto_tac (claset() addIs [realpow_gt_ge], simpset()));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   257
qed "realpow_ge_ge";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   258
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   259
Goal "[| (1::real) <= r; n <= N |] ==> r ^ n <= r ^ N";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   260
by (dres_inst_tac [("n","N")] le_imp_less_or_eq 1);
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   261
by (auto_tac (claset() addIs [realpow_gt_ge2], simpset()));
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   262
qed "realpow_ge_ge2";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   263
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   264
Goal "(1::real) < r ==> r <= r ^ Suc n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   265
by (dres_inst_tac [("n","1"),("N","Suc n")] 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   266
    realpow_ge_ge 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   267
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   268
qed_spec_mp "realpow_Suc_ge_self";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   269
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   270
Goal "(1::real) <= r ==> r <= r ^ Suc n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   271
by (dres_inst_tac [("n","1"),("N","Suc n")] 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   272
    realpow_ge_ge2 1);
10677
36625483213f further round of tidying
paulson
parents: 10648
diff changeset
   273
by Auto_tac;
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   274
qed_spec_mp "realpow_Suc_ge_self2";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   275
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   276
Goal "[| (1::real) < r; 0 < n |] ==> r <= r ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   277
by (dtac (less_not_refl2 RS  not0_implies_Suc) 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   278
by (auto_tac (claset() addSIs 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   279
    [realpow_Suc_ge_self],simpset()));
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   280
qed "realpow_ge_self";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   281
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   282
Goal "[| (1::real) <= r; 0 < n |] ==> r <= r ^ n";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   283
by (dtac (less_not_refl2 RS  not0_implies_Suc) 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   284
by (auto_tac (claset() addSIs [realpow_Suc_ge_self2],simpset()));
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   285
qed "realpow_ge_self2";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   286
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   287
Goal "0 < n --> (x::real) ^ (n - 1) * x = x ^ n";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   288
by (induct_tac "n" 1);
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   289
by (auto_tac (claset(),simpset() 
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   290
    addsimps [real_mult_commute]));
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   291
qed_spec_mp "realpow_minus_mult";
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   292
Addsimps [realpow_minus_mult];
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   293
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   294
Goal "r ~= 0 ==> r * inverse r ^Suc (Suc 0) = inverse (r::real)";
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   295
by (asm_simp_tac (simpset() addsimps [realpow_two,
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   296
                  real_mult_assoc RS sym]) 1);
10606
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   297
qed "realpow_two_mult_inverse";
e3229a37d53f converted rinv to inverse;
bauerg
parents: 10043
diff changeset
   298
Addsimps [realpow_two_mult_inverse];
7077
60b098bb8b8a heavily revised by Jacques: coercions have alphabetic names;
paulson
parents:
diff changeset
   299
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   300
(* 05/00 *)
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10919
diff changeset
   301
Goal "(-x)^Suc (Suc 0) = (x::real)^Suc (Suc 0)";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   302
by (Simp_tac 1);
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   303
qed "realpow_two_minus";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   304
Addsimps [realpow_two_minus];
7588
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7292
diff changeset
   305
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10919
diff changeset
   306
Goalw [real_diff_def] "(x::real)^Suc (Suc 0) - y^Suc (Suc 0) = (x - y) * (x + y)";
10712
351ba950d4d9 further tidying
paulson
parents: 10699
diff changeset
   307
by (simp_tac (simpset() addsimps 
12483
0a01efff43e9 new rewrite rules for use by arith_tac to take care of uminus.
nipkow
parents: 12330
diff changeset
   308
            [real_add_mult_distrib2, real_add_mult_distrib] @ real_mult_ac) 1);
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   309
qed "realpow_two_diff";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   310
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 10919
diff changeset
   311
Goalw [real_diff_def] "((x::real)^Suc (Suc 0) = y^Suc (Suc 0)) = (x = y | x = -y)";
10712
351ba950d4d9 further tidying
paulson
parents: 10699
diff changeset
   312
by (cut_inst_tac [("x","x"),("y","y")] realpow_two_diff 1);
351ba950d4d9 further tidying
paulson
parents: 10699
diff changeset
   313
by (auto_tac (claset(), simpset() delsimps [realpow_Suc]));
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   314
qed "realpow_two_disj";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   315
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   316
(* used in Transc *)
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   317
Goal  "[|(x::real) ~= 0; m <= n |] ==> x ^ (n - m) = x ^ n * inverse (x ^ m)";
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   318
by (auto_tac (claset(),
10712
351ba950d4d9 further tidying
paulson
parents: 10699
diff changeset
   319
       simpset() addsimps [le_eq_less_or_eq, less_iff_Suc_add, realpow_add,
351ba950d4d9 further tidying
paulson
parents: 10699
diff changeset
   320
                           realpow_not_zero] @ real_mult_ac));
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   321
qed "realpow_diff";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   322
10919
144ede948e58 renamings: real_of_nat, real_of_int -> (overloaded) real
paulson
parents: 10784
diff changeset
   323
Goal "real (m::nat) ^ n = real (m ^ n)";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   324
by (induct_tac "n" 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   325
by (auto_tac (claset(),
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   326
              simpset() addsimps [real_of_nat_one, real_of_nat_mult]));
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   327
qed "realpow_real_of_nat";
7588
26384af93359 Tidying to exploit the new arith_tac. RealBin no longer imports RealPow or
paulson
parents: 7292
diff changeset
   328
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   329
Goal "0 < real (Suc (Suc 0) ^ n)";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   330
by (induct_tac "n" 1);
9070
99d93349914b full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
   331
by (auto_tac (claset(),
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   332
          simpset() addsimps [real_of_nat_mult, real_0_less_mult_iff]));
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   333
qed "realpow_real_of_nat_two_pos";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   334
Addsimps [realpow_real_of_nat_two_pos];
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   335
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   336
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   337
Goal "(0::real) <= x --> 0 <= y --> x ^ Suc n <= y ^ Suc n --> x <= y";
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   338
by (induct_tac "n" 1);
9043
ca761fe227d8 First round of changes, towards installation of simprocs
paulson
parents: 9013
diff changeset
   339
by Auto_tac;
10752
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   340
by (asm_full_simp_tac (simpset() addsimps [linorder_not_less RS sym]) 1); 
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   341
by (swap_res_tac [real_mult_less_mono'] 1);
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   342
by Auto_tac;
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   343
by (auto_tac (claset(), simpset() addsimps [real_0_le_mult_iff]));
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   344
by (auto_tac (claset(), simpset() addsimps [linorder_not_less RS sym]));    
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   345
by (dres_inst_tac [("n","n")] realpow_gt_zero 1);   
c4f1bf2acf4c tidying, and separation of HOL-Hyperreal from HOL-Real
paulson
parents: 10715
diff changeset
   346
by Auto_tac;  
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   347
qed_spec_mp "realpow_increasing";
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   348
  
12018
ec054019c910 Numerals and simprocs for types real and hypreal. The abstract
paulson
parents: 11704
diff changeset
   349
Goal "[| (0::real) <= x; 0 <= y; x ^ Suc n = y ^ Suc n |] ==> x = y";
9070
99d93349914b full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
   350
by (blast_tac (claset() addIs [realpow_increasing, order_antisym, 
99d93349914b full_rename_numerals -> rename_numerals; tidied
paulson
parents: 9043
diff changeset
   351
			       order_eq_refl, sym]) 1);
9013
9dd0274f76af Updated files to remove 0r and 1r from theorems in descendant theories
fleuriot
parents: 8838
diff changeset
   352
qed_spec_mp "realpow_Suc_cancel_eq";
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   353
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   354
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   355
(*** Logical equivalences for inequalities ***)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   356
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   357
Goal "(x^n = 0) = (x = (0::real) & 0<n)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   358
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   359
by Auto_tac; 
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   360
qed "realpow_eq_0_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   361
Addsimps [realpow_eq_0_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   362
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   363
Goal "(0 < (abs x)^n) = (x ~= (0::real) | n=0)";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   364
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   365
by (auto_tac (claset(), simpset() addsimps [real_0_less_mult_iff]));  
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   366
qed "zero_less_realpow_abs_iff";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   367
Addsimps [zero_less_realpow_abs_iff];
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   368
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   369
Goal "(0::real) <= (abs x)^n";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   370
by (induct_tac "n" 1);
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   371
by (auto_tac (claset(), simpset() addsimps [real_0_le_mult_iff]));  
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   372
qed "zero_le_realpow_abs";
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents: 12018
diff changeset
   373
Addsimps [zero_le_realpow_abs];
12613
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   374
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   375
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   376
(*** Literal arithmetic involving powers, type real ***)
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   377
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   378
Goal "real (x::int) ^ n = real (x ^ n)";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   379
by (induct_tac "n" 1); 
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   380
by (ALLGOALS (asm_simp_tac (simpset() addsimps [nat_mult_distrib])));
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   381
qed "real_of_int_power";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   382
Addsimps [real_of_int_power RS sym];
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   383
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   384
Goal "(number_of v :: real) ^ n = real ((number_of v :: int) ^ n)";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   385
by (simp_tac (HOL_ss addsimps [real_number_of_def, real_of_int_power]) 1);
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   386
qed "power_real_number_of";
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   387
279facb4253a Literal arithmetic: raising numbers to powers (nat, int, real, hypreal)
paulson
parents: 12483
diff changeset
   388
Addsimps [inst "n" "number_of ?w" power_real_number_of];