| author | wenzelm | 
| Fri, 17 Dec 2010 18:10:37 +0100 | |
| changeset 41249 | 26f12f98f50a | 
| parent 38273 | d31a34569542 | 
| child 41413 | 64cd30d6b0b8 | 
| permissions | -rw-r--r-- | 
| 
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 | 5  | 
header {* Floating Point Representation of the Reals *}
 | 
6  | 
||
| 
29804
 
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
 
hoelzl 
parents: 
29667 
diff
changeset
 | 
7  | 
theory ComputeFloat  | 
| 
35032
 
7efe662e41b4
separate library theory for type classes combining lattices with various algebraic structures
 
haftmann 
parents: 
32491 
diff
changeset
 | 
8  | 
imports Complex_Main Lattice_Algebras  | 
| 
28952
 
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 | 10  | 
begin  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
11  | 
|
| 38273 | 12  | 
definition pow2 :: "int \<Rightarrow> real"  | 
13  | 
where "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
 | 
14  | 
|
| 38273 | 15  | 
definition float :: "int * int \<Rightarrow> real"  | 
16  | 
where "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 | 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])  | 
| 
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
 | 
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)  | 
| 29667 | 39  | 
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
 | 
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"])  | 
| 
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
 | 
45  | 
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
 | 
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 | 48  | 
qed  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
49  | 
qed  | 
| 19765 | 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 | 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)  | 
| 29667 | 60  | 
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
 | 
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 | 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 | 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"])  | 
| 29667 | 73  | 
apply (simp add: algebra_simps)  | 
| 
16782
 
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 | 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"])  | 
| 29667 | 84  | 
apply (auto simp add: algebra_simps)  | 
| 
16782
 
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)  | 
| 29667 | 92  | 
apply (auto simp add: algebra_simps)  | 
| 
16782
 
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 | 97  | 
lemma "float (a, e) + float (b, e) = float (a + b, e)"  | 
| 29667 | 98  | 
by (simp add: float_def algebra_simps)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
99  | 
|
| 38273 | 100  | 
definition int_of_real :: "real \<Rightarrow> int"  | 
101  | 
where "int_of_real x = (SOME y. real y = x)"  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21256 
diff
changeset
 | 
102  | 
|
| 38273 | 103  | 
definition real_is_int :: "real \<Rightarrow> bool"  | 
104  | 
where "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  | 
|
| 26313 | 112  | 
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
 | 
113  | 
by (simp add: pow2_def)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
114  | 
|
| 19765 | 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 | 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 | 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 | 192  | 
|
193  | 
  {
 | 
|
| 20485 | 194  | 
fix x :: int  | 
195  | 
have "real_is_int ((number_of \<Colon> int \<Rightarrow> real) x)"  | 
|
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 | 208  | 
apply (simp only: s of_int_add)  | 
209  | 
apply (rule real_is_int_add)  | 
|
210  | 
apply (simp add: neg1)  | 
|
211  | 
apply (simp only: rn)  | 
|
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 | 217  | 
fix x :: int  | 
218  | 
have "? u. x = u"  | 
|
219  | 
apply (rule exI[where x = "x"])  | 
|
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 | 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 | 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 | 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 | 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])  | 
| 29667 | 253  | 
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
 | 
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 | 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 | 261  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
262  | 
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
 | 
263  | 
by (rule zdiv_int)  | 
| 
16782
 
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_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
 | 
266  | 
by (rule zmod_int)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
267  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
268  | 
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
 | 
269  | 
by arith  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
270  | 
|
| 27366 | 271  | 
function norm_float :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where  | 
272  | 
"norm_float a b = (if a \<noteq> 0 \<and> even a then norm_float (a div 2) (b + 1)  | 
|
273  | 
else if a = 0 then (0, 0) else (a, b))"  | 
|
274  | 
by auto  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
275  | 
|
| 27366 | 276  | 
termination by (relation "measure (nat o abs o fst)")  | 
277  | 
(auto intro: abs_div_2_less)  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
278  | 
|
| 27366 | 279  | 
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
 | 
