src/HOL/Nonstandard_Analysis/HLog.thy
author nipkow
Wed, 10 Jan 2018 15:25:09 +0100
changeset 67399 eab6ce8368fa
parent 64604 2bf8cfc98c4d
child 70723 4e39d87c9737
permissions -rw-r--r--
ran isabelle update_op on all sources
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62479
716336f19aa9 clarified session;
wenzelm
parents: 61981
diff changeset
     1
(*  Title:      HOL/Nonstandard_Analysis/HLog.thy
716336f19aa9 clarified session;
wenzelm
parents: 61981
diff changeset
     2
    Author:     Jacques D. Fleuriot
716336f19aa9 clarified session;
wenzelm
parents: 61981
diff changeset
     3
    Copyright:  2000, 2001 University of Edinburgh
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     4
*)
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     5
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
     6
section \<open>Logarithms: Non-Standard Version\<close>
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     7
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
     8
theory HLog
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
     9
  imports HTranscendental
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    10
begin
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    11
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    12
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    13
(* should be in NSA.ML *)
61981
1b5845c62fa0 more symbols;
wenzelm
parents: 61975
diff changeset
    14
lemma epsilon_ge_zero [simp]: "0 \<le> \<epsilon>"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    15
  by (simp add: epsilon_def star_n_zero_num star_n_le)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    16
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    17
lemma hpfinite_witness: "\<epsilon> \<in> {x. 0 \<le> x \<and> x \<in> HFinite}"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    18
  by auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    19
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    20
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    21
definition powhr :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hypreal"  (infixr "powhr" 80)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 64604
diff changeset
    22
  where [transfer_unfold]: "x powhr a = starfun2 (powr) x a"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    23
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    24
definition hlog :: "hypreal \<Rightarrow> hypreal \<Rightarrow> hypreal"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    25
  where [transfer_unfold]: "hlog a x = starfun2 log a x"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    26
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    27
lemma powhr: "(star_n X) powhr (star_n Y) = star_n (\<lambda>n. (X n) powr (Y n))"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    28
  by (simp add: powhr_def starfun2_star_n)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    29
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    30
lemma powhr_one_eq_one [simp]: "\<And>a. 1 powhr a = 1"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    31
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    32
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    33
lemma powhr_mult: "\<And>a x y. 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> (x * y) powhr a = (x powhr a) * (y powhr a)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    34
  by transfer (simp add: powr_mult)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    35
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    36
lemma powhr_gt_zero [simp]: "\<And>a x. 0 < x powhr a \<longleftrightarrow> x \<noteq> 0"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    37
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    38
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: 58878
diff changeset
    39
