| author | dixon | 
| Thu, 02 Nov 2006 14:27:18 +0100 | |
| changeset 21145 | 87a03f9b7db2 | 
| parent 19765 | dfe940911617 | 
| child 21404 | eb85850d3eb7 | 
| permissions | -rw-r--r-- | 
| 13958 | 1 | (* Title : HLog.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: 
13958diff
changeset | 6 | header{*Logarithms: Non-Standard Version*}
 | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 7 | |
| 15131 | 8 | theory HLog | 
| 15140 | 9 | imports Log HTranscendental | 
| 15131 | 10 | begin | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 11 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 12 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 13 | (* should be in NSA.ML *) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 14 | lemma epsilon_ge_zero [simp]: "0 \<le> epsilon" | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 15 | by (simp add: epsilon_def star_n_zero_num star_n_le) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 16 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 17 | lemma hpfinite_witness: "epsilon : {x. 0 \<le> x & x : HFinite}"
 | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 18 | by auto | 
| 13958 | 19 | |
| 20 | ||
| 19765 | 21 | definition | 
| 22 | powhr :: "[hypreal,hypreal] => hypreal" (infixr "powhr" 80) | |
| 23 | "x powhr a = starfun2 (op powr) x a" | |
| 13958 | 24 | |
| 19765 | 25 | hlog :: "[hypreal,hypreal] => hypreal" | 
| 26 | "hlog a x = starfun2 log a x" | |
| 13958 | 27 | |
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 28 | declare powhr_def [transfer_unfold] | 
| 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 29 | declare hlog_def [transfer_unfold] | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 30 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 31 | lemma powhr: "(star_n X) powhr (star_n Y) = star_n (%n. (X n) powr (Y n))" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 32 | by (simp add: powhr_def starfun2_star_n) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 33 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 34 | lemma powhr_one_eq_one [simp]: "!!a. 1 powhr a = 1" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 35 | by (transfer, simp) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 36 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 37 | lemma powhr_mult: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 38 | "!!a x y. [| 0 < x; 0 < y |] ==> (x * y) powhr a = (x powhr a) * (y powhr a)" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 39 | by (transfer, rule powr_mult) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 40 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 41 | lemma powhr_gt_zero [simp]: "!!a x. 0 < x powhr a" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 42 | by (transfer, simp) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 43 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 44 | lemma powhr_not_zero [simp]: "x powhr a \<noteq> 0" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 45 | by (rule powhr_gt_zero [THEN hypreal_not_refl2, THEN not_sym]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 46 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 47 | lemma powhr_divide: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 48 | "!!a x y. [| 0 < x; 0 < y |] ==> (x / y) powhr a = (x powhr a)/(y powhr a)" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 49 | by (transfer, rule powr_divide) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 50 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 51 | lemma powhr_add: "!!a b x. x powhr (a + b) = (x powhr a) * (x powhr b)" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 52 | by (transfer, rule powr_add) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 53 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 54 | lemma powhr_powhr: "!!a b x. (x powhr a) powhr b = x powhr (a * b)" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 55 | by (transfer, rule powr_powr) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 56 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 57 | lemma powhr_powhr_swap: "!!a b x. (x powhr a) powhr b = (x powhr b) powhr a" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 58 | by (transfer, rule powr_powr_swap) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 59 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 60 | lemma powhr_minus: "!!a x. x powhr (-a) = inverse (x powhr a)" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 61 | by (transfer, rule powr_minus) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 62 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 63 | lemma powhr_minus_divide: "x powhr (-a) = 1/(x powhr a)" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14411diff
changeset | 64 | by (simp add: divide_inverse powhr_minus) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 65 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 66 | lemma powhr_less_mono: "!!a b x. [| a < b; 1 < x |] ==> x powhr a < x powhr b" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 67 | by (transfer, simp) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 68 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 69 | lemma powhr_less_cancel: "!!a b x. [| x powhr a < x powhr b; 1 < x |] ==> a < b" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 70 | by (transfer, simp) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 71 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 72 | lemma powhr_less_cancel_iff [simp]: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 73 | "1 < x ==> (x powhr a < x powhr b) = (a < b)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 74 | by (blast intro: powhr_less_cancel powhr_less_mono) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 75 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 76 | lemma powhr_le_cancel_iff [simp]: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 77 | "1 < x ==> (x powhr a \<le> x powhr b) = (a \<le> b)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 78 | by (simp add: linorder_not_less [symmetric]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 79 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 80 | lemma hlog: | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 81 | "hlog (star_n X) (star_n Y) = | 
| 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 82 | star_n (%n. log (X n) (Y n))" | 
| 17429 
e8d6ed3aacfe
merged Transfer.thy and StarType.thy into StarDef.thy; renamed Ifun2_of to starfun2; cleaned up
 huffman parents: 
17332diff
changeset | 83 | by (simp add: hlog_def starfun2_star_n) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 84 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 85 | lemma hlog_starfun_ln: "!!x. ( *f* ln) x = hlog (( *f* exp) 1) x" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 86 | by (transfer, rule log_ln) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 87 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 88 | lemma powhr_hlog_cancel [simp]: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 89 | "!!a x. [| 0 < a; a \<noteq> 1; 0 < x |] ==> a powhr (hlog a x) = x" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 90 | by (transfer, simp) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 91 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 92 | lemma hlog_powhr_cancel [simp]: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 93 | "!!a y. [| 0 < a; a \<noteq> 1 |] ==> hlog a (a powhr y) = y" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 94 | by (transfer, simp) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 95 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 96 | lemma hlog_mult: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 97 | "!!a x y. [| 0 < a; a \<noteq> 1; 0 < x; 0 < y |] | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 98 | ==> hlog a (x * y) = hlog a x + hlog a y" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 99 | by (transfer, rule log_mult) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 100 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 101 | lemma hlog_as_starfun: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 102 | "!!a x. [| 0 < a; a \<noteq> 1 |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 103 | by (transfer, simp add: log_def) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 104 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 105 | lemma hlog_eq_div_starfun_ln_mult_hlog: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 106 | "!!a b x. [| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |] | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 107 | ==> hlog a x = (( *f* ln) b/( *f*ln) a) * hlog b x" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 108 | by (transfer, rule log_eq_div_ln_mult_log) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 109 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 110 | lemma powhr_as_starfun: "!!a x. x powhr a = ( *f* exp) (a * ( *f* ln) x)" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 111 | by (transfer, simp add: powr_def) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 112 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 113 | lemma HInfinite_powhr: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 114 | "[| x : HInfinite; 0 < x; a : HFinite - Infinitesimal; | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 115 | 0 < a |] ==> x powhr a : HInfinite" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 116 | apply (auto intro!: starfun_ln_ge_zero starfun_ln_HInfinite HInfinite_HFinite_not_Infinitesimal_mult2 starfun_exp_HInfinite | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 117 | simp add: order_less_imp_le HInfinite_gt_zero_gt_one powhr_as_starfun zero_le_mult_iff) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 118 | done | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 119 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 120 | lemma hlog_hrabs_HInfinite_Infinitesimal: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 121 | "[| x : HFinite - Infinitesimal; a : HInfinite; 0 < a |] | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 122 | ==> hlog a (abs x) : Infinitesimal" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 123 | apply (frule HInfinite_gt_zero_gt_one) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 124 | apply (auto intro!: starfun_ln_HFinite_not_Infinitesimal | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 125 | HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2 | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 126 | simp add: starfun_ln_HInfinite not_Infinitesimal_not_zero | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14411diff
changeset | 127 | hlog_as_starfun hypreal_not_refl2 [THEN not_sym] divide_inverse) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 128 | done | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 129 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 130 | lemma hlog_HInfinite_as_starfun: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 131 | "[| a : HInfinite; 0 < a |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 132 | by (rule hlog_as_starfun, auto) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 133 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 134 | lemma hlog_one [simp]: "!!a. hlog a 1 = 0" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 135 | by (transfer, simp) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 136 | |
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 137 | lemma hlog_eq_one [simp]: "!!a. [| 0 < a; a \<noteq> 1 |] ==> hlog a a = 1" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 138 | by (transfer, rule log_eq_one) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 139 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 140 | lemma hlog_inverse: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 141 | "[| 0 < a; a \<noteq> 1; 0 < x |] ==> hlog a (inverse x) = - hlog a x" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 142 | apply (rule add_left_cancel [of "hlog a x", THEN iffD1]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 143 | apply (simp add: hlog_mult [symmetric]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 144 | done | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 145 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 146 | lemma hlog_divide: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 147 | "[| 0 < a; a \<noteq> 1; 0 < x; 0 < y|] ==> hlog a (x/y) = hlog a x - hlog a y" | 
| 14430 
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
 paulson parents: 
14411diff
changeset | 148 | by (simp add: hlog_mult hlog_inverse divide_inverse) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 149 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 150 | lemma hlog_less_cancel_iff [simp]: | 
| 17318 
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
 huffman parents: 
17298diff
changeset | 151 | "!!a x y. [| 1 < a; 0 < x; 0 < y |] ==> (hlog a x < hlog a y) = (x < y)" | 
| 17332 
4910cf8c0cd2
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
 huffman parents: 
17318diff
changeset | 152 | by (transfer, simp) | 
| 14411 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 153 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 154 | lemma hlog_le_cancel_iff [simp]: | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 155 | "[| 1 < a; 0 < x; 0 < y |] ==> (hlog a x \<le> hlog a y) = (x \<le> y)" | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 156 | by (simp add: linorder_not_less [symmetric]) | 
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 157 | |
| 
7851e526b8b7
converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
 paulson parents: 
13958diff
changeset | 158 | end |