| author | wenzelm | 
| Wed, 23 Mar 2022 12:21:13 +0100 | |
| changeset 75311 | 5960bae73afe | 
| parent 70723 | 4e39d87c9737 | 
| child 80521 | 5c691b178e08 | 
| permissions | -rw-r--r-- | 
| 62479 | 1 | (* Title: HOL/Nonstandard_Analysis/HLog.thy | 
| 2 | Author: Jacques D. Fleuriot | |
| 3 | Copyright: 2000, 2001 University of Edinburgh | |
| 27468 | 4 | *) | 
| 5 | ||
| 64604 | 6 | section \<open>Logarithms: Non-Standard Version\<close> | 
| 27468 | 7 | |
| 8 | theory HLog | |
| 64604 | 9 | imports HTranscendental | 
| 27468 | 10 | begin | 
| 11 | ||
| 64604 | 12 | definition powhr :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hypreal" (infixr "powhr" 80) | 
| 67399 | 13 | where [transfer_unfold]: "x powhr a = starfun2 (powr) x a" | 
| 64604 | 14 | |
| 15 | definition hlog :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hypreal" | |
| 16 | where [transfer_unfold]: "hlog a x = starfun2 log a x" | |
| 17 | ||
| 18 | lemma powhr: "(star_n X) powhr (star_n Y) = star_n (\<lambda>n. (X n) powr (Y n))" | |
| 19 | by (simp add: powhr_def starfun2_star_n) | |
| 27468 | 20 | |
| 64604 | 21 | lemma powhr_one_eq_one [simp]: "\<And>a. 1 powhr a = 1" | 
| 22 | by transfer simp | |
| 27468 | 23 | |
| 64604 | 24 | lemma powhr_mult: "\<And>a x y. 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> (x * y) powhr a = (x powhr a) * (y powhr a)" | 
| 25 | by transfer (simp add: powr_mult) | |
| 27468 | 26 | |
| 64604 | 27 | lemma powhr_gt_zero [simp]: "\<And>a x. 0 < x powhr a \<longleftrightarrow> x \<noteq> 0" | 
| 28 | by transfer simp | |
| 27468 | 29 | |
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
58878diff
changeset | 30 | lemma powhr_not_zero [simp]: "\<And>a x. x powhr a \<noteq> 0 \<longleftrightarrow> x \<noteq> 0" | 
| 64604 | 31 | by transfer simp | 
| 27468 | 32 | |
| 70723 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 33 | lemma powhr_divide: "\<And>a x y. 0 \<le> x \<Longrightarrow> 0 \<le> y \<Longrightarrow> (x / y) powhr a = (x powhr a) / (y powhr a)" | 
| 64604 | 34 | by transfer (rule powr_divide) | 
| 27468 | 35 | |
| 64604 | 36 | lemma powhr_add: "\<And>a b x. x powhr (a + b) = (x powhr a) * (x powhr b)" | 
| 37 | by transfer (rule powr_add) | |
| 27468 | 38 | |
| 64604 | 39 | lemma powhr_powhr: "\<And>a b x. (x powhr a) powhr b = x powhr (a * b)" | 
| 40 | by transfer (rule powr_powr) | |
| 27468 | 41 | |
| 64604 | 42 | lemma powhr_powhr_swap: "\<And>a b x. (x powhr a) powhr b = (x powhr b) powhr a" | 
| 43 | by transfer (rule powr_powr_swap) | |
| 27468 | 44 | |
| 64604 | 45 | lemma powhr_minus: "\<And>a x. x powhr (- a) = inverse (x powhr a)" | 
| 46 | by transfer (rule powr_minus) | |
| 27468 | 47 | |
| 64604 | 48 | lemma powhr_minus_divide: "x powhr (- a) = 1 / (x powhr a)" | 
| 49 | by (simp add: divide_inverse powhr_minus) | |
| 27468 | 50 | |
| 64604 | 51 | lemma powhr_less_mono: "\<And>a b x. a < b \<Longrightarrow> 1 < x \<Longrightarrow> x powhr a < x powhr b" | 
| 52 | by transfer simp | |
| 53 | ||
| 54 | lemma powhr_less_cancel: "\<And>a b x. x powhr a < x powhr b \<Longrightarrow> 1 < x \<Longrightarrow> a < b" | |
| 55 | by transfer simp | |
| 27468 | 56 | |
| 64604 | 57 | lemma powhr_less_cancel_iff [simp]: "1 < x \<Longrightarrow> x powhr a < x powhr b \<longleftrightarrow> a < b" | 
| 58 | by (blast intro: powhr_less_cancel powhr_less_mono) | |
| 27468 | 59 | |
| 64604 | 60 | lemma powhr_le_cancel_iff [simp]: "1 < x \<Longrightarrow> x powhr a \<le> x powhr b \<longleftrightarrow> a \<le> b" | 
| 61 | by (simp add: linorder_not_less [symmetric]) | |
| 27468 | 62 | |
| 64604 | 63 | lemma hlog: "hlog (star_n X) (star_n Y) = star_n (\<lambda>n. log (X n) (Y n))" | 
| 64 | by (simp add: hlog_def starfun2_star_n) | |
| 27468 | 65 | |
| 64604 | 66 | lemma hlog_starfun_ln: "\<And>x. ( *f* ln) x = hlog (( *f* exp) 1) x" | 
| 67 | by transfer (rule log_ln) | |
| 27468 | 68 | |
| 64604 | 69 | lemma powhr_hlog_cancel [simp]: "\<And>a x. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> a powhr (hlog a x) = x" | 
| 70 | by transfer simp | |
| 27468 | 71 | |
| 64604 | 72 | lemma hlog_powhr_cancel [simp]: "\<And>a y. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> hlog a (a powhr y) = y" | 
| 73 | by transfer simp | |
| 27468 | 74 | |
| 75 | lemma hlog_mult: | |
| 64604 | 76 | "\<And>a x y. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> hlog a (x * y) = hlog a x + hlog a y" | 
| 77 | by transfer (rule log_mult) | |
| 27468 | 78 | |
| 64604 | 79 | lemma hlog_as_starfun: "\<And>a x. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> hlog a x = ( *f* ln) x / ( *f* ln) a" | 
| 80 | by transfer (simp add: log_def) | |
| 27468 | 81 | |
| 82 | lemma hlog_eq_div_starfun_ln_mult_hlog: | |
| 64604 | 83 | "\<And>a b x. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < b \<Longrightarrow> b \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> | 
| 84 | hlog a x = (( *f* ln) b / ( *f* ln) a) * hlog b x" | |
| 85 | by transfer (rule log_eq_div_ln_mult_log) | |
| 27468 | 86 | |
| 64604 | 87 | lemma powhr_as_starfun: "\<And>a x. x powhr a = (if x = 0 then 0 else ( *f* exp) (a * ( *f* real_ln) x))" | 
| 88 | by transfer (simp add: powr_def) | |
| 27468 | 89 | |
| 90 | lemma HInfinite_powhr: | |
| 64604 | 91 | "x \<in> HInfinite \<Longrightarrow> 0 < x \<Longrightarrow> a \<in> HFinite - Infinitesimal \<Longrightarrow> 0 < a \<Longrightarrow> x powhr a \<in> HInfinite" | 
| 92 | by (auto intro!: starfun_ln_ge_zero starfun_ln_HInfinite | |
| 93 | HInfinite_HFinite_not_Infinitesimal_mult2 starfun_exp_HInfinite | |
| 94 | simp add: order_less_imp_le HInfinite_gt_zero_gt_one powhr_as_starfun zero_le_mult_iff) | |
| 27468 | 95 | |
| 96 | lemma hlog_hrabs_HInfinite_Infinitesimal: | |
| 64604 | 97 | "x \<in> HFinite - Infinitesimal \<Longrightarrow> a \<in> HInfinite \<Longrightarrow> 0 < a \<Longrightarrow> hlog a \<bar>x\<bar> \<in> Infinitesimal" | 
| 98 | apply (frule HInfinite_gt_zero_gt_one) | |
| 99 | apply (auto intro!: starfun_ln_HFinite_not_Infinitesimal | |
| 100 | HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2 | |
| 101 | simp add: starfun_ln_HInfinite not_Infinitesimal_not_zero | |
| 102 | hlog_as_starfun divide_inverse) | |
| 103 | done | |
| 27468 | 104 | |
| 64604 | 105 | lemma hlog_HInfinite_as_starfun: "a \<in> HInfinite \<Longrightarrow> 0 < a \<Longrightarrow> hlog a x = ( *f* ln) x / ( *f* ln) a" | 
| 106 | by (rule hlog_as_starfun) auto | |
| 27468 | 107 | |
| 64604 | 108 | lemma hlog_one [simp]: "\<And>a. hlog a 1 = 0" | 
| 109 | by transfer simp | |
| 27468 | 110 | |
| 64604 | 111 | lemma hlog_eq_one [simp]: "\<And>a. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> hlog a a = 1" | 
| 112 | by transfer (rule log_eq_one) | |
| 27468 | 113 | |
| 64604 | 114 | lemma hlog_inverse: "0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> hlog a (inverse x) = - hlog a x" | 
| 115 | by (rule add_left_cancel [of "hlog a x", THEN iffD1]) (simp add: hlog_mult [symmetric]) | |
| 27468 | 116 | |
| 64604 | 117 | lemma hlog_divide: "0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> hlog a (x / y) = hlog a x - hlog a y" | 
| 118 | by (simp add: hlog_mult hlog_inverse divide_inverse) | |
| 27468 | 119 | |
| 120 | lemma hlog_less_cancel_iff [simp]: | |
| 64604 | 121 | "\<And>a x y. 1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> hlog a x < hlog a y \<longleftrightarrow> x < y" | 
| 122 | by transfer simp | |
| 27468 | 123 | |
| 64604 | 124 | lemma hlog_le_cancel_iff [simp]: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> hlog a x \<le> hlog a y \<longleftrightarrow> x \<le> y" | 
| 125 | by (simp add: linorder_not_less [symmetric]) | |
| 27468 | 126 | |
| 127 | end |