| author | paulson | 
| Fri, 20 Aug 2004 12:20:09 +0200 | |
| changeset 15150 | c7af682b9ee5 | 
| parent 15140 | 322485b816ac | 
| child 16819 | 00d8f9300d13 | 
| permissions | -rw-r--r-- | 
| 12224 | 1 | (* Title : Log.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 2000,2001 University of Edinburgh | |
| 4 | *) | |
| 5 | ||
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 6 | header{*Logarithms: Standard Version*}
 | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 7 | |
| 15131 | 8 | theory Log | 
| 15140 | 9 | imports Transcendental | 
| 15131 | 10 | begin | 
| 12224 | 11 | |
| 12 | constdefs | |
| 13 | ||
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 14 | powr :: "[real,real] => real" (infixr "powr" 80) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 15 |     --{*exponentation with real exponent*}
 | 
| 12224 | 16 | "x powr a == exp(a * ln x)" | 
| 17 | ||
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 18 | log :: "[real,real] => real" | 
| 15053 | 19 |     --{*logarithm of @{term x} to base @{term a}*}
 | 
| 12224 | 20 | "log a x == ln x / ln a" | 
| 21 | ||
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 22 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 23 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 24 | lemma powr_one_eq_one [simp]: "1 powr a = 1" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 25 | by (simp add: powr_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 26 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 27 | lemma powr_zero_eq_one [simp]: "x powr 0 = 1" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 28 | by (simp add: powr_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 29 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 30 | lemma powr_one_gt_zero_iff [simp]: "(x powr 1 = x) = (0 < x)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 31 | by (simp add: powr_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 32 | declare powr_one_gt_zero_iff [THEN iffD2, simp] | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 33 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 34 | lemma powr_mult: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 35 | "[| 0 < x; 0 < y |] ==> (x * y) powr a = (x powr a) * (y powr a)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 36 | by (simp add: powr_def exp_add [symmetric] ln_mult right_distrib) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 37 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 38 | lemma powr_gt_zero [simp]: "0 < x powr a" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 39 | by (simp add: powr_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 40 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 41 | lemma powr_not_zero [simp]: "x powr a \<noteq> 0" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 42 | by (simp add: powr_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 43 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 44 | lemma powr_divide: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 45 | "[| 0 < x; 0 < y |] ==> (x / y) powr a = (x powr a)/(y powr a)" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14411diff
changeset | 46 | apply (simp add: divide_inverse positive_imp_inverse_positive powr_mult) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 47 | apply (simp add: powr_def exp_minus [symmetric] exp_add [symmetric] ln_inverse) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 48 | done | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 49 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 50 | lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 51 | by (simp add: powr_def exp_add [symmetric] left_distrib) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 52 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 53 | lemma powr_powr: "(x powr a) powr b = x powr (a * b)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 54 | by (simp add: powr_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 55 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 56 | lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 57 | by (simp add: powr_powr real_mult_commute) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 58 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 59 | lemma powr_minus: "x powr (-a) = inverse (x powr a)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 60 | by (simp add: powr_def exp_minus [symmetric]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 61 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 62 | lemma powr_minus_divide: "x powr (-a) = 1/(x powr a)" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14411diff
changeset | 63 | by (simp add: divide_inverse powr_minus) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 64 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 65 | lemma powr_less_mono: "[| a < b; 1 < x |] ==> x powr a < x powr b" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 66 | by (simp add: powr_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 67 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 68 | lemma powr_less_cancel: "[| x powr a < x powr b; 1 < x |] ==> a < b" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 69 | by (simp add: powr_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 70 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 71 | lemma powr_less_cancel_iff [simp]: "1 < x ==> (x powr a < x powr b) = (a < b)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 72 | by (blast intro: powr_less_cancel powr_less_mono) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 73 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 74 | lemma powr_le_cancel_iff [simp]: "1 < x ==> (x powr a \<le> x powr b) = (a \<le> b)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 75 | by (simp add: linorder_not_less [symmetric]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 76 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 77 | lemma log_ln: "ln x = log (exp(1)) x" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 78 | by (simp add: log_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 79 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 80 | lemma powr_log_cancel [simp]: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 81 | "[| 0 < a; a \<noteq> 1; 0 < x |] ==> a powr (log a x) = x" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 82 | by (simp add: powr_def log_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 83 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 84 | lemma log_powr_cancel [simp]: "[| 0 < a; a \<noteq> 1 |] ==> log a (a powr y) = y" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 85 | by (simp add: log_def powr_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 86 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 87 | lemma log_mult: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 88 | "[| 0 < a; a \<noteq> 1; 0 < x; 0 < y |] | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 89 | ==> log a (x * y) = log a x + log a y" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14411diff
changeset | 90 | by (simp add: log_def ln_mult divide_inverse left_distrib) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 91 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 92 | lemma log_eq_div_ln_mult_log: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 93 | "[| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |] | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 94 | ==> log a x = (ln b/ln a) * log b x" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14411diff
changeset | 95 | by (simp add: log_def divide_inverse) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 96 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 97 | text{*Base 10 logarithms*}
 | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 98 | lemma log_base_10_eq1: "0 < x ==> log 10 x = (ln (exp 1) / ln 10) * ln x" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 99 | by (simp add: log_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 100 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 101 | lemma log_base_10_eq2: "0 < x ==> log 10 x = (log 10 (exp 1)) * ln x" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 102 | by (simp add: log_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 103 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 104 | lemma log_one [simp]: "log a 1 = 0" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 105 | by (simp add: log_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 106 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 107 | lemma log_eq_one [simp]: "[| 0 < a; a \<noteq> 1 |] ==> log a a = 1" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 108 | by (simp add: log_def) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 109 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 110 | lemma log_inverse: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 111 | "[| 0 < a; a \<noteq> 1; 0 < x |] ==> log a (inverse x) = - log a x" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 112 | apply (rule_tac a1 = "log a x" in add_left_cancel [THEN iffD1]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 113 | apply (simp add: log_mult [symmetric]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 114 | done | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 115 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 116 | lemma log_divide: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 117 | "[|0 < a; a \<noteq> 1; 0 < x; 0 < y|] ==> log a (x/y) = log a x - log a y" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14411diff
changeset | 118 | by (simp add: log_mult divide_inverse log_inverse) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 119 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 120 | lemma log_less_cancel_iff [simp]: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 121 | "[| 1 < a; 0 < x; 0 < y |] ==> (log a x < log a y) = (x < y)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 122 | apply safe | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 123 | apply (rule_tac [2] powr_less_cancel) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 124 | apply (drule_tac a = "log a x" in powr_less_mono, auto) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 125 | done | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 126 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 127 | lemma log_le_cancel_iff [simp]: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 128 | "[| 1 < a; 0 < x; 0 < y |] ==> (log a x \<le> log a y) = (x \<le> y)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 129 | by (simp add: linorder_not_less [symmetric]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 130 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 131 | |
| 15085 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 132 | subsection{*Further Results Courtesy Jeremy Avigad*}
 | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 133 | |
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 134 | lemma powr_realpow: "0 < x ==> x powr (real n) = x^n" | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 135 | apply (induct n, simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 136 | apply (subgoal_tac "real(Suc n) = real n + 1") | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 137 | apply (erule ssubst) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 138 | apply (subst powr_add, simp, simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 139 | done | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 140 | |
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 141 | lemma powr_realpow2: "0 <= x ==> 0 < n ==> x^n = (if (x = 0) then 0 | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 142 | else x powr (real n))" | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 143 | apply (case_tac "x = 0", simp, simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 144 | apply (rule powr_realpow [THEN sym], simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 145 | done | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 146 | |
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 147 | lemma ln_pwr: "0 < x ==> 0 < y ==> ln(x powr y) = y * ln x" | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 148 | by (unfold powr_def, simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 149 | |
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 150 | lemma ln_bound: "1 <= x ==> ln x <= x" | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 151 | apply (subgoal_tac "ln(1 + (x - 1)) <= x - 1") | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 152 | apply simp | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 153 | apply (rule ln_add_one_self_le_self, simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 154 | done | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 155 | |
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 156 | lemma powr_mono: "a <= b ==> 1 <= x ==> x powr a <= x powr b" | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 157 | apply (case_tac "x = 1", simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 158 | apply (case_tac "a = b", simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 159 | apply (rule order_less_imp_le) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 160 | apply (rule powr_less_mono, auto) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 161 | done | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 162 | |
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 163 | lemma ge_one_powr_ge_zero: "1 <= x ==> 0 <= a ==> 1 <= x powr a" | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 164 | apply (subst powr_zero_eq_one [THEN sym]) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 165 | apply (rule powr_mono, assumption+) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 166 | done | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 167 | |
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 168 | lemma powr_less_mono2: "0 < a ==> 0 < x ==> x < y ==> x powr a < | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 169 | y powr a" | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 170 | apply (unfold powr_def) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 171 | apply (rule exp_less_mono) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 172 | apply (rule mult_strict_left_mono) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 173 | apply (subst ln_less_cancel_iff, assumption) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 174 | apply (rule order_less_trans) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 175 | prefer 2 | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 176 | apply assumption+ | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 177 | done | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 178 | |
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 179 | lemma powr_mono2: "0 <= a ==> 0 < x ==> x <= y ==> x powr a <= y powr a"; | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 180 | apply (case_tac "a = 0", simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 181 | apply (case_tac "x = y", simp) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 182 | apply (rule order_less_imp_le) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 183 | apply (rule powr_less_mono2, auto) | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 184 | done | 
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 185 | |
| 
5693a977a767
removed some [iff] declarations from RealDef.thy, concerning inequalities
 paulson parents: 
15053diff
changeset | 186 | |
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 187 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 188 | ML | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 189 | {*
 | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 190 | val powr_one_eq_one = thm "powr_one_eq_one"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 191 | val powr_zero_eq_one = thm "powr_zero_eq_one"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 192 | val powr_one_gt_zero_iff = thm "powr_one_gt_zero_iff"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 193 | val powr_mult = thm "powr_mult"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 194 | val powr_gt_zero = thm "powr_gt_zero"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 195 | val powr_not_zero = thm "powr_not_zero"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 196 | val powr_divide = thm "powr_divide"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 197 | val powr_add = thm "powr_add"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 198 | val powr_powr = thm "powr_powr"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 199 | val powr_powr_swap = thm "powr_powr_swap"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 200 | val powr_minus = thm "powr_minus"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 201 | val powr_minus_divide = thm "powr_minus_divide"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 202 | val powr_less_mono = thm "powr_less_mono"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 203 | val powr_less_cancel = thm "powr_less_cancel"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 204 | val powr_less_cancel_iff = thm "powr_less_cancel_iff"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 205 | val powr_le_cancel_iff = thm "powr_le_cancel_iff"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 206 | val log_ln = thm "log_ln"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 207 | val powr_log_cancel = thm "powr_log_cancel"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 208 | val log_powr_cancel = thm "log_powr_cancel"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 209 | val log_mult = thm "log_mult"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 210 | val log_eq_div_ln_mult_log = thm "log_eq_div_ln_mult_log"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 211 | val log_base_10_eq1 = thm "log_base_10_eq1"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 212 | val log_base_10_eq2 = thm "log_base_10_eq2"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 213 | val log_one = thm "log_one"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 214 | val log_eq_one = thm "log_eq_one"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 215 | val log_inverse = thm "log_inverse"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 216 | val log_divide = thm "log_divide"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 217 | val log_less_cancel_iff = thm "log_less_cancel_iff"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 218 | val log_le_cancel_iff = thm "log_le_cancel_iff"; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 219 | *} | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
12224diff
changeset | 220 | |
| 12224 | 221 | end |