src/HOL/Matrix/ComputeFloat.thy
author wenzelm
Fri, 05 Feb 2010 20:19:40 +0100
changeset 35003 e0d01e77c7b1
parent 32491 d5d8bea0cd94
child 35032 7efe662e41b4
permissions -rw-r--r--
eliminated self intersection and non-integer coordinates;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
     1
(*  Title:  HOL/Tools/ComputeFloat.thy
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
     2
    Author: Steven Obua
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
     3
*)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
     4
20717
2244b0d719a0 add header
huffman
parents: 20485
diff changeset
     5
header {* Floating Point Representation of the Reals *}
2244b0d719a0 add header
huffman
parents: 20485
diff changeset
     6
29804
e15b74577368 Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents: 29667
diff changeset
     7
theory ComputeFloat
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 27366
diff changeset
     8
imports Complex_Main
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 27366
diff changeset
     9
uses "~~/src/Tools/float.ML" ("~~/src/HOL/Tools/float_arith.ML")
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
    10
begin
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    11
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    12
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
    13
  pow2 :: "int \<Rightarrow> real" where
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    14
  "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
    15
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
    16
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
    17
  float :: "int * int \<Rightarrow> real" where
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    18
  "float x = real (fst x) * pow2 (snd x)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    19
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    20
lemma pow2_0[simp]: "pow2 0 = 1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    21
by (simp add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    22
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    23
lemma pow2_1[simp]: "pow2 1 = 2"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    24
by (simp add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    25
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    26
lemma pow2_neg: "pow2 x = inverse (pow2 (-x))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    27
by (simp add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    28
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    29
lemma pow2_add1: "pow2 (1 + a) = 2 * (pow2 a)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    30
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    31
  have h: "! n. nat (2 + int n) - Suc 0 = nat (1 + int n)" by arith
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    32
  have g: "! a b. a - -1 = a + (1::int)" by arith
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    33
  have pos: "! n. pow2 (int n + 1) = 2 * pow2 (int n)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    34
    apply (auto, induct_tac n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    35
    apply (simp_all add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    36
    apply (rule_tac m1="2" and n1="nat (2 + int na)" in ssubst[OF realpow_num_eq_if])
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23365
diff changeset
    37
    by (auto simp add: h)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    38
  show ?thesis
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    39
  proof (induct a)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    40
    case (1 n)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29040
diff changeset
    41
    from pos show ?case by (simp add: algebra_simps)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    42
  next
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    43
    case (2 n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    44
    show ?case
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    45
      apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    46
      apply (subst pow2_neg[of "- int n"])
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23365
diff changeset
    47
      apply (subst pow2_neg[of "-1 - int n"])
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    48
      apply (auto simp add: g pos)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    49
      done
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    50
  qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    51
qed
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    52
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    53
lemma pow2_add: "pow2 (a+b) = (pow2 a) * (pow2 b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    54
proof (induct b)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    55
  case (1 n)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    56
  show ?case
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    57
  proof (induct n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    58
    case 0
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    59
    show ?case by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    60
  next
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    61
    case (Suc m)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29040
diff changeset
    62
    show ?case by (auto simp add: algebra_simps pow2_add1 prems)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    63
  qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    64
next
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    65
  case (2 n)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    66
  show ?case
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    67
  proof (induct n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    68
    case 0
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    69
    show ?case
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    70
      apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    71
      apply (subst pow2_neg[of "a + -1"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    72
      apply (subst pow2_neg[of "-1"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    73
      apply (simp)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    74
      apply (insert pow2_add1[of "-a"])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29040
diff changeset
    75
      apply (simp add: algebra_simps)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    76
      apply (subst pow2_neg[of "-a"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    77
      apply (simp)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    78
      done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    79
    case (Suc m)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    80
    have a: "int m - (a + -2) =  1 + (int m - a + 1)" by arith
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    81
    have b: "int m - -2 = 1 + (int m + 1)" by arith
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    82
    show ?case
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    83
      apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    84
      apply (subst pow2_neg[of "a + (-2 - int m)"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    85
      apply (subst pow2_neg[of "-2 - int m"])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29040
diff changeset
    86
      apply (auto simp add: algebra_simps)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    87
      apply (subst a)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    88
      apply (subst b)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    89
      apply (simp only: pow2_add1)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    90
      apply (subst pow2_neg[of "int m - a + 1"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    91
      apply (subst pow2_neg[of "int m + 1"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    92
      apply auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    93
      apply (insert prems)
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29040
diff changeset
    94
      apply (auto simp add: algebra_simps)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    95
      done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    96
  qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    97
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    98
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    99
lemma "float (a, e) + float (b, e) = float (a + b, e)"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29040
diff changeset
   100
by (simp add: float_def algebra_simps)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   101
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   102
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
   103
  int_of_real :: "real \<Rightarrow> int" where
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   104
  "int_of_real x = (SOME y. real y = x)"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
   105
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
   106
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21256
diff changeset
   107
  real_is_int :: "real \<Rightarrow> bool" where
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   108
  "real_is_int x = (EX (u::int). x = real u)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   109
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   110
lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   111
by (auto simp add: real_is_int_def int_of_real_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   112
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   113
lemma float_transfer: "real_is_int ((real a)*(pow2 c)) \<Longrightarrow> float (a, b) = float (int_of_real ((real a)*(pow2 c)), b - c)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   114
by (simp add: float_def real_is_int_def2 pow2_add[symmetric])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   115
26313
8590bf5ef343 avoid rebinding of existing facts;
wenzelm
parents: 26086
diff changeset
   116
lemma pow2_int: "pow2 (int c) = 2^c"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   117
by (simp add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   118
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   119
lemma float_transfer_nat: "float (a, b) = float (a * 2^c, b - int c)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   120
by (simp add: float_def pow2_int[symmetric] pow2_add[symmetric])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   121
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   122
lemma real_is_int_real[simp]: "real_is_int (real (x::int))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   123
by (auto simp add: real_is_int_def int_of_real_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   124
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   125
lemma int_of_real_real[simp]: "int_of_real (real x) = x"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   126
by (simp add: int_of_real_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   127
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   128
lemma real_int_of_real[simp]: "real_is_int x \<Longrightarrow> real (int_of_real x) = x"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   129
by (auto simp add: int_of_real_def real_is_int_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   130
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   131
lemma real_is_int_add_int_of_real: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a+b)) = (int_of_real a) + (int_of_real b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   132
by (auto simp add: int_of_real_def real_is_int_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   133
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   134
lemma real_is_int_add[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a+b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   135
apply (subst real_is_int_def2)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   136
apply (simp add: real_is_int_add_int_of_real real_int_of_real)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   137
done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   138
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   139
lemma int_of_real_sub: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> (int_of_real (a-b)) = (int_of_real a) - (int_of_real b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   140
by (auto simp add: int_of_real_def real_is_int_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   141
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   142
lemma real_is_int_sub[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a-b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   143
apply (subst real_is_int_def2)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   144
apply (simp add: int_of_real_sub real_int_of_real)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   145
done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   146
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   147
lemma real_is_int_rep: "real_is_int x \<Longrightarrow> ?! (a::int). real a = x"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   148
by (auto simp add: real_is_int_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   149
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   150
lemma int_of_real_mult:
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   151
  assumes "real_is_int a" "real_is_int b"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   152
  shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   153
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   154
  from prems have a: "?! (a'::int). real a' = a" by (rule_tac real_is_int_rep, auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   155
  from prems have b: "?! (b'::int). real b' = b" by (rule_tac real_is_int_rep, auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   156
  from a obtain a'::int where a':"a = real a'" by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   157
  from b obtain b'::int where b':"b = real b'" by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   158
  have r: "real a' * real b' = real (a' * b')" by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   159
  show ?thesis
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   160
    apply (simp add: a' b')
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   161
    apply (subst r)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   162
    apply (simp only: int_of_real_real)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   163
    done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   164
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   165
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   166
lemma real_is_int_mult[simp]: "real_is_int a \<Longrightarrow> real_is_int b \<Longrightarrow> real_is_int (a*b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   167
apply (subst real_is_int_def2)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   168
apply (simp add: int_of_real_mult)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   169
done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   170
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   171
lemma real_is_int_0[simp]: "real_is_int (0::real)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   172
by (simp add: real_is_int_def int_of_real_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   173
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   174
lemma real_is_int_1[simp]: "real_is_int (1::real)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   175
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   176
  have "real_is_int (1::real) = real_is_int(real (1::int))" by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   177
  also have "\<dots> = True" by (simp only: real_is_int_real)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   178
  ultimately show ?thesis by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   179
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   180
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   181
lemma real_is_int_n1: "real_is_int (-1::real)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   182
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   183
  have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   184
  also have "\<dots> = True" by (simp only: real_is_int_real)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   185
  ultimately show ?thesis by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   186
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   187
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   188
lemma real_is_int_number_of[simp]: "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   189
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   190
  have neg1: "real_is_int (-1::real)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   191
  proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   192
    have "real_is_int (-1::real) = real_is_int(real (-1::int))" by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   193
    also have "\<dots> = True" by (simp only: real_is_int_real)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   194
    ultimately show ?thesis by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   195
  qed
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   196
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   197
  {
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   198
    fix x :: int
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   199
    have "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   200
      unfolding number_of_eq
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   201
      apply (induct x)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   202
      apply (induct_tac n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   203
      apply (simp)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   204
      apply (simp)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   205
      apply (induct_tac n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   206
      apply (simp add: neg1)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   207
    proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   208
      fix n :: nat
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   209
      assume rn: "(real_is_int (of_int (- (int (Suc n)))))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   210
      have s: "-(int (Suc (Suc n))) = -1 + - (int (Suc n))" by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   211
      show "real_is_int (of_int (- (int (Suc (Suc n)))))"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   212
        apply (simp only: s of_int_add)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   213
        apply (rule real_is_int_add)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   214
        apply (simp add: neg1)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   215
        apply (simp only: rn)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   216
        done
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   217
    qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   218
  }
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   219
  note Abs_Bin = this
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   220
  {
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   221
    fix x :: int
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   222
    have "? u. x = u"
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   223
      apply (rule exI[where x = "x"])
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   224
      apply (simp)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   225
      done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   226
  }
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   227
  then obtain u::int where "x = u" by auto
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   228
  with Abs_Bin show ?thesis by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   229
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   230
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   231
lemma int_of_real_0[simp]: "int_of_real (0::real) = (0::int)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   232
by (simp add: int_of_real_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   233
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   234
lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   235
proof -
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   236
  have 1: "(1::real) = real (1::int)" by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   237
  show ?thesis by (simp only: 1 int_of_real_real)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   238
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   239
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   240
lemma int_of_real_number_of[simp]: "int_of_real (number_of b) = number_of b"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   241
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   242
  have "real_is_int (number_of b)" by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   243
  then have uu: "?! u::int. number_of b = real u" by (auto simp add: real_is_int_rep)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   244
  then obtain u::int where u:"number_of b = real u" by auto
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   245
  have "number_of b = real ((number_of b)::int)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   246
    by (simp add: number_of_eq real_of_int_def)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   247
  have ub: "number_of b = real ((number_of b)::int)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   248
    by (simp add: number_of_eq real_of_int_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   249
  from uu u ub have unb: "u = number_of b"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   250
    by blast
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   251
  have "int_of_real (number_of b) = u" by (simp add: u)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   252
  with unb show ?thesis by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   253
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   254
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   255
lemma float_transfer_even: "even a \<Longrightarrow> float (a, b) = float (a div 2, b+1)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   256
  apply (subst float_transfer[where a="a" and b="b" and c="-1", simplified])
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29040
diff changeset
   257
  apply (simp_all add: pow2_def even_def real_is_int_def algebra_simps)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   258
  apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   259
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   260
  fix q::int
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   261
  have a:"b - (-1\<Colon>int) = (1\<Colon>int) + b" by arith
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   262
  show "(float (q, (b - (-1\<Colon>int)))) = (float (q, ((1\<Colon>int) + b)))"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   263
    by (simp add: a)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   264
qed
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   265
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   266
lemma int_div_zdiv: "int (a div b) = (int a) div (int b)"
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23365
diff changeset
   267
by (rule zdiv_int)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   268
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   269
lemma int_mod_zmod: "int (a mod b) = (int a) mod (int b)"
23431
25ca91279a9b change simp rules for of_nat to work like int did previously (reorient of_nat_Suc, remove of_nat_mult [simp]); preserve original variable names in legacy int theorems
huffman
parents: 23365
diff changeset
   270
by (rule zmod_int)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   271
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   272
lemma abs_div_2_less: "a \<noteq> 0 \<Longrightarrow> a \<noteq> -1 \<Longrightarrow> abs((a::int) div 2) < abs a"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   273
by arith
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   274
27366
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   275
function norm_float :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   276
  "norm_float a b = (if a \<noteq> 0 \<and> even a then norm_float (a div 2) (b + 1)
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   277
    else if a = 0 then (0, 0) else (a, b))"
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   278
by auto
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   279
27366
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   280
termination by (relation "measure (nat o abs o fst)")
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   281
  (auto intro: abs_div_2_less)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   282
27366
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   283
lemma norm_float: "float x = float (split norm_float x)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   284
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   285
  {
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   286
    fix a b :: int
27366
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   287
    have norm_float_pair: "float (a, b) = float (norm_float a b)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   288
    proof (induct a b rule: norm_float.induct)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   289
      case (1 u v)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   290
      show ?case
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   291
      proof cases
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   292
        assume u: "u \<noteq> 0 \<and> even u"
27366
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   293
        with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   294
        with u have "float (u,v) = float (u div 2, v+1)" by (simp add: float_transfer_even)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   295
        then show ?thesis
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   296
          apply (subst norm_float.simps)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   297
          apply (simp add: ind)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   298
          done
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   299
      next
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   300
        assume "~(u \<noteq> 0 \<and> even u)"
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   301
        then show ?thesis
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   302
          by (simp add: prems float_def)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   303
      qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   304
    qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   305
  }
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   306
  note helper = this
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   307
  have "? a b. x = (a,b)" by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   308
  then obtain a b where "x = (a, b)" by blast
27366
d0cda1ea705e dropped recdef
haftmann
parents: 26313
diff changeset
   309
  then show ?thesis by (simp add: helper)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   310
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   311
24301
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   312
lemma float_add_l0: "float (0, e) + x = x"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   313
  by (simp add: float_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   314
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   315
lemma float_add_r0: "x + float (0, e) = x"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   316
  by (simp add: float_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   317
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   318
lemma float_add:
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   319
  "float (a1, e1) + float (a2, e2) =
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   320
  (if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   321
  else float (a1*2^(nat (e1-e2))+a2, e2))"
29667
53103fc8ffa3 Replaced group_ and ring_simps by algebra_simps;
nipkow
parents: 29040
diff changeset
   322
  apply (simp add: float_def algebra_simps)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   323
  apply (auto simp add: pow2_int[symmetric] pow2_add[symmetric])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   324
  done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   325
24301
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   326
lemma float_add_assoc1:
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   327
  "(x + float (y1, e1)) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   328
  by simp
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   329
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   330
lemma float_add_assoc2:
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   331
  "(float (y1, e1) + x) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   332
  by simp
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   333
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   334
lemma float_add_assoc3:
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   335
  "float (y1, e1) + (x + float (y2, e2)) = (float (y1, e1) + float (y2, e2)) + x"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   336
  by simp
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   337
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   338
lemma float_add_assoc4:
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   339
  "float (y1, e1) + (float (y2, e2) + x) = (float (y1, e1) + float (y2, e2)) + x"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   340
  by simp
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   341
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   342
lemma float_mult_l0: "float (0, e) * x = float (0, 0)"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   343
  by (simp add: float_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   344
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   345
lemma float_mult_r0: "x * float (0, e) = float (0, 0)"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   346
  by (simp add: float_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   347
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   348
definition 
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   349
  lbound :: "real \<Rightarrow> real"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   350
where
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   351
  "lbound x = min 0 x"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   352
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   353
definition
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   354
  ubound :: "real \<Rightarrow> real"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   355
where
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   356
  "ubound x = max 0 x"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   357
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   358
lemma lbound: "lbound x \<le> x"   
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   359
  by (simp add: lbound_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   360
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   361
lemma ubound: "x \<le> ubound x"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   362
  by (simp add: ubound_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   363
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   364
lemma float_mult:
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   365
  "float (a1, e1) * float (a2, e2) =
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   366
  (float (a1 * a2, e1 + e2))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   367
  by (simp add: float_def pow2_add)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   368
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   369
lemma float_minus:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   370
  "- (float (a,b)) = float (-a, b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   371
  by (simp add: float_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   372
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   373
lemma zero_less_pow2:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   374
  "0 < pow2 x"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   375
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   376
  {
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   377
    fix y
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   378
    have "0 <= y \<Longrightarrow> 0 < pow2 y"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   379
      by (induct y, induct_tac n, simp_all add: pow2_add)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   380
  }
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   381
  note helper=this
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   382
  show ?thesis
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   383
    apply (case_tac "0 <= x")
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   384
    apply (simp add: helper)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   385
    apply (subst pow2_neg)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   386
    apply (simp add: helper)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   387
    done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   388
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   389
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   390
lemma zero_le_float:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   391
  "(0 <= float (a,b)) = (0 <= a)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   392
  apply (auto simp add: float_def)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   393
  apply (auto simp add: zero_le_mult_iff zero_less_pow2)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   394
  apply (insert zero_less_pow2[of b])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   395
  apply (simp_all)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   396
  done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   397
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   398
lemma float_le_zero:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   399
  "(float (a,b) <= 0) = (a <= 0)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   400
  apply (auto simp add: float_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   401
  apply (auto simp add: mult_le_0_iff)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   402
  apply (insert zero_less_pow2[of b])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   403
  apply auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   404
  done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   405
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   406
lemma float_abs:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   407
  "abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   408
  apply (auto simp add: abs_if)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   409
  apply (simp_all add: zero_le_float[symmetric, of a b] float_minus)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   410
  done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   411
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   412
lemma float_zero:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   413
  "float (0, b) = 0"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   414
  by (simp add: float_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   415
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   416
lemma float_pprt:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   417
  "pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   418
  by (auto simp add: zero_le_float float_le_zero float_zero)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   419
24301
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   420
lemma pprt_lbound: "pprt (lbound x) = float (0, 0)"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   421
  apply (simp add: float_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   422
  apply (rule pprt_eq_0)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   423
  apply (simp add: lbound_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   424
  done
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   425
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   426
lemma nprt_ubound: "nprt (ubound x) = float (0, 0)"
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   427
  apply (simp add: float_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   428
  apply (rule nprt_eq_0)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   429
  apply (simp add: ubound_def)
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   430
  done
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   431
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   432
lemma float_nprt:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   433
  "nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   434
  by (auto simp add: zero_le_float float_le_zero float_zero)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   435
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   436
lemma norm_0_1: "(0::_::number_ring) = Numeral0 & (1::_::number_ring) = Numeral1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   437
  by auto
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   438
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   439
lemma add_left_zero: "0 + a = (a::'a::comm_monoid_add)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   440
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   441
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   442
lemma add_right_zero: "a + 0 = (a::'a::comm_monoid_add)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   443
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   444
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   445
lemma mult_left_one: "1 * a = (a::'a::semiring_1)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   446
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   447
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   448
lemma mult_right_one: "a * 1 = (a::'a::semiring_1)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   449
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   450
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   451
lemma int_pow_0: "(a::int)^(Numeral0) = 1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   452
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   453
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   454
lemma int_pow_1: "(a::int)^(Numeral1) = a"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   455
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   456
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   457
lemma zero_eq_Numeral0_nring: "(0::'a::number_ring) = Numeral0"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   458
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   459
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   460
lemma one_eq_Numeral1_nring: "(1::'a::number_ring) = Numeral1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   461
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   462
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   463
lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   464
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   465
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   466
lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   467
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   468
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   469
lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   470
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   471
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   472
lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   473
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   474
  have 1:"((-1)::nat) = 0"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   475
    by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   476
  show ?thesis by (simp add: 1)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   477
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   478
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   479
lemma fst_cong: "a=a' \<Longrightarrow> fst (a,b) = fst (a',b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   480
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   481
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   482
lemma snd_cong: "b=b' \<Longrightarrow> snd (a,b) = snd (a,b')"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   483
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   484
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   485
lemma lift_bool: "x \<Longrightarrow> x=True"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   486
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   487
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   488
lemma nlift_bool: "~x \<Longrightarrow> x=False"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   489
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   490
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   491
lemma not_false_eq_true: "(~ False) = True" by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   492
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   493
lemma not_true_eq_false: "(~ True) = False" by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   494
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   495
lemmas binarith =
26076
b9c716a9fb5f added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
huffman
parents: 24653
diff changeset
   496
  normalize_bin_simps
b9c716a9fb5f added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
huffman
parents: 24653
diff changeset
   497
  pred_bin_simps succ_bin_simps
b9c716a9fb5f added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
huffman
parents: 24653
diff changeset
   498
  add_bin_simps minus_bin_simps mult_bin_simps
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   499
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   500
lemma int_eq_number_of_eq:
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   501
  "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)"
28967
3bdb1eae352c enable eq_bin_simps for simplifying equalities on numerals
huffman
parents: 28963
diff changeset
   502
  by (rule eq_number_of_eq)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   503
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   504
lemma int_iszero_number_of_Pls: "iszero (Numeral0::int)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   505
  by (simp only: iszero_number_of_Pls)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   506
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   507
lemma int_nonzero_number_of_Min: "~(iszero ((-1)::int))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   508
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   509
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26076
diff changeset
   510
lemma int_iszero_number_of_Bit0: "iszero ((number_of (Int.Bit0 w))::int) = iszero ((number_of w)::int)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   511
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   512
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26076
diff changeset
   513
lemma int_iszero_number_of_Bit1: "\<not> iszero ((number_of (Int.Bit1 w))::int)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   514
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   515
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   516
lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)"
29040
286c669d3a7a move all neg-related lemmas to NatBin; make type of neg specific to int
huffman
parents: 28967
diff changeset
   517
  unfolding neg_def number_of_is_id by simp
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   518
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   519
lemma int_not_neg_number_of_Pls: "\<not> (neg (Numeral0::int))"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   520
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   521
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   522
lemma int_neg_number_of_Min: "neg (-1::int)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   523
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   524
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26076
diff changeset
   525
lemma int_neg_number_of_Bit0: "neg ((number_of (Int.Bit0 w))::int) = neg ((number_of w)::int)"
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26076
diff changeset
   526
  by simp
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26076
diff changeset
   527
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26076
diff changeset
   528
lemma int_neg_number_of_Bit1: "neg ((number_of (Int.Bit1 w))::int) = neg ((number_of w)::int)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   529
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   530
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   531
lemma int_le_number_of_eq: "(((number_of x)::int) \<le> number_of y) = (\<not> neg ((number_of (y + (uminus x)))::int))"
28963
f6d9e0e0b153 fix proofs related to simplification of inequalities on numerals
huffman
parents: 28952
diff changeset
   532
  unfolding neg_def number_of_is_id by (simp add: not_less)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   533
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   534
lemmas intarithrel =
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   535
  int_eq_number_of_eq
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26076
diff changeset
   536
  lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_Bit0
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26076
diff changeset
   537
  lift_bool[OF int_iszero_number_of_Bit1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min]
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 26076
diff changeset
   538
  int_neg_number_of_Bit0 int_neg_number_of_Bit1 int_le_number_of_eq
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   539
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   540
lemma int_number_of_add_sym: "((number_of v)::int) + number_of w = number_of (v + w)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   541
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   542
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   543
lemma int_number_of_diff_sym: "((number_of v)::int) - number_of w = number_of (v + (uminus w))"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   544
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   545
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   546
lemma int_number_of_mult_sym: "((number_of v)::int) * number_of w = number_of (v * w)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   547
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   548
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   549
lemma int_number_of_minus_sym: "- ((number_of v)::int) = number_of (uminus v)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   550
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   551
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   552
lemmas intarith = int_number_of_add_sym int_number_of_minus_sym int_number_of_diff_sym int_number_of_mult_sym
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   553
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   554
lemmas natarith = add_nat_number_of diff_nat_number_of mult_nat_number_of eq_nat_number_of less_nat_number_of
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   555
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   556
lemmas powerarith = nat_number_of zpower_number_of_even
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   557
  zpower_number_of_odd[simplified zero_eq_Numeral0_nring one_eq_Numeral1_nring]
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   558
  zpower_Pls zpower_Min
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   559
24301
6c7f226b24c3 changed floatarith lemmas
obua
parents: 24124
diff changeset
   560
lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0 
24653
3d3ebc0c927c changed lemmas
obua
parents: 24301
diff changeset
   561
          float_minus float_abs zero_le_float float_pprt float_nprt pprt_lbound nprt_ubound
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   562
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   563
(* for use with the compute oracle *)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   564
lemmas arith = binarith intarith intarithrel natarith powerarith floatarith not_false_eq_true not_true_eq_false
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   565
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 27366
diff changeset
   566
use "~~/src/HOL/Tools/float_arith.ML"
20771
89bec28a03c8 proper use of float.ML;
wenzelm
parents: 20717
diff changeset
   567
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   568
end