src/HOL/Hyperreal/HLog.thy
author huffman
Mon, 12 Sep 2005 23:14:41 +0200
changeset 17332 4910cf8c0cd2
parent 17318 bc1c75855f3d
child 17429 e8d6ed3aacfe
permissions -rw-r--r--
added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     1
(*  Title       : HLog.thy
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     3
    Copyright   : 2000,2001 University of Edinburgh
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     4
*)
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
     5
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
     6
header{*Logarithms: Non-Standard Version*}
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
     7
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14468
diff changeset
     8
theory HLog
15140
322485b816ac import -> imports
nipkow
parents: 15131
diff changeset
     9
imports Log HTranscendental
15131
c69542757a4d New theory header syntax.
nipkow
parents: 14468
diff changeset
    10
begin
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    11
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    12
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    13
(* should be in NSA.ML *)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff 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: 17298
diff 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: 13958
diff changeset
    16
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    17
lemma hpfinite_witness: "epsilon : {x. 0 \<le> x & x : HFinite}"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    18
by auto
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    19
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    20
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    21
constdefs
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    22
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    23
    powhr  :: "[hypreal,hypreal] => hypreal"     (infixr "powhr" 80)
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    24
    "x powhr a == Ifun2_of (op powr) x a"
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    25
  
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    26
    hlog :: "[hypreal,hypreal] => hypreal"
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    27
    "hlog a x == Ifun2_of log a x"
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    28
17332
4910cf8c0cd2 added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents: 17318
diff changeset
    29
declare powhr_def [transfer_unfold]
4910cf8c0cd2 added theorem attributes transfer_intro, transfer_unfold, transfer_refold; simplified some proofs; some rearranging
huffman
parents: 17318
diff changeset
    30
declare hlog_def [transfer_unfold]
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    31
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    32
lemma powhr: "(star_n X) powhr (star_n Y) = star_n (%n. (X n) powr (Y n))"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    33
by (simp add: powhr_def Ifun2_of_def star_of_def Ifun_star_n)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    34
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    35
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: 17318
diff changeset
    36
by (transfer, simp)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    37
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    38
lemma powhr_mult:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    39
  "!!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: 17318
diff changeset
    40
by (transfer, rule powr_mult)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    41
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    42
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: 17318
diff changeset
    43
by (transfer, simp)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    44
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    45
lemma powhr_not_zero [simp]: "x powhr a \<noteq> 0"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    46
by (rule powhr_gt_zero [THEN hypreal_not_refl2, THEN not_sym])
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    47
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    48
lemma powhr_divide:
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    49
  "!!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: 17318
diff changeset
    50
by (transfer, rule powr_divide)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    51
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    52
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: 17318
diff changeset
    53
by (transfer, rule powr_add)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    54
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    55
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: 17318
diff changeset
    56
by (transfer, rule powr_powr)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    57
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    58
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: 17318
diff changeset
    59
by (transfer, rule powr_powr_swap)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    60
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    61
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: 17318
diff changeset
    62
by (transfer, rule powr_minus)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    63
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    64
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: 14411
diff changeset
    65
by (simp add: divide_inverse powhr_minus)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    66
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    67
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: 17318
diff changeset
    68
by (transfer, simp)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    69
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    70
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: 17318
diff changeset
    71
by (transfer, simp)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    72
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    73
lemma powhr_less_cancel_iff [simp]:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    74
     "1 < x ==> (x powhr a < x powhr b) = (a < b)"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    75
by (blast intro: powhr_less_cancel powhr_less_mono)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    76
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    77
lemma powhr_le_cancel_iff [simp]:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    78
     "1 < x ==> (x powhr a \<le> x powhr b) = (a \<le> b)"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    79
by (simp add: linorder_not_less [symmetric])
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    80
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    81
lemma hlog:
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    82
     "hlog (star_n X) (star_n Y) =  
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    83
      star_n (%n. log (X n) (Y n))"
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    84
by (simp add: hlog_def Ifun2_of_def star_of_def Ifun_star_n)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    85
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    86
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: 17318
diff changeset
    87
by (transfer, rule log_ln)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    88
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    89
lemma powhr_hlog_cancel [simp]:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    90
     "!!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: 17318
diff changeset
    91
by (transfer, simp)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    92
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    93
lemma hlog_powhr_cancel [simp]:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    94
     "!!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: 17318
diff changeset
    95
by (transfer, simp)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    96
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    97
lemma hlog_mult:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
    98
     "!!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: 13958
diff changeset
    99
      ==> 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: 17318
diff changeset
   100
by (transfer, rule log_mult)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   101
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   102
lemma hlog_as_starfun:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   103
     "!!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: 17318
diff changeset
   104
by (transfer, simp add: log_def)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   105
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   106
lemma hlog_eq_div_starfun_ln_mult_hlog:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   107
     "!!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: 13958
