| author | paulson | 
| Mon, 18 Dec 2000 12:21:54 +0100 | |
| changeset 10689 | 5c44de6aadf4 | 
| parent 10677 | 36625483213f | 
| child 10690 | cd80241125b0 | 
| permissions | -rw-r--r-- | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 1 | (* Title : RealPow.ML | 
| 7219 | 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 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 6 | *) | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 7 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 8 | Goal "(#0::real) ^ (Suc n) = #0"; | 
| 10677 | 9 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 10 | qed "realpow_zero"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 11 | Addsimps [realpow_zero]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 12 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 13 | Goal "r ~= (#0::real) --> r ^ n ~= #0"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 14 | by (induct_tac "n" 1); | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 15 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 16 | qed_spec_mp "realpow_not_zero"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 17 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 18 | Goal "r ^ n = (#0::real) ==> r = #0"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 19 | by (rtac ccontr 1); | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 20 | by (auto_tac (claset() addDs [realpow_not_zero], simpset())); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 21 | qed "realpow_zero_zero"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 22 | |
| 10648 | 23 | Goal "inverse ((r::real) ^ n) = (inverse r) ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 24 | by (induct_tac "n" 1); | 
| 10648 | 25 | by (auto_tac (claset(), simpset() addsimps [real_inverse_distrib])); | 
| 26 | qed "realpow_inverse"; | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 27 | |
| 8838 | 28 | Goal "abs (r::real) ^ n = abs (r ^ n)"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 29 | by (induct_tac "n" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 30 | by (auto_tac (claset(),simpset() addsimps | 
| 8838 | 31 | [abs_mult,abs_one])); | 
| 32 | qed "realpow_abs"; | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 33 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 34 | Goal "(r::real) ^ (n + m) = (r ^ n) * (r ^ m)"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 35 | by (induct_tac "n" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 36 | by (auto_tac (claset(),simpset() addsimps real_mult_ac)); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 37 | qed "realpow_add"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 38 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 39 | Goal "(r::real) ^ 1 = r"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 40 | by (Simp_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 41 | qed "realpow_one"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 42 | Addsimps [realpow_one]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 43 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 44 | Goal "(r::real) ^ 2 = r * r"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 45 | by (Simp_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 46 | qed "realpow_two"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 47 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 48 | Goal "(#0::real) < r --> #0 <= r ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 49 | by (induct_tac "n" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 50 | by (auto_tac (claset() addDs [real_less_imp_le] | 
| 9428 | 51 | addIs [rename_numerals real_le_mult_order], | 
| 9070 | 52 | simpset())); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 53 | qed_spec_mp "realpow_ge_zero"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 54 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 55 | Goal "(#0::real) < r --> #0 < r ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 56 | by (induct_tac "n" 1); | 
| 9428 | 57 | by (auto_tac (claset() addIs [rename_numerals real_mult_order], | 
| 9070 | 58 | simpset() addsimps [real_zero_less_one])); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 59 | qed_spec_mp "realpow_gt_zero"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 60 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 61 | Goal "(#0::real) <= r --> #0 <= r ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 62 | by (induct_tac "n" 1); | 
| 9428 | 63 | by (auto_tac (claset() addIs [rename_numerals real_le_mult_order], | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 64 | simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 65 | qed_spec_mp "realpow_ge_zero2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 66 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 67 | Goal "(#0::real) < x & x <= y --> x ^ n <= y ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 68 | by (induct_tac "n" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 69 | by (auto_tac (claset() addSIs [real_mult_le_mono], | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 70 | simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 71 | by (asm_simp_tac (simpset() addsimps [realpow_ge_zero]) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 72 | qed_spec_mp "realpow_le"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 73 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 74 | Goal "(#0::real) <= x & x <= y --> x ^ n <= y ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 75 | by (induct_tac "n" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 76 | by (auto_tac (claset() addSIs [real_mult_le_mono4], | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 77 | simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 78 | by (asm_simp_tac (simpset() addsimps [realpow_ge_zero2]) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 79 | qed_spec_mp "realpow_le2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 80 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 81 | 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 | 82 | by (induct_tac "n" 1); | 
| 9428 | 83 | by (auto_tac (claset() addIs [rename_numerals real_mult_less_mono, gr0I] | 
| 9070 | 84 | addDs [realpow_gt_zero], | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 85 | simpset())); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 86 | qed_spec_mp "realpow_less"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 87 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 88 | Goal "#1 ^ n = (#1::real)"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 89 | by (induct_tac "n" 1); | 
| 10677 | 90 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 91 | qed "realpow_eq_one"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 92 | Addsimps [realpow_eq_one]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 93 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 94 | Goal "abs(-(#1 ^ n)) = (#1::real)"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 95 | by (simp_tac (simpset() addsimps | 
| 8838 | 96 | [abs_minus_cancel,abs_one]) 1); | 
| 97 | qed "abs_minus_realpow_one"; | |
| 98 | Addsimps [abs_minus_realpow_one]; | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 99 | |
| 10677 | 100 | Goal "abs((#-1) ^ n) = (#1::real)"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 101 | by (induct_tac "n" 1); | 
| 8838 | 102 | by (auto_tac (claset(),simpset() addsimps [abs_mult, | 
| 103 | abs_minus_cancel,abs_one])); | |
| 104 | qed "abs_realpow_minus_one"; | |
| 105 | Addsimps [abs_realpow_minus_one]; | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 106 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 107 | Goal "((r::real) * s) ^ n = (r ^ n) * (s ^ n)"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 108 | by (induct_tac "n" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 109 | by (auto_tac (claset(),simpset() addsimps real_mult_ac)); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 110 | qed "realpow_mult"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 111 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 112 | Goal "(#0::real) <= r ^ 2"; | 
| 9428 | 113 | by (simp_tac (simpset() addsimps [rename_numerals real_le_square]) 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 114 | qed "realpow_two_le"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 115 | Addsimps [realpow_two_le]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 116 | |
| 8838 | 117 | Goal "abs((x::real) ^ 2) = x ^ 2"; | 
| 9070 | 118 | by (simp_tac (simpset() addsimps [abs_eqI1, | 
| 9428 | 119 | rename_numerals real_le_square]) 1); | 
| 8838 | 120 | qed "abs_realpow_two"; | 
| 121 | Addsimps [abs_realpow_two]; | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 122 | |
| 8838 | 123 | Goal "abs(x::real) ^ 2 = x ^ 2"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 124 | by (simp_tac (simpset() addsimps | 
| 9428 | 125 | [realpow_abs,abs_eqI1] delsimps [thm "realpow_Suc"]) 1); | 
| 8838 | 126 | qed "realpow_two_abs"; | 
| 127 | Addsimps [realpow_two_abs]; | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 128 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 129 | Goal "(#1::real) < r ==> #1 < r ^ 2"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 130 | by Auto_tac; | 
| 9428 | 131 | by (cut_facts_tac [rename_numerals real_zero_less_one] 1); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 132 | by (forw_inst_tac [("R1.0","#0")] real_less_trans 1);
 | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 133 | by (assume_tac 1); | 
| 9070 | 134 | by (dres_inst_tac [("z","r"),("x","#1")] 
 | 
| 9428 | 135 | (rename_numerals real_mult_less_mono1) 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 136 | by (auto_tac (claset() addIs [real_less_trans],simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 137 | qed "realpow_two_gt_one"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 138 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 139 | Goal "(#1::real) < r --> #1 <= r ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 140 | by (induct_tac "n" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 141 | by (auto_tac (claset() addSDs [real_le_imp_less_or_eq], | 
| 7588 
26384af93359
Tidying to exploit the new arith_tac.  RealBin no longer imports RealPow or
 paulson parents: 
7292diff
changeset | 142 | simpset())); | 
| 9428 | 143 | by (dtac (rename_numerals (real_zero_less_one RS real_mult_less_mono)) 1); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 144 | by (dtac sym 4); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 145 | by (auto_tac (claset() addSIs [real_less_imp_le], | 
| 7588 
26384af93359
Tidying to exploit the new arith_tac.  RealBin no longer imports RealPow or
 paulson parents: 
7292diff
changeset | 146 | simpset() addsimps [real_zero_less_one])); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 147 | qed_spec_mp "realpow_ge_one"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 148 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 149 | Goal "(#1::real) < r ==> #1 < r ^ (Suc n)"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 150 | by (forw_inst_tac [("n","n")] realpow_ge_one 1);
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 151 | by (dtac real_le_imp_less_or_eq 1 THEN Step_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 152 | by (dtac sym 2); | 
| 9428 | 153 | by (ftac (rename_numerals (real_zero_less_one RS real_less_trans)) 1); | 
| 9070 | 154 | by (dres_inst_tac [("y","r ^ n")] 
 | 
| 9428 | 155 | (rename_numerals real_mult_less_mono2) 1); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 156 | by (assume_tac 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 157 | by (auto_tac (claset() addDs [real_less_trans], | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 158 | simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 159 | qed "realpow_Suc_gt_one"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 160 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 161 | Goal "(#1::real) <= r ==> #1 <= r ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 162 | by (dtac real_le_imp_less_or_eq 1); | 
| 7588 
26384af93359
Tidying to exploit the new arith_tac.  RealBin no longer imports RealPow or
 paulson parents: 
7292diff
changeset | 163 | by (auto_tac (claset() addDs [realpow_ge_one], simpset())); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 164 | qed "realpow_ge_one2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 165 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 166 | Goal "(#0::real) < r ==> #0 < r ^ Suc n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 167 | by (forw_inst_tac [("n","n")] realpow_ge_zero 1);
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 168 | by (forw_inst_tac [("n1","n")]
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 169 | ((real_not_refl2 RS not_sym) RS realpow_not_zero RS not_sym) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 170 | by (auto_tac (claset() addSDs [real_le_imp_less_or_eq] | 
| 9428 | 171 | addIs [rename_numerals real_mult_order], | 
| 9070 | 172 | simpset())); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 173 | qed "realpow_Suc_gt_zero"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 174 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 175 | Goal "(#0::real) <= r ==> #0 <= r ^ Suc n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 176 | by (etac (real_le_imp_less_or_eq RS disjE) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 177 | by (etac (realpow_ge_zero) 1); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 178 | by (dtac sym 1); | 
| 7588 
26384af93359
Tidying to exploit the new arith_tac.  RealBin no longer imports RealPow or
 paulson parents: 
7292diff
changeset | 179 | by (Asm_simp_tac 1); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 180 | qed "realpow_Suc_ge_zero"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 181 | |
| 7588 
26384af93359
Tidying to exploit the new arith_tac.  RealBin no longer imports RealPow or
 paulson parents: 
7292diff
changeset | 182 | Goal "(#1::real) <= #2 ^ n"; | 
| 
26384af93359
Tidying to exploit the new arith_tac.  RealBin no longer imports RealPow or
 paulson parents: 
7292diff
changeset | 183 | by (res_inst_tac [("j","#1 ^ n")] real_le_trans 1);
 | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 184 | by (rtac realpow_le 2); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 185 | by (auto_tac (claset() addIs [real_less_imp_le],simpset())); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 186 | qed "two_realpow_ge_one"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 187 | |
| 7588 
26384af93359
Tidying to exploit the new arith_tac.  RealBin no longer imports RealPow or
 paulson parents: 
7292diff
changeset | 188 | Goal "real_of_nat n < #2 ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 189 | by (induct_tac "n" 1); | 
| 9070 | 190 | by Auto_tac; | 
| 191 | by (stac real_mult_2 1); | |
| 192 | by (rtac real_add_less_le_mono 1); | |
| 7292 | 193 | by (auto_tac (claset(), | 
| 9070 | 194 | simpset() addsimps [two_realpow_ge_one])); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 195 | qed "two_realpow_gt"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 196 | Addsimps [two_realpow_gt,two_realpow_ge_one]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 197 | |
| 10677 | 198 | Goal "(#-1) ^ (#2*n) = (#1::real)"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 199 | by (induct_tac "n" 1); | 
| 10677 | 200 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 201 | qed "realpow_minus_one"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 202 | Addsimps [realpow_minus_one]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 203 | |
| 10677 | 204 | Goal "(#-1) ^ Suc (#2*n) = -(#1::real)"; | 
| 205 | by Auto_tac; | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 206 | qed "realpow_minus_one_odd"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 207 | Addsimps [realpow_minus_one_odd]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 208 | |
| 10677 | 209 | Goal "(#-1) ^ Suc (Suc (#2*n)) = (#1::real)"; | 
| 210 | by Auto_tac; | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 211 | qed "realpow_minus_one_even"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 212 | Addsimps [realpow_minus_one_even]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 213 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 214 | 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 | 215 | by (induct_tac "n" 1); | 
| 10677 | 216 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 217 | qed_spec_mp "realpow_Suc_less"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 218 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 219 | 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 | 220 | by (induct_tac "n" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 221 | by (auto_tac (claset() addIs [real_less_imp_le] addSDs | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 222 | [real_le_imp_less_or_eq],simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 223 | qed_spec_mp "realpow_Suc_le"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 224 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 225 | Goal "(#0::real) <= #0 ^ n"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 226 | by (case_tac "n" 1); | 
| 10677 | 227 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 228 | qed "realpow_zero_le"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 229 | Addsimps [realpow_zero_le]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 230 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 231 | 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 | 232 | by (blast_tac (claset() addSIs [real_less_imp_le, | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 233 | realpow_Suc_less]) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 234 | qed_spec_mp "realpow_Suc_le2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 235 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 236 | 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 | 237 | by (etac (real_le_imp_less_or_eq RS disjE) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 238 | by (rtac realpow_Suc_le2 1); | 
| 10677 | 239 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 240 | qed "realpow_Suc_le3"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 241 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 242 | 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 | 243 | by (induct_tac "N" 1); | 
| 10677 | 244 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 245 | by (ALLGOALS(forw_inst_tac [("n","na")] realpow_ge_zero2));
 | 
| 9428 | 246 | by (ALLGOALS(dtac (rename_numerals real_mult_le_mono3))); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 247 | by (REPEAT(assume_tac 1)); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 248 | by (REPEAT(assume_tac 3)); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 249 | by (auto_tac (claset(),simpset() addsimps | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 250 | [less_Suc_eq])); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 251 | qed_spec_mp "realpow_less_le"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 252 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 253 | 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 | 254 | 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 | 255 | by (auto_tac (claset() addIs [realpow_less_le], | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 256 | simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 257 | qed "realpow_le_le"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 258 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 259 | Goal "[| #0 < r; r < (#1::real) |] ==> r ^ Suc n <= r"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 260 | by (dres_inst_tac [("n","1"),("N","Suc n")] 
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 261 | (real_less_imp_le RS realpow_le_le) 1); | 
| 10677 | 262 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 263 | qed "realpow_Suc_le_self"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 264 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 265 | Goal "[| #0 < r; r < (#1::real) |] ==> r ^ Suc n < #1"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 266 | by (blast_tac (claset() addIs [realpow_Suc_le_self, | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 267 | real_le_less_trans]) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 268 | qed "realpow_Suc_less_one"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 269 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 270 | Goal "(#1::real) <= r --> r ^ n <= r ^ Suc n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 271 | by (induct_tac "n" 1); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 272 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 273 | qed_spec_mp "realpow_le_Suc"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 274 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 275 | Goal "(#1::real) < r --> r ^ n < r ^ Suc n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 276 | by (induct_tac "n" 1); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 277 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 278 | qed_spec_mp "realpow_less_Suc"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 279 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 280 | Goal "(#1::real) < r --> r ^ n <= r ^ Suc n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 281 | by (blast_tac (claset() addSIs [real_less_imp_le, | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 282 | realpow_less_Suc]) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 283 | qed_spec_mp "realpow_le_Suc2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 284 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 285 | Goal "(#1::real) < r & n < N --> r ^ n <= r ^ N"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 286 | by (induct_tac "N" 1); | 
| 10677 | 287 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 288 | by (ALLGOALS(forw_inst_tac [("n","na")] realpow_ge_one));
 | 
| 9428 | 289 | by (ALLGOALS(dtac (rename_numerals real_mult_self_le))); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 290 | by (assume_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 291 | by (assume_tac 2); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 292 | by (auto_tac (claset() addIs [real_le_trans], | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 293 | simpset() addsimps [less_Suc_eq])); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 294 | qed_spec_mp "realpow_gt_ge"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 295 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 296 | Goal "(#1::real) <= r & n < N --> r ^ n <= r ^ N"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 297 | by (induct_tac "N" 1); | 
| 10677 | 298 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 299 | by (ALLGOALS(forw_inst_tac [("n","na")] realpow_ge_one2));
 | 
| 9428 | 300 | by (ALLGOALS(dtac (rename_numerals real_mult_self_le2))); | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 301 | by (assume_tac 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 302 | by (assume_tac 2); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 303 | by (auto_tac (claset() addIs [real_le_trans], | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 304 | simpset() addsimps [less_Suc_eq])); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 305 | qed_spec_mp "realpow_gt_ge2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 306 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 307 | Goal "[| (#1::real) < r; n <= N |] ==> r ^ n <= r ^ N"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 308 | 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 | 309 | by (auto_tac (claset() addIs [realpow_gt_ge],simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 310 | qed "realpow_ge_ge"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 311 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 312 | Goal "[| (#1::real) <= r; n <= N |] ==> r ^ n <= r ^ N"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 313 | 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 | 314 | by (auto_tac (claset() addIs [realpow_gt_ge2],simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 315 | qed "realpow_ge_ge2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 316 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 317 | Goal "(#1::real) < r ==> r <= r ^ Suc n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 318 | by (dres_inst_tac [("n","1"),("N","Suc n")] 
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 319 | realpow_ge_ge 1); | 
| 10677 | 320 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 321 | qed_spec_mp "realpow_Suc_ge_self"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 322 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 323 | Goal "(#1::real) <= r ==> r <= r ^ Suc n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 324 | by (dres_inst_tac [("n","1"),("N","Suc n")] 
 | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 325 | realpow_ge_ge2 1); | 
| 10677 | 326 | by Auto_tac; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 327 | qed_spec_mp "realpow_Suc_ge_self2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 328 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 329 | Goal "[| (#1::real) < r; 0 < n |] ==> r <= r ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 330 | by (dtac (less_not_refl2 RS not0_implies_Suc) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 331 | by (auto_tac (claset() addSIs | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 332 | [realpow_Suc_ge_self],simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 333 | qed "realpow_ge_self"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 334 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 335 | Goal "[| (#1::real) <= r; 0 < n |] ==> r <= r ^ n"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 336 | by (dtac (less_not_refl2 RS not0_implies_Suc) 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 337 | by (auto_tac (claset() addSIs [realpow_Suc_ge_self2],simpset())); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 338 | qed "realpow_ge_self2"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 339 | |
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 340 | Goal "0 < n --> (x::real) ^ (n - 1) * x = x ^ n"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 341 | by (induct_tac "n" 1); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 342 | by (auto_tac (claset(),simpset() | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 343 | addsimps [real_mult_commute])); | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 344 | qed_spec_mp "realpow_minus_mult"; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 345 | Addsimps [realpow_minus_mult]; | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 346 | |
| 10606 | 347 | Goal "r ~= #0 ==> r * inverse r ^ 2 = inverse (r::real)"; | 
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 348 | by (asm_simp_tac (simpset() addsimps [realpow_two, | 
| 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 349 | real_mult_assoc RS sym]) 1); | 
| 10606 | 350 | qed "realpow_two_mult_inverse"; | 
| 351 | Addsimps [realpow_two_mult_inverse]; | |
| 7077 
60b098bb8b8a
heavily revised by Jacques: coercions have alphabetic names;
 paulson parents: diff
changeset | 352 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 353 | (* 05/00 *) | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 354 | Goal "(-x) ^ 2 = (x::real) ^ 2"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 355 | by (Simp_tac 1); | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 356 | qed "realpow_two_minus"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 357 | Addsimps [realpow_two_minus]; | 
| 7588 
26384af93359
Tidying to exploit the new arith_tac.  RealBin no longer imports RealPow or
 paulson parents: 
7292diff
changeset | 358 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 359 | Goal "(x::real) ^ 2 + - (y ^ 2) = (x + -y) * (x + y)"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 360 | by (simp_tac (simpset() addsimps [real_add_mult_distrib2, | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 361 | real_add_mult_distrib, real_minus_mult_eq2 RS sym] | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 362 | @ real_mult_ac) 1); | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 363 | qed "realpow_two_add_minus"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 364 | |
| 9428 | 365 | Goalw [real_diff_def] | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 366 | "(x::real) ^ 2 - y ^ 2 = (x - y) * (x + y)"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 367 | by (simp_tac (simpset() addsimps [realpow_two_add_minus] | 
| 9428 | 368 | delsimps [thm "realpow_Suc"]) 1); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 369 | qed "realpow_two_diff"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 370 | |
| 9428 | 371 | Goalw [real_diff_def] | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 372 | "((x::real) ^ 2 = y ^ 2) = (x = y | x = -y)"; | 
| 9428 | 373 | by (auto_tac (claset(),simpset() delsimps [thm "realpow_Suc"])); | 
| 374 | by (dtac (rename_numerals (real_eq_minus_iff RS iffD1 RS sym)) 1); | |
| 375 | by (auto_tac (claset() addDs [rename_numerals real_add_minus_eq_minus], | |
| 376 | simpset() addsimps [realpow_two_add_minus] delsimps [thm "realpow_Suc"])); | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 377 | qed "realpow_two_disj"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 378 | |
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 379 | (* used in Transc *) | 
| 10606 | 380 | 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: 
9013diff
changeset | 381 | by (auto_tac (claset(), | 
| 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 382 | simpset() addsimps [le_eq_less_or_eq, | 
| 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 383 | less_iff_Suc_add,realpow_add, | 
| 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 384 | realpow_not_zero] @ real_mult_ac)); | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 385 | qed "realpow_diff"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 386 | |
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 387 | Goal "real_of_nat (m) ^ n = real_of_nat (m ^ n)"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 388 | by (induct_tac "n" 1); | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 389 | by (auto_tac (claset(), | 
| 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 390 | 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: 
8838diff
changeset | 391 | qed "realpow_real_of_nat"; | 
| 7588 
26384af93359
Tidying to exploit the new arith_tac.  RealBin no longer imports RealPow or
 paulson parents: 
7292diff
changeset | 392 | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 393 | Goal "#0 < real_of_nat (2 ^ n)"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 394 | by (induct_tac "n" 1); | 
| 9070 | 395 | by (auto_tac (claset(), | 
| 396 | simpset() addsimps [real_of_nat_mult, real_zero_less_mult_iff])); | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 397 | qed "realpow_real_of_nat_two_pos"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 398 | Addsimps [realpow_real_of_nat_two_pos]; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 399 | |
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 400 | |
| 9070 | 401 | 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: 
8838diff
changeset | 402 | by (induct_tac "n" 1); | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 403 | by Auto_tac; | 
| 9070 | 404 | by (rtac real_leI 1); | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 405 | by (auto_tac (claset(), | 
| 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 406 | simpset() addsimps [inst "x" "#0" order_le_less, | 
| 9070 | 407 | real_mult_le_0_iff])); | 
| 10606 | 408 | by (subgoal_tac "inverse x * (x * (x * x ^ n)) <= inverse y * (y * (y * y ^ n))" 1); | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 409 | by (rtac real_mult_le_mono 2); | 
| 9070 | 410 | by (asm_full_simp_tac (simpset() addsimps [realpow_ge_zero, real_0_le_mult_iff]) 4); | 
| 10606 | 411 | by (asm_full_simp_tac (simpset() addsimps [real_inverse_le_iff]) 3); | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 412 | by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1); | 
| 10606 | 413 | by (rtac real_inverse_gt_zero 1); | 
| 9043 
ca761fe227d8
First round of changes, towards installation of simprocs
 paulson parents: 
9013diff
changeset | 414 | by Auto_tac; | 
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 415 | qed_spec_mp "realpow_increasing"; | 
| 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 416 | |
| 9070 | 417 | Goal "[| (#0::real) <= x; #0 <= y; x ^ Suc n = y ^ Suc n |] ==> x = y"; | 
| 418 | by (blast_tac (claset() addIs [realpow_increasing, order_antisym, | |
| 419 | order_eq_refl, sym]) 1); | |
| 9013 
9dd0274f76af
Updated files to remove 0r and 1r from theorems in descendant theories
 fleuriot parents: 
8838diff
changeset | 420 | qed_spec_mp "realpow_Suc_cancel_eq"; |