src/HOL/Hyperreal/HLog.thy
author paulson
Wed, 27 Jul 2005 11:28:18 +0200
changeset 16924 04246269386e
parent 15140 322485b816ac
child 17298 ad73fb6144cf
permissions -rw-r--r--
removed the dependence on abs_mult
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"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    15
by (simp add: epsilon_def hypreal_zero_num hypreal_le)
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)
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    24
    "x powhr a == ( *f* exp) (a * ( *f* ln) x)"
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"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    27
    "hlog a x == Abs_hypreal(\<Union>A \<in> Rep_hypreal(a).\<Union>X \<in> Rep_hypreal(x).
13958
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    28
			     hyprel `` {%n. log (A n) (X n)})"
c1c67582c9b5 New material on integration, etc. Moving Hyperreal/ex
paulson
parents:
diff changeset
    29
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    30
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    31
lemma powhr: 
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    32
    "(Abs_hypreal(hyprel `` {X})) powhr (Abs_hypreal(hyprel `` {Y})) =  
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    33
     Abs_hypreal(hyprel `` {%n.  (X n) powr (Y n)})"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    34
by (simp add: powhr_def starfun hypreal_mult powr_def)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    35
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    36
lemma powhr_one_eq_one [simp]: "1 powhr a = 1"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    37
apply (cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    38
apply (simp add: powhr hypreal_one_num)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    39
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    40
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    41
lemma powhr_mult:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    42
     "[| 0 < x; 0 < y |] ==> (x * y) powhr a = (x powhr a) * (y powhr a)"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    43
apply (cases x, cases y, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    44
apply (simp add: powhr hypreal_zero_num hypreal_mult hypreal_less, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    45
apply (simp add: powr_mult) 
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    46
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    47
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    48
lemma powhr_gt_zero [simp]: "0 < x powhr a"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    49
apply (cases a, cases x)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    50
apply (simp add: hypreal_zero_def powhr hypreal_less hypreal_zero_num)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    51
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    52
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    53
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
    54
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
    55
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    56
lemma hypreal_divide: 
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    57
     "(Abs_hypreal(hyprel `` {X}))/(Abs_hypreal(hyprel `` {Y})) =  
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    58
      (Abs_hypreal(hyprel `` {%n. (X n)/(Y n)}))"
14430
5cb24165a2e1 new material from Avigad, and simplified treatment of division by 0
paulson
parents: 14411
diff changeset
    59
by (simp add: divide_inverse hypreal_zero_num hypreal_inverse hypreal_mult) 
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    60
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    61
lemma powhr_divide:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    62
     "[| 0 < x; 0 < y |] ==> (x / y) powhr a = (x powhr a)/(y powhr a)"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    63
apply (cases x, cases y, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    64
apply (simp add: powhr hypreal_divide hypreal_zero_num hypreal_less, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    65
apply (simp add: powr_divide)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    66
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    67
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    68
lemma powhr_add: "x powhr (a + b) = (x powhr a) * (x powhr b)"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    69
apply (cases x, cases b, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    70
apply (simp add: powhr hypreal_add hypreal_mult powr_add)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    71
done
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_powhr: "(x powhr a) powhr b = x powhr (a * b)"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    74
apply (cases x, cases b, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    75
apply (simp add: powhr hypreal_mult powr_powr)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    76
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    77
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    78
lemma powhr_powhr_swap: "(x powhr a) powhr b = (x powhr b) powhr a"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    79
apply (cases x, cases b, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    80
apply (simp add: powhr powr_powr_swap)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    81
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    82
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    83
lemma powhr_minus: "x powhr (-a) = inverse (x powhr a)"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    84
apply (cases x, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    85
apply (simp add: powhr hypreal_minus hypreal_inverse hypreal_less powr_minus)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    86
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    87
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    88
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
    89
by (simp add: divide_inverse powhr_minus)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    90
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    91
lemma powhr_less_mono: "[| a < b; 1 < x |] ==> x powhr a < x powhr b"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    92
apply (cases x, cases a, cases b)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    93
apply (simp add: powhr hypreal_one_num hypreal_less, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    94
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    95
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    96
lemma powhr_less_cancel: "[| x powhr a < x powhr b; 1 < x |] ==> a < b"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
    97
apply (cases x, cases a, cases b)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    98
apply (simp add: powhr hypreal_one_num hypreal_less, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
    99
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   100
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   101
lemma powhr_less_cancel_iff [simp]:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   102
     "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
   103
by (blast intro: powhr_less_cancel powhr_less_mono)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   104
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   105
lemma powhr_le_cancel_iff [simp]:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   106
     "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
   107
by (simp add: linorder_not_less [symmetric])
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   108
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   109
lemma hlog: 
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   110
     "hlog (Abs_hypreal(hyprel `` {X})) (Abs_hypreal(hyprel `` {Y})) =  
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   111
      Abs_hypreal(hyprel `` {%n. log (X n) (Y n)})"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   112
apply (simp add: hlog_def)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   113
apply (rule arg_cong [where f=Abs_hypreal], auto, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   114
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   115
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   116
lemma hlog_starfun_ln: "( *f* ln) x = hlog (( *f* exp) 1) x"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   117
apply (cases x)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   118
apply (simp add: starfun hlog log_ln hypreal_one_num)
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 powhr_hlog_cancel [simp]:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   122
     "[| 0 < a; a \<noteq> 1; 0 < x |] ==> a powhr (hlog a x) = x"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   123
apply (cases x, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   124
apply (simp add: hlog powhr hypreal_zero_num hypreal_less hypreal_one_num, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   125
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   126
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   127
lemma hlog_powhr_cancel [simp]:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   128
     "[| 0 < a; a \<noteq> 1 |] ==> hlog a (a powhr y) = y"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   129
apply (cases y, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   130
apply (simp add: hlog powhr hypreal_zero_num hypreal_less hypreal_one_num, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   131
apply (auto intro: log_powr_cancel) 
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   132
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   133
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   134
lemma hlog_mult:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   135
     "[| 0 < a; a \<noteq> 1; 0 < x; 0 < y  |]  
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   136
      ==> hlog a (x * y) = hlog a x + hlog a y"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   137
apply (cases x, cases y, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   138
apply (simp add: hlog powhr hypreal_zero_num hypreal_one_num hypreal_less hypreal_add hypreal_mult, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   139
apply (simp add: log_mult)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   140
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   141
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   142
lemma hlog_as_starfun:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   143
     "[| 0 < a; a \<noteq> 1 |] ==> hlog a x = ( *f* ln) x / ( *f* ln) a"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   144
apply (cases x, cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   145
apply (simp add: hlog starfun hypreal_zero_num hypreal_one_num hypreal_divide log_def)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   146
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   147
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   148
lemma hlog_eq_div_starfun_ln_mult_hlog:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   149
     "[| 0 < a; a \<noteq> 1; 0 < b; b \<noteq> 1; 0 < x |]  
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   150
      ==> hlog a x = (( *f* ln) b/( *f*ln) a) * hlog b x"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   151
apply (cases x, cases a, cases b)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   152
apply (simp add: hlog starfun hypreal_zero_num hypreal_one_num hypreal_less hypreal_divide hypreal_mult, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   153
apply (auto dest: log_eq_div_ln_mult_log) 
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   154
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   155
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   156
lemma powhr_as_starfun: "x powhr a = ( *f* exp) (a * ( *f* ln) x)"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   157
apply (cases a, cases x)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   158
apply (simp add: powhr starfun hypreal_mult powr_def)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   159
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   160
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   161
lemma HInfinite_powhr:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   162
     "[| x : HInfinite; 0 < x; a : HFinite - Infinitesimal;  
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   163
         0 < a |] ==> x powhr a : HInfinite"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   164
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
   165
       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
   166
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   167
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   168
lemma hlog_hrabs_HInfinite_Infinitesimal:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   169
     "[| x : HFinite - Infinitesimal; a : HInfinite; 0 < a |]  
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   170
      ==> hlog a (abs x) : Infinitesimal"
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   171
apply (frule HInfinite_gt_zero_gt_one)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   172
apply (auto intro!: starfun_ln_HFinite_not_Infinitesimal
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   173
            HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2 
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   174
        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
   175
          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
   176
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   177
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   178
lemma hlog_HInfinite_as_starfun:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   179
     "[| 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
   180
by (rule hlog_as_starfun, auto)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   181
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   182
lemma hlog_one [simp]: "hlog a 1 = 0"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   183
apply (cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   184
apply (simp add: hypreal_one_num hypreal_zero_num hlog)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   185
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   186
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   187
lemma hlog_eq_one [simp]: "[| 0 < a; a \<noteq> 1 |] ==> hlog a a = 1"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   188
apply (cases a)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   189
apply (simp add: hypreal_one_num hypreal_zero_num hlog hypreal_less, ultra)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   190
apply (auto intro: log_eq_one) 
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   191
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   192
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   193
lemma hlog_inverse:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   194
     "[| 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
   195
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
   196
apply (simp add: hlog_mult [symmetric])
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   197
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   198
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   199
lemma hlog_divide:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   200
     "[| 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
   201
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
   202
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   203
lemma hlog_less_cancel_iff [simp]:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   204
     "[| 1 < a; 0 < x; 0 < y |] ==> (hlog a x < hlog a y) = (x < y)"
14468
6be497cacab5 heavy tidying
paulson
parents: 14430
diff changeset
   205
apply (cases a, cases x, cases y)
14411
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   206
apply (auto simp add: hlog hypreal_less hypreal_zero_num hypreal_one_num, ultra+)
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   207
done
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   208
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   209
lemma hlog_le_cancel_iff [simp]:
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   210
     "[| 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
   211
by (simp add: linorder_not_less [symmetric])
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   212
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   213
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   214
ML
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   215
{*
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   216
val powhr = thm "powhr";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   217
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
   218
val powhr_mult = thm "powhr_mult";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   219
val powhr_gt_zero = thm "powhr_gt_zero";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   220
val powhr_not_zero = thm "powhr_not_zero";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   221
val hypreal_divide = thm "hypreal_divide";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   222
val powhr_divide = thm "powhr_divide";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   223
val powhr_add = thm "powhr_add";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   224
val powhr_powhr = thm "powhr_powhr";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   225
val powhr_powhr_swap = thm "powhr_powhr_swap";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   226
val powhr_minus = thm "powhr_minus";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   227
val powhr_minus_divide = thm "powhr_minus_divide";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   228
val powhr_less_mono = thm "powhr_less_mono";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   229
val powhr_less_cancel = thm "powhr_less_cancel";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   230
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
   231
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
   232
val hlog = thm "hlog";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   233
val hlog_starfun_ln = thm "hlog_starfun_ln";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   234
val powhr_hlog_cancel = thm "powhr_hlog_cancel";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   235
val hlog_powhr_cancel = thm "hlog_powhr_cancel";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   236
val hlog_mult = thm "hlog_mult";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   237
val hlog_as_starfun = thm "hlog_as_starfun";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   238
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
   239
val powhr_as_starfun = thm "powhr_as_starfun";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   240
val HInfinite_powhr = thm "HInfinite_powhr";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   241
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
   242
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
   243
val hlog_one = thm "hlog_one";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   244
val hlog_eq_one = thm "hlog_eq_one";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   245
val hlog_inverse = thm "hlog_inverse";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   246
val hlog_divide = thm "hlog_divide";
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   247
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
   248
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
   249
*}
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   250
7851e526b8b7 converted Hyperreal/Log and Hyperreal/HLog to Isar scripts
paulson
parents: 13958
diff changeset
   251
end