src/Tools/isac/BaseDefinitions/BaseDefinitions.thy
author wenzelm
Thu, 16 Sep 2021 17:23:54 +0200
changeset 60405 d4ebe139100d
parent 60360 49680d595342
child 60409 432482c14d96
permissions -rw-r--r--
separate realpow constant, with additional cases not covered by Transcendental.powr;
walther@59866
     1
(* Title:  "BaseDefinitions/BaseDefinitions.thy"
wneuper@59586
     2
   Author: Walther Neuper 190823
wneuper@59586
     3
   (c) due to copyright terms
walther@59866
     4
walther@59887
     5
Know_Store holds Theory_Data (problems, methods, etc) and requires respective definitions.
wneuper@59586
     6
*)
walther@59887
     7
theory BaseDefinitions imports Know_Store
wneuper@59586
     8
begin
wenzelm@60405
     9
wenzelm@60405
    10
subsection \<open>Power re-defined for a specific type\<close>
wenzelm@60405
    11
wenzelm@60405
    12
definition even_real :: "real \<Rightarrow> bool"
wenzelm@60405
    13
  where "even_real x \<longleftrightarrow> \<bar>x\<bar> \<in> \<nat> \<and> even (inv real \<bar>x\<bar>)"
wenzelm@60405
    14
definition odd_real :: "real \<Rightarrow> bool"
wenzelm@60405
    15
  where "odd_real x \<longleftrightarrow> \<bar>x\<bar> \<in> \<nat> \<and> odd (inv real \<bar>x\<bar>)"
wenzelm@60405
    16
wenzelm@60405
    17
definition realpow :: "real \<Rightarrow> real \<Rightarrow> real"  (infixr "\<up>" 80)
wenzelm@60405
    18
  where "x \<up> n =
wenzelm@60405
    19
    (if x \<ge> 0 then x powr n
wenzelm@60405
    20
     else if n \<ge> 0 \<and> even_real n then (- x) powr n
wenzelm@60405
    21
     else if n \<ge> 0 \<and> odd_real n then - ((- x) powr n)
wenzelm@60405
    22
     else if n < 0 \<and> even_real n then 1 / (- x) powr (- n)
wenzelm@60405
    23
     else if n < 0 \<and> odd_real n then - 1 / ((- x) powr (- n))
wenzelm@60405
    24
     else undefined)"
wenzelm@60405
    25
wenzelm@60405
    26
lemma realpow_simps [simp]:
wenzelm@60405
    27
  "0 \<up> a = 0"
wenzelm@60405
    28
  "1 \<up> a = 1"
wenzelm@60405
    29
  "numeral m \<up> numeral n = (numeral m ^ numeral n :: real)"
wenzelm@60405
    30
  "x > 0 \<Longrightarrow> x \<up> 0 = 1"
wenzelm@60405
    31
  "x \<ge> 0 \<Longrightarrow> x \<up> 1 = x"
wenzelm@60405
    32
  "x \<ge> 0 \<Longrightarrow> x \<up> numeral n = x ^ numeral n"
wenzelm@60405
    33
  "x > 0 \<Longrightarrow> x \<up> - numeral n = 1 / (x ^ numeral n)"
wenzelm@60405
    34
  by (simp_all add: realpow_def powr_neg_numeral)
wenzelm@60405
    35
wenzelm@60405
    36
lemma inv_real_eq:
wenzelm@60405
    37
  "inv real 0 = 0"
wenzelm@60405
    38
  "inv real 1 = 1"
wenzelm@60405
    39
  "inv real (numeral n) = numeral n"
wenzelm@60405
    40
  by (simp_all add: inv_f_eq)
wenzelm@60405
    41
wenzelm@60405
    42
lemma realpow_uminus_simps [simp]:
wenzelm@60405
    43
  "(- numeral m) \<up> 0 = 1"
wenzelm@60405
    44
  "(- numeral m) \<up> 1 = - numeral m"
wenzelm@60405
    45
  "even (numeral n :: nat) \<Longrightarrow> (- numeral m) \<up> (numeral n) = numeral m ^ numeral n"
wenzelm@60405
    46
  "odd (numeral n :: nat) \<Longrightarrow> (- numeral m) \<up> (numeral n) = - (numeral m ^ numeral n)"
wenzelm@60405
    47
  "even (numeral n :: nat) \<Longrightarrow> (- numeral m) \<up> (- numeral n) = 1 / (numeral m ^ numeral n)"
wenzelm@60405
    48
  "odd (numeral n :: nat) \<Longrightarrow> (- numeral m) \<up> (- numeral n) = - 1 / (numeral m ^ numeral n)"
wenzelm@60405
    49
  by (simp_all add: realpow_def even_real_def odd_real_def inv_real_eq)
wenzelm@60405
    50
wenzelm@60405
    51
walther@59659
    52
  ML_file termC.sml
walther@59911
    53
  ML_file substitution.sml
walther@59659
    54
  ML_file contextC.sml
walther@59659
    55
  ML_file environment.sml
walther@59630
    56
wneuper@59586
    57
ML \<open>
wneuper@59586
    58
\<close> ML \<open>
walther@60317
    59
\<close> ML \<open>
wneuper@59586
    60
\<close>
wneuper@59586
    61
end