| author | paulson |
| Thu, 05 Feb 2004 10:45:28 +0100 | |
| changeset 14377 | f454b3004f8f |
| parent 13958 | c1c67582c9b5 |
| child 14411 | 7851e526b8b7 |
| permissions | -rw-r--r-- |
(* Title : HLog.thy Author : Jacques D. Fleuriot Copyright : 2000,2001 University of Edinburgh Description : hyperreal base logarithms *) HLog = Log + HTranscendental + constdefs powhr :: [hypreal,hypreal] => hypreal (infixr 80) "x powhr a == ( *f* exp) (a * ( *f* ln) x)" hlog :: [hypreal,hypreal] => hypreal "hlog a x == Abs_hypreal(UN A: Rep_hypreal(a).UN X: Rep_hypreal(x). hyprel `` {%n. log (A n) (X n)})" end