280  | 
proof -  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
281  | 
  {
 | 
| 19765 | 282  | 
fix a b :: int  | 
| 27366 | 283  | 
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
 | 
284  | 
proof (induct a b rule: norm_float.induct)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
285  | 
case (1 u v)  | 
| 19765 | 286  | 
show ?case  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
287  | 
proof cases  | 
| 19765 | 288  | 
assume u: "u \<noteq> 0 \<and> even u"  | 
| 27366 | 289  | 
with prems have ind: "float (u div 2, v + 1) = float (norm_float (u div 2) (v + 1))" by auto  | 
| 19765 | 290  | 
with u have "float (u,v) = float (u div 2, v+1)" by (simp add: float_transfer_even)  | 
291  | 
then show ?thesis  | 
|
292  | 
apply (subst norm_float.simps)  | 
|
293  | 
apply (simp add: ind)  | 
|
294  | 
done  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
295  | 
next  | 
| 19765 | 296  | 
assume "~(u \<noteq> 0 \<and> even u)"  | 
297  | 
then show ?thesis  | 
|
298  | 
by (simp add: prems float_def)  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
299  | 
qed  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
300  | 
qed  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
301  | 
}  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
302  | 
note helper = this  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
303  | 
have "? a b. x = (a,b)" by auto  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
304  | 
then obtain a b where "x = (a, b)" by blast  | 
| 27366 | 305  | 
then show ?thesis by (simp add: helper)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
306  | 
qed  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
307  | 
|
| 24301 | 308  | 
lemma float_add_l0: "float (0, e) + x = x"  | 
309  | 
by (simp add: float_def)  | 
|
310  | 
||
311  | 
lemma float_add_r0: "x + float (0, e) = x"  | 
|
312  | 
by (simp add: float_def)  | 
|
313  | 
||
| 19765 | 314  | 
lemma float_add:  | 
315  | 
"float (a1, e1) + float (a2, e2) =  | 
|
316  | 
(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
 | 
317  | 
else float (a1*2^(nat (e1-e2))+a2, e2))"  | 
| 29667 | 318  | 
apply (simp add: float_def algebra_simps)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
319  | 
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
 | 
320  | 
done  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
321  | 
|
| 24301 | 322  | 
lemma float_add_assoc1:  | 
323  | 
"(x + float (y1, e1)) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"  | 
|
324  | 
by simp  | 
|
325  | 
||
326  | 
lemma float_add_assoc2:  | 
|
327  | 
"(float (y1, e1) + x) + float (y2, e2) = (float (y1, e1) + float (y2, e2)) + x"  | 
|
328  | 
by simp  | 
|
329  | 
||
330  | 
lemma float_add_assoc3:  | 
|
331  | 
"float (y1, e1) + (x + float (y2, e2)) = (float (y1, e1) + float (y2, e2)) + x"  | 
|
332  | 
by simp  | 
|
333  | 
||
334  | 
lemma float_add_assoc4:  | 
|
335  | 
"float (y1, e1) + (float (y2, e2) + x) = (float (y1, e1) + float (y2, e2)) + x"  | 
|
336  | 
by simp  | 
|
337  | 
||
338  | 
lemma float_mult_l0: "float (0, e) * x = float (0, 0)"  | 
|
339  | 
by (simp add: float_def)  | 
|
340  | 
||
341  | 
lemma float_mult_r0: "x * float (0, e) = float (0, 0)"  | 
|
342  | 
by (simp add: float_def)  | 
|
343  | 
||
| 38273 | 344  | 
definition lbound :: "real \<Rightarrow> real"  | 
345  | 
where "lbound x = min 0 x"  | 
|
| 24301 | 346  | 
|
| 38273 | 347  | 
definition ubound :: "real \<Rightarrow> real"  | 
348  | 
where "ubound x = max 0 x"  | 
|
| 24301 | 349  | 
|
350  | 
lemma lbound: "lbound x \<le> x"  | 
|
351  | 
by (simp add: lbound_def)  | 
|
352  | 
||
353  | 
lemma ubound: "x \<le> ubound x"  | 
|
354  | 
by (simp add: ubound_def)  | 
|
355  | 
||
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
356  | 
lemma float_mult:  | 
| 19765 | 357  | 
"float (a1, e1) * float (a2, e2) =  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
358  | 
(float (a1 * a2, e1 + e2))"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
359  | 
by (simp add: float_def pow2_add)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
360  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
361  | 
lemma float_minus:  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
362  | 
"- (float (a,b)) = float (-a, b)"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
363  | 
by (simp add: float_def)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
364  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
365  | 
lemma zero_less_pow2:  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
366  | 
"0 < pow2 x"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
367  | 
proof -  | 
| 
 
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  | 
fix y  | 
| 19765 | 370  | 
have "0 <= y \<Longrightarrow> 0 < pow2 y"  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
371  | 
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
 | 