lemma powhr_not_zero [simp]: "\<And>a x. x powhr a \<noteq> 0 \<longleftrightarrow> x \<noteq> 0"
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    40
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    41
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    42
lemma powhr_divide: "\<And>a x y. 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> (x / y) powhr a = (x powhr a) / (y powhr a)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    43
  by transfer (rule powr_divide)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    44
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    45
lemma powhr_add: "\<And>a b x. x powhr (a + b) = (x powhr a) * (x powhr b)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    46
  by transfer (rule powr_add)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    47
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    48
lemma powhr_powhr: "\<And>a b x. (x powhr a) powhr b = x powhr (a * b)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    49
  by transfer (rule powr_powr)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    50
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    51
lemma powhr_powhr_swap: "\<And>a b x. (x powhr a) powhr b = (x powhr b) powhr a"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    52
  by transfer (rule powr_powr_swap)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    53
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    54
lemma powhr_minus: "\<And>a x. x powhr (- a) = inverse (x powhr a)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    55
  by transfer (rule powr_minus)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    56
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    57
lemma powhr_minus_divide: "x powhr (- a) = 1 / (x powhr a)"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    58
  by (simp add: divide_inverse powhr_minus)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    59
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    60
lemma powhr_less_mono: "\<And>a b x. a < b \<Longrightarrow> 1 < x \<Longrightarrow> x powhr a < x powhr b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    61
  by transfer simp
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    62
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    63
lemma powhr_less_cancel: "\<And>a b x. x powhr a < x powhr b \<Longrightarrow> 1 < x \<Longrightarrow> a < b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    64
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    65
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    66
lemma powhr_less_cancel_iff [simp]: "1 < x \<Longrightarrow> x powhr a < x powhr b \<longleftrightarrow> a < b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    67
  by (blast intro: powhr_less_cancel powhr_less_mono)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    68
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    69
lemma powhr_le_cancel_iff [simp]: "1 < x \<Longrightarrow> x powhr a \<le> x powhr b \<longleftrightarrow> a \<le> b"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    70
  by (simp add: linorder_not_less [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    71
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    72
lemma hlog: "hlog (star_n X) (star_n Y) = star_n (\<lambda>n. log (X n) (Y n))"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    73
  by (simp add: hlog_def starfun2_star_n)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    74
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    75
lemma hlog_starfun_ln: "\<And>x. ( *f* ln) x = hlog (( *f* exp) 1) x"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    76
  by transfer (rule log_ln)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    77
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    78
lemma powhr_hlog_cancel [simp]: "\<And>a x. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> a powhr (hlog a x) = x"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    79
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    80
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    81
lemma hlog_powhr_cancel [simp]: "\<And>a y. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> hlog a (a powhr y) = y"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    82
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    83
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    84
lemma hlog_mult:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    85
  "\<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"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    86
  by transfer (rule log_mult)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    87
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    88
lemma hlog_as_starfun: "\<And>a x. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> hlog a x = ( *f* ln) x / ( *f* ln) a"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    89
  by transfer (simp add: log_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    90
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    91
lemma hlog_eq_div_starfun_ln_mult_hlog:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    92
  "\<And>a b x. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < b \<Longrightarrow> b \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow>
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    93
    hlog a x = (( *f* ln) b / ( *f* ln) a) * hlog b x"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    94
  by transfer (rule log_eq_div_ln_mult_log)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    95
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    96
lemma powhr_as_starfun: "\<And>a x. x powhr a = (if x = 0 then 0 else ( *f* exp) (a * ( *f* real_ln) x))"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
    97
  by transfer (simp add: powr_def)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    98
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
    99
lemma HInfinite_powhr:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   100
  "x \<in> HInfinite \<Longrightarrow> 0 < x \<Longrightarrow> a \<in> HFinite - Infinitesimal \<Longrightarrow> 0 < a \<Longrightarrow> x powhr a \<in> HInfinite"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   101
  by (auto intro!: starfun_ln_ge_zero starfun_ln_HInfinite
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   102
        HInfinite_HFinite_not_Infinitesimal_mult2 starfun_exp_HInfinite
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   103
      simp add: order_less_imp_le HInfinite_gt_zero_gt_one powhr_as_starfun zero_le_mult_iff)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   104
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   105
lemma hlog_hrabs_HInfinite_Infinitesimal:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   106
  "x \<in> HFinite - Infinitesimal \<Longrightarrow> a \<in> HInfinite \<Longrightarrow> 0 < a \<Longrightarrow> hlog a \<bar>x\<bar> \<in> Infinitesimal"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   107
  apply (frule HInfinite_gt_zero_gt_one)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   108
   apply (auto intro!: starfun_ln_HFinite_not_Infinitesimal
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   109
      HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   110
      simp add: starfun_ln_HInfinite not_Infinitesimal_not_zero
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   111
      hlog_as_starfun divide_inverse)
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   112
  done
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   113
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   114
lemma hlog_HInfinite_as_starfun: "a \<in> HInfinite \<Longrightarrow> 0 < a \<Longrightarrow> hlog a x = ( *f* ln) x / ( *f* ln) a"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   115
  by (rule hlog_as_starfun) auto
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   116
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   117
lemma hlog_one [simp]: "\<And>a. hlog a 1 = 0"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   118
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   119
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   120
lemma hlog_eq_one [simp]: "\<And>a. 0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> hlog a a = 1"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   121
  by transfer (rule log_eq_one)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   122
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   123
lemma hlog_inverse: "0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> 0 < x \<Longrightarrow> hlog a (inverse x) = - hlog a x"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   124
  by (rule add_left_cancel [of "hlog a x", THEN iffD1]) (simp add: hlog_mult [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   125
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   126
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"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   127
  by (simp add: hlog_mult hlog_inverse divide_inverse)
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   128
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   129
lemma hlog_less_cancel_iff [simp]:
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   130
  "\<And>a x y. 1 < a \<Longrightarrow> 0 < x \<Longrightarrow> 0 < y \<Longrightarrow> hlog a x < hlog a y \<longleftrightarrow> x < y"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   131
  by transfer simp
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   132
64604
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   133
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"
2bf8cfc98c4d misc tuning and modernization;
wenzelm
parents: 62479
diff changeset
   134
  by (simp add: linorder_not_less [symmetric])
27468
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   135
0783dd1dc13d move nonstandard analysis theories to NSA directory
huffman
parents:
diff changeset
   136
end