12224
|
1 |
(* Title : Log.ML
|
|
2 |
Author : Jacques D. Fleuriot
|
|
3 |
Copyright : 2000,2001 University of Edinburgh
|
|
4 |
Description : standard logarithms only
|
|
5 |
*)
|
|
6 |
|
|
7 |
|
|
8 |
Goalw [powr_def] "1 powr a = 1";
|
|
9 |
by (Simp_tac 1);
|
|
10 |
qed "powr_one_eq_one";
|
|
11 |
Addsimps [powr_one_eq_one];
|
|
12 |
|
|
13 |
Goalw [powr_def] "x powr 0 = 1";
|
|
14 |
by (Simp_tac 1);
|
|
15 |
qed "powr_zero_eq_one";
|
|
16 |
Addsimps [powr_zero_eq_one];
|
|
17 |
|
|
18 |
Goalw [powr_def]
|
|
19 |
"(x powr 1 = x) = (0 < x)";
|
|
20 |
by (Simp_tac 1);
|
|
21 |
qed "powr_one_gt_zero_iff";
|
|
22 |
Addsimps [powr_one_gt_zero_iff];
|
|
23 |
Addsimps [powr_one_gt_zero_iff RS iffD2];
|
|
24 |
|
|
25 |
Goalw [powr_def]
|
|
26 |
"[| 0 < x; 0 < y |] ==> (x * y) powr a = (x powr a) * (y powr a)";
|
|
27 |
by (asm_simp_tac (simpset() addsimps [exp_add RS sym,ln_mult,
|
|
28 |
real_add_mult_distrib2]) 1);
|
|
29 |
qed "powr_mult";
|
|
30 |
|
|
31 |
Goalw [powr_def] "0 < x powr a";
|
|
32 |
by (Simp_tac 1);
|
|
33 |
qed "powr_gt_zero";
|
|
34 |
Addsimps [powr_gt_zero];
|
|
35 |
|
|
36 |
Goalw [powr_def] "x powr a ~= 0";
|
|
37 |
by (Simp_tac 1);
|
|
38 |
qed "powr_not_zero";
|
|
39 |
Addsimps [powr_not_zero];
|
|
40 |
|
|
41 |
Goal "[| 0 < x; 0 < y |] ==> (x / y) powr a = (x powr a)/(y powr a)";
|
|
42 |
by (asm_simp_tac (simpset() addsimps [real_divide_def,real_inverse_gt_0,
|
|
43 |
powr_mult]) 1);
|
|
44 |
by (asm_simp_tac (simpset() addsimps [powr_def,exp_minus RS sym,
|
|
45 |
exp_add RS sym,ln_inverse]) 1);
|
|
46 |
qed "powr_divide";
|
|
47 |
|
|
48 |
Goalw [powr_def] "x powr (a + b) = (x powr a) * (x powr b)";
|
|
49 |
by (asm_simp_tac (simpset() addsimps [exp_add RS sym,
|
|
50 |
real_add_mult_distrib]) 1);
|
|
51 |
qed "powr_add";
|
|
52 |
|
|
53 |
Goalw [powr_def] "(x powr a) powr b = x powr (a * b)";
|
|
54 |
by (simp_tac (simpset() addsimps real_mult_ac) 1);
|
|
55 |
qed "powr_powr";
|
|
56 |
|
|
57 |
Goal "(x powr a) powr b = (x powr b) powr a";
|
|
58 |
by (simp_tac (simpset() addsimps [powr_powr,real_mult_commute]) 1);
|
|
59 |
qed "powr_powr_swap";
|
|
60 |
|
|
61 |
Goalw [powr_def] "x powr (-a) = inverse (x powr a)";
|
|
62 |
by (simp_tac (simpset() addsimps [exp_minus RS sym]) 1);
|
|
63 |
qed "powr_minus";
|
|
64 |
|
|
65 |
Goalw [real_divide_def] "x powr (-a) = 1/(x powr a)";
|
|
66 |
by (simp_tac (simpset() addsimps [powr_minus]) 1);
|
|
67 |
qed "powr_minus_divide";
|
|
68 |
|
|
69 |
Goalw [powr_def]
|
|
70 |
"[| a < b; 1 < x |] ==> x powr a < x powr b";
|
|
71 |
by Auto_tac;
|
|
72 |
qed "powr_less_mono";
|
|
73 |
|
|
74 |
Goalw [powr_def]
|
|
75 |
"[| x powr a < x powr b; 1 < x |] ==> a < b";
|
|
76 |
by (auto_tac (claset() addDs [ln_gt_zero],
|
|
77 |
simpset() addsimps [rename_numerals real_mult_less_cancel2]));
|
|
78 |
qed "powr_less_cancel";
|
|
79 |
|
|
80 |
Goal "1 < x ==> (x powr a < x powr b) = (a < b)";
|
|
81 |
by (blast_tac (claset() addIs [powr_less_cancel,powr_less_mono]) 1);
|
|
82 |
qed "powr_less_cancel_iff";
|
|
83 |
Addsimps [powr_less_cancel_iff];
|
|
84 |
|
|
85 |
Goalw [real_le_def] "1 < x ==> (x powr a <= x powr b) = (a <= b)";
|
|
86 |
by (Auto_tac);
|
|
87 |
qed "powr_le_cancel_iff";
|
|
88 |
Addsimps [powr_le_cancel_iff];
|
|
89 |
|
|
90 |
Goalw [log_def] "ln x = log (exp(1)) x";
|
|
91 |
by (Simp_tac 1);
|
|
92 |
qed "log_ln";
|
|
93 |
|
|
94 |
Goalw [powr_def,log_def]
|
|
95 |
"[| 0 < a; a ~= 1; 0 < x |] ==> a powr (log a x) = x";
|
|
96 |
by Auto_tac;
|
|
97 |
qed "powr_log_cancel";
|
|
98 |
Addsimps [powr_log_cancel];
|
|
99 |
|
|
100 |
Goalw [log_def,powr_def]
|
|
101 |
"[| 0 < a; a ~= 1 |] ==> log a (a powr y) = y";
|
|
102 |
by (auto_tac (claset(),simpset() addsimps [real_divide_def,
|
|
103 |
real_mult_assoc]));
|
|
104 |
qed "log_powr_cancel";
|
|
105 |
Addsimps [log_powr_cancel];
|
|
106 |
|
|
107 |
Goalw [log_def]
|
|
108 |
"[| 0 < a; a ~= 1; 0 < x; 0 < y |] \
|
|
109 |
\ ==> log a (x * y) = log a x + log a y";
|
|
110 |
by (auto_tac (claset(),simpset() addsimps [ln_mult,real_divide_def,
|
|
111 |
real_add_mult_distrib]));
|
|
112 |
qed "log_mult";
|
|
113 |
|
|
114 |
Goalw [log_def,real_divide_def]
|
|
115 |
"[| 0 < a; a ~= 1; 0 < b; b ~= 1; 0 < x |] \
|
|
116 |
\ ==> log a x = (ln b/ln a) * log b x";
|
|
117 |
by Auto_tac;
|
|
118 |
qed "log_eq_div_ln_mult_log";
|
|
119 |
|
|
120 |
(* specific case *)
|
|
121 |
Goal "0 < x ==> log 10 x = (ln (exp 1) / ln 10) * ln x";
|
|
122 |
by (auto_tac (claset(),simpset() addsimps [log_def]));
|
|
123 |
qed "log_base_10_eq1";
|
|
124 |
|
|
125 |
Goal "0 < x ==> log 10 x = (log 10 (exp 1)) * ln x";
|
|
126 |
by (auto_tac (claset(),simpset() addsimps [log_def]));
|
|
127 |
qed "log_base_10_eq2";
|
|
128 |
|
|
129 |
Goalw [log_def] "log a 1 = 0";
|
|
130 |
by Auto_tac;
|
|
131 |
qed "log_one";
|
|
132 |
Addsimps [log_one];
|
|
133 |
|
|
134 |
Goalw [log_def]
|
|
135 |
"[| 0 < a; a ~= 1 |] ==> log a a = 1";
|
|
136 |
by Auto_tac;
|
|
137 |
qed "log_eq_one";
|
|
138 |
Addsimps [log_eq_one];
|
|
139 |
|
|
140 |
Goal "[| 0 < a; a ~= 1; 0 < x |] ==> log a (inverse x) = - log a x";
|
|
141 |
by (res_inst_tac [("x1","log a x")] (real_add_left_cancel RS iffD1) 1);
|
|
142 |
by (auto_tac (claset(),simpset() addsimps [log_mult RS sym]));
|
|
143 |
qed "log_inverse";
|
|
144 |
|
|
145 |
Goal "[|0 < a; a ~= 1; 0 < x; 0 < y|] \
|
|
146 |
\ ==> log a (x/y) = log a x - log a y";
|
|
147 |
by (auto_tac (claset(),simpset() addsimps [log_mult,real_divide_def,
|
|
148 |
log_inverse]));
|
|
149 |
qed "log_divide";
|
|
150 |
|
|
151 |
Goal "[| 1 < a; 0 < x; 0 < y |] ==> (log a x < log a y) = (x < y)";
|
|
152 |
by (Step_tac 1);
|
|
153 |
by (rtac powr_less_cancel 2);
|
|
154 |
by (dres_inst_tac [("a","log a x")] powr_less_mono 1);
|
|
155 |
by Auto_tac;
|
|
156 |
qed "log_less_cancel_iff";
|
|
157 |
Addsimps [log_less_cancel_iff];
|
|
158 |
|
|
159 |
Goal "[| 1 < a; 0 < x; 0 < y |] ==> (log a x <= log a y) = (x <= y)";
|
|
160 |
by (auto_tac (claset(),simpset() addsimps [real_le_def]));
|
|
161 |
qed "log_le_cancel_iff";
|
|
162 |
Addsimps [log_le_cancel_iff];
|
|
163 |
|