372  | 
}  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
373  | 
note helper=this  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
374  | 
show ?thesis  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
375  | 
apply (case_tac "0 <= x")  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
376  | 
apply (simp add: helper)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
377  | 
apply (subst pow2_neg)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
378  | 
apply (simp add: helper)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
379  | 
done  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
380  | 
qed  | 
| 
 
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 zero_le_float:  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
383  | 
"(0 <= float (a,b)) = (0 <= a)"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
384  | 
apply (auto simp add: float_def)  | 
| 19765 | 385  | 
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
 | 
386  | 
apply (insert zero_less_pow2[of b])  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
387  | 
apply (simp_all)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
388  | 
done  | 
| 
 
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_le_zero:  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
391  | 
"(float (a,b) <= 0) = (a <= 0)"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
392  | 
apply (auto simp add: float_def)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
393  | 
apply (auto simp add: mult_le_0_iff)  | 
| 
 
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 auto  | 
| 
 
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_abs:  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
399  | 
"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
 | 
400  | 
apply (auto simp add: abs_if)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
401  | 
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
 | 
402  | 
done  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
403  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
404  | 
lemma float_zero:  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
405  | 
"float (0, b) = 0"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
406  | 
by (simp add: float_def)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
407  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
408  | 
lemma float_pprt:  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
409  | 
"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
 | 
410  | 
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
 | 
411  | 
|
| 24301 | 412  | 
lemma pprt_lbound: "pprt (lbound x) = float (0, 0)"  | 
413  | 
apply (simp add: float_def)  | 
|
414  | 
apply (rule pprt_eq_0)  | 
|
415  | 
apply (simp add: lbound_def)  | 
|
416  | 
done  | 
|
417  | 
||
418  | 
lemma nprt_ubound: "nprt (ubound x) = float (0, 0)"  | 
|
419  | 
apply (simp add: float_def)  | 
|
420  | 
apply (rule nprt_eq_0)  | 
|
421  | 
apply (simp add: ubound_def)  | 
|
422  | 
done  | 
|
423  | 
||
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
424  | 
lemma float_nprt:  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
425  | 
"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
 | 
426  | 
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
 | 
427  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
428  | 
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
 | 
429  | 
by auto  | 
| 19765 | 430  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
431  | 
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
 | 
432  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
433  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
434  | 
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
 | 
435  | 
by simp  | 
| 
 
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 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
 | 
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 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
 | 
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 int_pow_0: "(a::int)^(Numeral0) = 1"  | 
| 
 
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 int_pow_1: "(a::int)^(Numeral1) = a"  | 
| 
 
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 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
 | 
450  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
451  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
452  | 
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
 | 