diff changeset
   108
      ==> 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: 17318
diff changeset
   109
by (transfer, rule log_eq_div_ln_mult_log)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   110
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   111
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: 17318
diff changeset
   112
by (transfer, simp add: powr_def)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   113
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   114
lemma HInfinite_powhr:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   115
     "[| x : HInfinite; 0 < x; a : HFinite - Infinitesimal;  
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   116
         0 < a |] ==> x powhr a : HInfinite"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   117
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: 13958
diff changeset
   118
       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: 13958
diff changeset
   119
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   120
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   121
lemma hlog_hrabs_HInfinite_Infinitesimal:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   122
     "[| x : HFinite - Infinitesimal; a : HInfinite; 0 < a |]  
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   123
      ==> hlog a (abs x) : Infinitesimal"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   124
apply (frule HInfinite_gt_zero_gt_one)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   125
apply (auto intro!: starfun_ln_HFinite_not_Infinitesimal
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   126
            HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2 
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   127
        simp add: starfun_ln_HInfinite not_Infinitesimal_not_zero 
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14411
diff changeset
   128
          hlog_as_starfun hypreal_not_refl2 [THEN not_sym] divide_inverse)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   129
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   130
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   131
lemma hlog_HInfinite_as_starfun:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   132
     "[| 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: 13958
diff changeset
   133
by (rule hlog_as_starfun, auto)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   134
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   135
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: 17318
diff changeset
   136
by (transfer, simp)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   137
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   138
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: 17318
diff changeset
   139
by (transfer, rule log_eq_one)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   140
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   141
lemma hlog_inverse:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   142
     "[| 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: 13958
diff changeset
   143
apply (rule add_left_cancel [of "hlog a x", THEN iffD1])
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   144
apply (simp add: hlog_mult [symmetric])
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   145
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   146
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   147
lemma hlog_divide:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   148
     "[| 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: 14411
diff changeset
   149
by (simp add: hlog_mult hlog_inverse divide_inverse)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   150
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   151
lemma hlog_less_cancel_iff [simp]:
17318
bc1c75855f3d starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents: 17298
diff changeset
   152
     "!!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: 17318
diff changeset
   153
by (transfer, simp)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   154
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   155
lemma hlog_le_cancel_iff [simp]:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   156
     "[| 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: 13958
diff changeset
   157
by (simp add: linorder_not_less [symmetric])
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   158
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   159
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   160
ML
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   161
{*
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   162
val powhr = thm "powhr";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   163
val powhr_one_eq_one = thm "powhr_one_eq_one";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   164
val powhr_mult = thm "powhr_mult";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   165
val powhr_gt_zero = thm "powhr_gt_zero";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   166
val powhr_not_zero = thm "powhr_not_zero";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   167
val powhr_divide = thm "powhr_divide";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   168
val powhr_add = thm "powhr_add";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   169
val powhr_powhr = thm "powhr_powhr";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   170
val powhr_powhr_swap = thm "powhr_powhr_swap";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   171
val powhr_minus = thm "powhr_minus";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   172
val powhr_minus_divide = thm "powhr_minus_divide";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   173
val powhr_less_mono = thm "powhr_less_mono";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   174
val powhr_less_cancel = thm "powhr_less_cancel";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   175
val powhr_less_cancel_iff = thm "powhr_less_cancel_iff";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   176
val powhr_le_cancel_iff = thm "powhr_le_cancel_iff";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   177
val hlog = thm "hlog";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   178
val hlog_starfun_ln = thm "hlog_starfun_ln";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   179
val powhr_hlog_cancel = thm "powhr_hlog_cancel";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   180
val hlog_powhr_cancel = thm "hlog_powhr_cancel";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   181
val hlog_mult = thm "hlog_mult";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   182
val hlog_as_starfun = thm "hlog_as_starfun";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   183
val hlog_eq_div_starfun_ln_mult_hlog = thm "hlog_eq_div_starfun_ln_mult_hlog";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   184
val powhr_as_starfun = thm "powhr_as_starfun";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   185
val HInfinite_powhr = thm "HInfinite_powhr";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   186
val hlog_hrabs_HInfinite_Infinitesimal = thm "hlog_hrabs_HInfinite_Infinitesimal";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   187
val hlog_HInfinite_as_starfun = thm "hlog_HInfinite_as_starfun";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   188
val hlog_one = thm "hlog_one";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   189
val hlog_eq_one = thm "hlog_eq_one";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   190
val hlog_inverse = thm "hlog_inverse";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   191
val hlog_divide = thm "hlog_divide";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   192
val hlog_less_cancel_iff = thm "hlog_less_cancel_iff";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   193
val hlog_le_cancel_iff = thm "hlog_le_cancel_iff";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   194
*}
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   195
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   196
end