src/HOL/Real/Float.thy
author huffman
Tue, 26 Sep 2006 22:37:51 +0200
changeset 20717 2244b0d719a0
parent 20485 3078fd2eec7b
child 20771 89bec28a03c8
permissions -rw-r--r--
add header
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
     1
(*  Title: HOL/Real/Float.thy
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
     2
    ID:    $Id$
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
     3
    Author: Steven Obua
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
     4
*)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
     5
20717
2244b0d719a0 add header
huffman
parents: 20485
diff changeset
     6
header {* Floating Point Representation of the Reals *}
2244b0d719a0 add header
huffman
parents: 20485
diff changeset
     7
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
     8
theory Float
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
     9
imports Real
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
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    13
  pow2 :: "int \<Rightarrow> real"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    14
  "pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    15
  float :: "int * int \<Rightarrow> real"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    16
  "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
    17
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    18
lemma pow2_0[simp]: "pow2 0 = 1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    19
by (simp add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    20
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    21
lemma pow2_1[simp]: "pow2 1 = 2"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    22
by (simp add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    23
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    24
lemma pow2_neg: "pow2 x = inverse (pow2 (-x))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    25
by (simp add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    26
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    27
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
    28
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    29
  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
    30
  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
    31
  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
    32
    apply (auto, induct_tac n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    33
    apply (simp_all add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    34
    apply (rule_tac m1="2" and n1="nat (2 + int na)" in ssubst[OF realpow_num_eq_if])
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19765
diff changeset
    35
    by (auto simp add: h)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    36
  show ?thesis
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    37
  proof (induct a)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    38
    case (1 n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    39
    from pos show ?case by (simp add: ring_eq_simps)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    40
  next
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    41
    case (2 n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    42
    show ?case
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    43
      apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    44
      apply (subst pow2_neg[of "- int n"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    45
      apply (subst pow2_neg[of "-1 - int n"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    46
      apply (auto simp add: g pos)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    47
      done
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    48
  qed
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    49
qed
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    50
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    51
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
    52
proof (induct b)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    53
  case (1 n)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    54
  show ?case
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    55
  proof (induct n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    56
    case 0
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    57
    show ?case by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    58
  next
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    59
    case (Suc m)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    60
    show ?case by (auto simp add: ring_eq_simps pow2_add1 prems)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    61
  qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    62
next
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    63
  case (2 n)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    64
  show ?case
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    65
  proof (induct n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    66
    case 0
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    67
    show ?case
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    68
      apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    69
      apply (subst pow2_neg[of "a + -1"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    70
      apply (subst pow2_neg[of "-1"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    71
      apply (simp)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    72
      apply (insert pow2_add1[of "-a"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    73
      apply (simp add: ring_eq_simps)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    74
      apply (subst pow2_neg[of "-a"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    75
      apply (simp)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    76
      done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    77
    case (Suc m)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    78
    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
    79
    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
    80
    show ?case
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    81
      apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    82
      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
    83
      apply (subst pow2_neg[of "-2 - int m"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    84
      apply (auto simp add: ring_eq_simps)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    85
      apply (subst a)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    86
      apply (subst b)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    87
      apply (simp only: pow2_add1)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    88
      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
    89
      apply (subst pow2_neg[of "int m + 1"])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    90
      apply auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    91
      apply (insert prems)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    92
      apply (auto simp add: ring_eq_simps)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    93
      done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    94
  qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    95
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    96
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
    97
lemma "float (a, e) + float (b, e) = float (a + b, e)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    98
by (simp add: float_def ring_eq_simps)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
    99
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   100
definition
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   101
  int_of_real :: "real \<Rightarrow> int"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   102
  "int_of_real x = (SOME y. real y = x)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   103
  real_is_int :: "real \<Rightarrow> bool"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   104
  "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
   105
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   106
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
   107
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
   108
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   109
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
   110
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
   111
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   112
lemma pow2_int: "pow2 (int c) = (2::real)^c"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   113
by (simp add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   114
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   115
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
   116
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
   117
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   118
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
   119
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
   120
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   121
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
   122
by (simp add: int_of_real_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   123
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   124
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
   125
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
   126
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   127
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
   128
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
   129
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   130
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
   131
apply (subst real_is_int_def2)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   132
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
   133
done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   134
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   135
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
   136
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
   137
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   138
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
   139
apply (subst real_is_int_def2)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   140
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
   141
done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   142
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   143
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
   144
by (auto simp add: real_is_int_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   145
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   146
lemma int_of_real_mult:
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   147
  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
   148
  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
   149
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   150
  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
   151
  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
   152
  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
   153
  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
   154
  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
   155
  show ?thesis
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   156
    apply (simp add: a' b')
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   157
    apply (subst r)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   158
    apply (simp only: int_of_real_real)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   159
    done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   160
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   161
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   162
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
   163
apply (subst real_is_int_def2)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   164
apply (simp add: int_of_real_mult)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   165
done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   166
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   167
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
   168
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
   169
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   170
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
   171
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   172
  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
   173
  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
   174
  ultimately show ?thesis by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   175
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   176
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   177
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
   178
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   179
  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
   180
  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
   181
  ultimately show ?thesis by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   182
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   183
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   184
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
   185
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   186
  have neg1: "real_is_int (-1::real)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   187
  proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   188
    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
   189
    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
   190
    ultimately show ?thesis by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   191
  qed
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   192
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   193
  {
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   194
    fix x :: int
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   195
    have "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   196
      unfolding number_of_eq
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   197
      apply (induct x)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   198
      apply (induct_tac n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   199
      apply (simp)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   200
      apply (simp)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   201
      apply (induct_tac n)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   202
      apply (simp add: neg1)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   203
    proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   204
      fix n :: nat
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   205
      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
   206
      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
   207
      show "real_is_int (of_int (- (int (Suc (Suc n)))))"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   208
        apply (simp only: s of_int_add)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   209
        apply (rule real_is_int_add)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   210
        apply (simp add: neg1)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   211
        apply (simp only: rn)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   212
        done
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   213
    qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   214
  }
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   215
  note Abs_Bin = this
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   216
  {
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   217
    fix x :: int
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   218
    have "? u. x = u"
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   219
      apply (rule exI[where x = "x"])
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   220
      apply (simp)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   221
      done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   222
  }
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   223
  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
   224
  with Abs_Bin show ?thesis by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   225
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   226
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   227
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
   228
by (simp add: int_of_real_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   229
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   230
lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)"
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   231
proof -
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   232
  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
   233
  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
   234
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   235
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   236
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
   237
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   238
  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
   239
  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
   240
  then obtain u::int where u:"number_of b = real u" by auto
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   241
  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
   242
    by (simp add: number_of_eq real_of_int_def)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   243
  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
   244
    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
   245
  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
   246
    by blast
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   247
  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
   248
  with unb show ?thesis by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   249
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   250
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   251
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
   252
  apply (subst float_transfer[where a="a" and b="b" and c="-1", simplified])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   253
  apply (simp_all add: pow2_def even_def real_is_int_def ring_eq_simps)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   254
  apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   255
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   256
  fix q::int
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   257
  have a:"b - (-1\<Colon>int) = (1\<Colon>int) + b" by arith
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   258
  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
   259
    by (simp add: a)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   260
qed
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   261
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   262
consts
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   263
  norm_float :: "int*int \<Rightarrow> int*int"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   264
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   265
lemma int_div_zdiv: "int (a div b) = (int a) div (int b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   266
apply (subst split_div, auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   267
apply (subst split_zdiv, auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   268
apply (rule_tac a="int (b * i) + int j" and b="int b" and r="int j" and r'=ja in IntDiv.unique_quotient)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   269
apply (auto simp add: IntDiv.quorem_def int_eq_of_nat)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   270
done
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 int_mod_zmod: "int (a mod b) = (int a) mod (int b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   273
apply (subst split_mod, auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   274
apply (subst split_zmod, auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   275
apply (rule_tac a="int (b * i) + int j" and b="int b" and q="int i" and q'=ia in IntDiv.unique_remainder)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   276
apply (auto simp add: IntDiv.quorem_def int_eq_of_nat)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   277
done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   278
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   279
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
   280
by arith
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   281
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   282
lemma terminating_norm_float: "\<forall>a. (a::int) \<noteq> 0 \<and> even a \<longrightarrow> a \<noteq> 0 \<and> \<bar>a div 2\<bar> < \<bar>a\<bar>"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   283
apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   284
apply (rule abs_div_2_less)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   285
apply (auto)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   286
done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   287
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   288
ML {* simp_depth_limit := 2 *}
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   289
recdef norm_float "measure (% (a,b). nat (abs a))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   290
  "norm_float (a,b) = (if (a \<noteq> 0) & (even a) then norm_float (a div 2, b+1) else (if a=0 then (0,0) else (a,b)))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   291
(hints simp: terminating_norm_float)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   292
ML {* simp_depth_limit := 1000 *}
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   293
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   294
lemma norm_float: "float x = float (norm_float x)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   295
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   296
  {
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   297
    fix a b :: int
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   298
    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
   299
    proof (induct a b rule: norm_float.induct)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   300
      case (1 u v)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   301
      show ?case
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   302
      proof cases
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   303
        assume u: "u \<noteq> 0 \<and> even u"
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   304
        with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2, v + 1))" by auto
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   305
        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
   306
        then show ?thesis
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   307
          apply (subst norm_float.simps)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   308
          apply (simp add: ind)
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   309
          done
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   310
      next
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   311
        assume "~(u \<noteq> 0 \<and> even u)"
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   312
        then show ?thesis
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   313
          by (simp add: prems float_def)
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   314
      qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   315
    qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   316
  }
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   317
  note helper = this
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   318
  have "? a b. x = (a,b)" by auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   319
  then obtain a b where "x = (a, b)" by blast
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   320
  then show ?thesis by (simp only: helper)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   321
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   322
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   323
lemma pow2_int: "pow2 (int n) = 2^n"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   324
  by (simp add: pow2_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   325
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   326
lemma float_add:
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   327
  "float (a1, e1) + float (a2, e2) =
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   328
  (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
   329
  else float (a1*2^(nat (e1-e2))+a2, e2))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   330
  apply (simp add: float_def ring_eq_simps)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   331
  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
   332
  done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   333
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   334
lemma float_mult:
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   335
  "float (a1, e1) * float (a2, e2) =
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   336
  (float (a1 * a2, e1 + e2))"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   337
  by (simp add: float_def pow2_add)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   338
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   339
lemma float_minus:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   340
  "- (float (a,b)) = float (-a, b)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   341
  by (simp add: float_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   342
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   343
lemma zero_less_pow2:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   344
  "0 < pow2 x"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   345
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   346
  {
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   347
    fix y
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   348
    have "0 <= y \<Longrightarrow> 0 < pow2 y"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   349
      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
   350
  }
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   351
  note helper=this
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   352
  show ?thesis
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   353
    apply (case_tac "0 <= x")
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   354
    apply (simp add: helper)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   355
    apply (subst pow2_neg)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   356
    apply (simp add: helper)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   357
    done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   358
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   359
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   360
lemma zero_le_float:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   361
  "(0 <= float (a,b)) = (0 <= a)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   362
  apply (auto simp add: float_def)
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   363
  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
   364
  apply (insert zero_less_pow2[of b])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   365
  apply (simp_all)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   366
  done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   367
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   368
lemma float_le_zero:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   369
  "(float (a,b) <= 0) = (a <= 0)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   370
  apply (auto simp add: float_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   371
  apply (auto simp add: mult_le_0_iff)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   372
  apply (insert zero_less_pow2[of b])
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   373
  apply auto
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   374
  done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   375
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   376
lemma float_abs:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   377
  "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
   378
  apply (auto simp add: abs_if)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   379
  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
   380
  done
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   381
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   382
lemma float_zero:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   383
  "float (0, b) = 0"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   384
  by (simp add: float_def)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   385
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   386
lemma float_pprt:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   387
  "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
   388
  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
   389
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   390
lemma float_nprt:
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   391
  "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
   392
  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
   393
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   394
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
   395
  by auto
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   396
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   397
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
   398
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   399
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   400
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
   401
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   402
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   403
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
   404
  by simp
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 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
   407
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   408
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   409
lemma int_pow_0: "(a::int)^(Numeral0) = 1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   410
  by simp
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 int_pow_1: "(a::int)^(Numeral1) = a"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   413
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   414
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   415
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
   416
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   417
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   418
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
   419
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   420
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   421
lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   422
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   423
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   424
lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   425
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   426
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   427
lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   428
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   429
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   430
lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   431
proof -
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   432
  have 1:"((-1)::nat) = 0"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   433
    by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   434
  show ?thesis by (simp add: 1)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   435
qed
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   436
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   437
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
   438
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   439
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   440
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
   441
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   442
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   443
lemma lift_bool: "x \<Longrightarrow> x=True"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   444
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   445
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   446
lemma nlift_bool: "~x \<Longrightarrow> x=False"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   447
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   448
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   449
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
   450
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   451
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
   452
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   453
lemmas binarith =
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   454
  Pls_0_eq Min_1_eq
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   455
  pred_Pls pred_Min pred_1 pred_0
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   456
  succ_Pls succ_Min succ_1 succ_0
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   457
  add_Pls add_Min add_BIT_0 add_BIT_10
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   458
  add_BIT_11 minus_Pls minus_Min minus_1
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   459
  minus_0 mult_Pls mult_Min mult_num1 mult_num0
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   460
  add_Pls_right add_Min_right
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   461
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   462
lemma int_eq_number_of_eq:
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   463
  "(((number_of v)::int)=(number_of w)) = iszero ((number_of (v + uminus w))::int)"
16782
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
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   466
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
   467
  by (simp only: iszero_number_of_Pls)
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 int_nonzero_number_of_Min: "~(iszero ((-1)::int))"
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 int_iszero_number_of_0: "iszero ((number_of (w BIT bit.B0))::int) = iszero ((number_of w)::int)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   473
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   474
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   475
lemma int_iszero_number_of_1: "\<not> iszero ((number_of (w BIT bit.B1))::int)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   476
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   477
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   478
lemma int_less_number_of_eq_neg: "(((number_of x)::int) < number_of y) = neg ((number_of (x + (uminus y)))::int)"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   479
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   480
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   481
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
   482
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   483
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   484
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
   485
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   486
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   487
lemma int_neg_number_of_BIT: "neg ((number_of (w BIT x))::int) = neg ((number_of w)::int)"
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   488
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   489
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   490
lemma int_le_number_of_eq: "(((number_of x)::int) \<le> number_of y) = (\<not> neg ((number_of (y + (uminus x)))::int))"
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   491
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   492
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   493
lemmas intarithrel =
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   494
  int_eq_number_of_eq
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   495
  lift_bool[OF int_iszero_number_of_Pls] nlift_bool[OF int_nonzero_number_of_Min] int_iszero_number_of_0
16782
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   496
  lift_bool[OF int_iszero_number_of_1] int_less_number_of_eq_neg nlift_bool[OF int_not_neg_number_of_Pls] lift_bool[OF int_neg_number_of_Min]
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   497
  int_neg_number_of_BIT int_le_number_of_eq
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   498
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   499
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
   500
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   501
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   502
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
   503
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   504
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   505
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
   506
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   507
20485
3078fd2eec7b got rid of Numeral.bin type
haftmann
parents: 20217
diff changeset
   508
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
   509
  by simp
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   510
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   511
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
   512
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   513
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
   514
19765
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   515
lemmas powerarith = nat_number_of zpower_number_of_even
dfe940911617 misc cleanup;
wenzelm
parents: 16890
diff changeset
   516
  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
   517
  zpower_Pls zpower_Min
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   518
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   519
lemmas floatarith[simplified norm_0_1] = float_add float_mult float_minus float_abs zero_le_float float_pprt float_nprt
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   520
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   521
(* for use with the compute oracle *)
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   522
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
   523
b214f21ae396 - use TableFun instead of homebrew binary tree in am_interpreter.ML
obua
parents:
diff changeset
   524
end