453  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
454  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
455  | 
lemma zero_eq_Numeral0_nat: "(0::nat) = Numeral0"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
456  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
457  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
458  | 
lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
459  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
460  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
461  | 
lemma zpower_Pls: "(z::int)^Numeral0 = Numeral1"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
462  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
463  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
464  | 
lemma zpower_Min: "(z::int)^((-1)::nat) = Numeral1"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
465  | 
proof -  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
466  | 
have 1:"((-1)::nat) = 0"  | 
| 
 
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  | 
show ?thesis by (simp add: 1)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
469  | 
qed  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
470  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
471  | 
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
 | 
472  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
473  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
474  | 
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
 | 
475  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
476  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
477  | 
lemma lift_bool: "x \<Longrightarrow> x=True"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
478  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
479  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
480  | 
lemma nlift_bool: "~x \<Longrightarrow> x=False"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
481  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
482  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
483  | 
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
 | 
484  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
485  | 
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
 | 
486  | 
|
| 19765 | 487  | 
lemmas binarith =  | 
| 
26076
 
b9c716a9fb5f
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 
huffman 
parents: 
24653 
diff
changeset
 | 
488  | 
normalize_bin_simps  | 
| 
 
b9c716a9fb5f
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 
huffman 
parents: 
24653 
diff
changeset
 | 
489  | 
pred_bin_simps succ_bin_simps  | 
| 
 
b9c716a9fb5f
added lemma lists {normalize,succ,pred,minus,add,mult}_bin_simps
 
huffman 
parents: 
24653 
diff
changeset
 | 
490  | 
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
 | 
491  | 
|
| 20485 | 492  | 
lemma int_eq_number_of_eq:  | 
493  | 
"(((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
 | 
494  | 
by (rule eq_number_of_eq)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
495  | 
|
| 19765 | 496  | 
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
 | 
497  | 
by (simp only: iszero_number_of_Pls)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
498  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
499  | 
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
 | 
500  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
501  | 
|
| 
26086
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26076 
diff
changeset
 | 
502  | 
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
 | 
503  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
504  | 
|
| 
26086
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26076 
diff
changeset
 | 
505  | 
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
 | 
506  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
507  | 
|
| 20485 | 508  | 
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
 | 
509  | 
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
 | 
510  | 
|
| 19765 | 511  | 
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
 | 
512  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
513  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
514  | 
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
 | 
515  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
516  | 
|
| 
26086
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26076 
diff
changeset
 | 
517  | 
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
 | 
518  | 
by simp  | 
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26076 
diff
changeset
 | 
519  | 
|
| 
 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 
huffman 
parents: 
26076 
diff
changeset
 | 
520  | 
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
 | 
521  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
522  | 
|
| 20485 | 523  | 
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
 | 
524  | 
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
 | 
525  | 
|
| 19765 | 526  | 
lemmas intarithrel =  | 
527  | 
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
 | 
528  | 
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
 | 
529  | 
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
 | 
530  | 
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
 | 
531  | 
|
| 20485 | 532  | 
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
 | 
533  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
534  | 
|
| 20485 | 535  | 
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
 | 
536  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
537  | 
|
| 20485 | 538  | 
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
 | 
539  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
540  | 
|
| 20485 | 541  | 
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
 | 
542  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
543  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
544  | 
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
 | 
545  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
546  | 
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
 | 
547  | 
|
| 19765 | 548  | 
lemmas powerarith = nat_number_of zpower_number_of_even  | 
549  | 
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
 | 
550  | 
zpower_Pls zpower_Min  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
551  | 
|
| 24301 | 552  | 
lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0  | 
| 24653 | 553  | 
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
 | 
554  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
555  | 
(* for use with the compute oracle *)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
556  | 
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
 | 
557  | 
|
| 
28952
 
15a4b2cf8c34
made repository layout more coherent with logical distribution structure; stripped some $Id$s
 
haftmann 
parents: 
27366 
diff
changeset
 | 
558  | 
use "~~/src/HOL/Tools/float_arith.ML"  | 
| 20771 | 559  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
560  | 
end  |