src/HOL/Hyperreal/HLog.thy
changeset 17332 4910cf8c0cd2
parent 17318 bc1c75855f3d
child 17429 e8d6ed3aacfe
--- a/src/HOL/Hyperreal/HLog.thy	Mon Sep 12 23:13:46 2005 +0200
+++ b/src/HOL/Hyperreal/HLog.thy	Mon Sep 12 23:14:41 2005 +0200
@@ -26,47 +26,49 @@
     hlog :: "[hypreal,hypreal] => hypreal"
     "hlog a x == Ifun2_of log a x"
 
+declare powhr_def [transfer_unfold]
+declare hlog_def [transfer_unfold]
 
 lemma powhr: "(star_n X) powhr (star_n Y) = star_n (%n. (X n) powr (Y n))"
 by (simp add: powhr_def Ifun2_of_def star_of_def Ifun_star_n)
 
 lemma powhr_one_eq_one [simp]: "!!a. 1 powhr a = 1"
-by (unfold powhr_def, transfer, simp)
+by (transfer, simp)
 
 lemma powhr_mult:
   "!!a x y. [| 0 < x; 0 < y |] ==> (x * y) powhr a = (x powhr a) * (y powhr a)"
-by (unfold powhr_def, transfer, rule powr_mult)
+by (transfer, rule powr_mult)
 
 lemma powhr_gt_zero [simp]: "!!a x. 0 < x powhr a"
-by (unfold powhr_def, transfer, simp)
+by (transfer, simp)
 
 lemma powhr_not_zero [simp]: "x powhr a \<noteq> 0"
 by (rule powhr_gt_zero [THEN hypreal_not_refl2, THEN not_sym])
 
 lemma powhr_divide:
   "!!a x y. [| 0 < x; 0 < y |] ==> (x / y) powhr a = (x powhr a)/(y powhr a)"
-by (unfold powhr_def, transfer, rule powr_divide)
+by (transfer, rule powr_divide)
 
 lemma powhr_add: "!!a b x. x powhr (a + b) = (x powhr a) * (x powhr b)"
-by (unfold powhr_def, transfer, rule powr_add)
+by (transfer, rule powr_add)
 
 lemma powhr_powhr: "!!a b x. (x powhr a) powhr b = x powhr (a * b)"
-by (unfold powhr_def, transfer, rule powr_powr)
+by (transfer, rule powr_powr)
 
 lemma powhr_powhr_swap: "!!a b x. (x powhr a) powhr b = (x powhr b) powhr a"
-by (unfold powhr_def, transfer, rule powr_powr_swap)
+by (transfer, rule powr_powr_swap)
 
 lemma powhr_minus: "!!a x. x powhr (-a) = inverse (x powhr a)"
-by (unfold powhr_def, transfer, rule powr_minus)
+by (transfer, rule powr_minus)
 
 lemma powhr_minus_divide: "x powhr (-a) = 1/(x powhr a)"
 by (simp add: divide_inverse powhr_minus)
 
 lemma powhr_less_mono: "!!a b x. [| a < b; 1 < x |] ==> x powhr a < x powhr b"
-by (unfold powhr_def, transfer, simp)
+by (transfer, simp)
 
 lemma powhr_less_cancel: "!!a b x. [| x powhr a < x powhr b; 1 < x |] ==> a < b"
-by (unfold powhr_def, transfer, simp)
+by (transfer, simp)
 
 lemma powhr_less_cancel_iff [simp]:
      "1 < x ==> (x powhr a < x powhr b) = (a < b)"
@@ -82,32 +84,32 @@
 by (simp add: hlog_def Ifun2_of_def star_of_def Ifun_star_n)
 
 lemma hlog_starfun_ln: "!!x. ( *f* ln) x = hlog (( *f* exp) 1) x"
-by (unfold hlog_def, transfer, rule log_ln)
+by (transfer, rule log_ln)
 
 lemma powhr_hlog_cancel [simp]:
      "!!a x. [| 0 < a; a \<noteq> 1; 0 < x |] ==> a powhr (hlog a x) = x"
-by (unfold powhr_def hlog_def, transfer, simp)
+by (transfer, simp)
 
 lemma hlog_powhr_cancel [simp]:
      "!!a y. [| 0 < a; a \<noteq> 1 |] ==> hlog a (a powhr y) = y"
-by (unfold powhr_def hlog_def, transfer, simp)
+by (transfer, simp)
 
 lemma hlog_mult:
      "!!a x y. [| 0 < a; a \<noteq> 1; 0 < x; 0 < y  |]  
       ==> hlog a (x * y) = hlog a x + hlog a y"
-by (unfold powhr_def hlog_def, transfer, rule log_mult)
+by (transfer, rule log_mult)
 
 lemma hlog_as_starfun:
      "!!a x. [| 0 < a; a \<noteq> 1 |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a"
-by (unfold hlog_def, transfer, simp add: log_def)
+by (transfer, simp add: log_def)
 
 lemma hlog_eq_div_starfun_ln_mult_hlog:
      "!!a b x. [| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |]  
       ==> hlog a x = (( *f* ln) b/( *f*ln) a) * hlog b x"
-by (unfold hlog_def, transfer, rule log_eq_div_ln_mult_log)
+by (transfer, rule log_eq_div_ln_mult_log)
 
 lemma powhr_as_starfun: "!!a x. x powhr a = ( *f* exp) (a * ( *f* ln) x)"
-by (unfold powhr_def, transfer, simp add: powr_def)
+by (transfer, simp add: powr_def)
 
 lemma HInfinite_powhr:
      "[| x : HInfinite; 0 < x; a : HFinite - Infinitesimal;  
@@ -131,10 +133,10 @@
 by (rule hlog_as_starfun, auto)
 
 lemma hlog_one [simp]: "!!a. hlog a 1 = 0"
-by (unfold hlog_def, transfer, simp)
+by (transfer, simp)
 
 lemma hlog_eq_one [simp]: "!!a. [| 0 < a; a \<noteq> 1 |] ==> hlog a a = 1"
-by (unfold hlog_def, transfer, rule log_eq_one)
+by (transfer, rule log_eq_one)
 
 lemma hlog_inverse:
      "[| 0 < a; a \<noteq> 1; 0 < x |] ==> hlog a (inverse x) = - hlog a x"
@@ -148,7 +150,7 @@
 
 lemma hlog_less_cancel_iff [simp]:
      "!!a x y. [| 1 < a; 0 < x; 0 < y |] ==> (hlog a x < hlog a y) = (x < y)"
-by (unfold hlog_def, transfer, simp)
+by (transfer, simp)
 
 lemma hlog_le_cancel_iff [simp]:
      "[| 1 < a; 0 < x; 0 < y |] ==> (hlog a x \<le> hlog a y) = (x \<le> y)"