| author | blanchet | 
| Sun, 09 Sep 2012 21:22:31 +0200 | |
| changeset 49238 | 2267901ae911 | 
| parent 48891 | c0eafbd55de3 | 
| child 54489 | 03ff4d1e6784 | 
| permissions | -rw-r--r-- | 
| 47455 | 1  | 
(* Title: HOL/Matrix_LP/ComputeFloat.thy  | 
| 41959 | 2  | 
Author: Steven Obua  | 
| 
16782
 
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  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
38273 
diff
changeset
 | 
8  | 
imports Complex_Main "~~/src/HOL/Library/Lattice_Algebras"  | 
| 20485 | 9  | 
begin  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
10  | 
|
| 48891 | 11  | 
ML_file "~~/src/Tools/float.ML"  | 
12  | 
||
| 38273 | 13  | 
definition int_of_real :: "real \<Rightarrow> int"  | 
14  | 
where "int_of_real x = (SOME y. real y = x)"  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21256 
diff
changeset
 | 
15  | 
|
| 38273 | 16  | 
definition real_is_int :: "real \<Rightarrow> bool"  | 
17  | 
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
 | 
18  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
19  | 
lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))"  | 
| 
45495
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
20  | 
by (auto simp add: real_is_int_def int_of_real_def)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
22  | 
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
 | 
23  | 
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
 | 
24  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
25  | 
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
 | 
26  | 
by (simp add: int_of_real_def)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
28  | 
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
 | 
29  | 
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
 | 
30  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
31  | 
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
 | 
32  | 
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
 | 
33  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
34  | 
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
 | 
35  | 
apply (subst real_is_int_def2)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
36  | 
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
 | 
37  | 
done  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
38  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
39  | 
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
 | 
40  | 
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
 | 
41  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
42  | 
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
 | 
43  | 
apply (subst real_is_int_def2)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
44  | 
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
 | 
45  | 
done  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
47  | 
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
 | 
48  | 
by (auto simp add: real_is_int_def)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
49  | 
|
| 19765 | 50  | 
lemma int_of_real_mult:  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
51  | 
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
 | 
52  | 
shows "(int_of_real (a*b)) = (int_of_real a) * (int_of_real b)"  | 
| 
45495
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
53  | 
using assms  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
54  | 
by (auto simp add: real_is_int_def real_of_int_mult[symmetric]  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
55  | 
simp del: real_of_int_mult)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
56  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
57  | 
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
 | 
58  | 
apply (subst real_is_int_def2)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
59  | 
apply (simp add: int_of_real_mult)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
60  | 
done  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
61  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
62  | 
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
 | 
63  | 
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
 | 
64  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
65  | 
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
 | 
66  | 
proof -  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
67  | 
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
 | 
68  | 
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
 | 
69  | 
ultimately show ?thesis by auto  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
70  | 
qed  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
71  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
72  | 
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
 | 
73  | 
proof -  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
74  | 
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
 | 
75  | 
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
 | 
76  | 
ultimately show ?thesis by auto  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
77  | 
qed  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
78  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
79  | 
lemma real_is_int_numeral[simp]: "real_is_int (numeral x)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
80  | 
by (auto simp: real_is_int_def intro!: exI[of _ "numeral x"])  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
81  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
82  | 
lemma real_is_int_neg_numeral[simp]: "real_is_int (neg_numeral x)"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
83  | 
by (auto simp: real_is_int_def intro!: exI[of _ "neg_numeral x"])  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
84  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
85  | 
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
 | 
86  | 
by (simp add: int_of_real_def)  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
87  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
88  | 
lemma int_of_real_1[simp]: "int_of_real (1::real) = (1::int)"  | 
| 19765 | 89  | 
proof -  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
90  | 
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
 | 
91  | 
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
 | 
92  | 
qed  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
93  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
94  | 
lemma int_of_real_numeral[simp]: "int_of_real (numeral b) = numeral b"  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
95  | 
unfolding int_of_real_def  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
96  | 
by (intro some_equality)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
97  | 
(auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject)  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
98  | 
|
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
99  | 
lemma int_of_real_neg_numeral[simp]: "int_of_real (neg_numeral b) = neg_numeral b"  | 
| 
45495
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
100  | 
unfolding int_of_real_def  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
101  | 
by (intro some_equality)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
102  | 
(auto simp add: real_of_int_inject[symmetric] simp del: real_of_int_inject)  | 
| 19765 | 103  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
104  | 
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
 | 
105  | 
by (rule zdiv_int)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
107  | 
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
 | 
108  | 
by (rule zmod_int)  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
110  | 
lemma 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
 | 
111  | 
by arith  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
112  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
113  | 
lemma norm_0_1: "(1::_::numeral) = Numeral1"  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
114  | 
by auto  | 
| 19765 | 115  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
116  | 
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
 | 
117  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
118  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
119  | 
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
 | 
120  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
121  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
122  | 
lemma 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
 | 
123  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
124  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
125  | 
lemma 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
 | 
126  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
127  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
128  | 
lemma int_pow_0: "(a::int)^0 = 1"  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
129  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
130  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
131  | 
lemma int_pow_1: "(a::int)^(Numeral1) = a"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
132  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
133  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
134  | 
lemma one_eq_Numeral1_nring: "(1::'a::numeral) = Numeral1"  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
135  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
136  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
137  | 
lemma one_eq_Numeral1_nat: "(1::nat) = Numeral1"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
138  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
139  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
140  | 
lemma zpower_Pls: "(z::int)^0 = Numeral1"  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
141  | 
by simp  | 
| 
 
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 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
 | 
144  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
145  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
146  | 
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
 | 
147  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
148  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
149  | 
lemma lift_bool: "x \<Longrightarrow> x=True"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
150  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
151  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
152  | 
lemma nlift_bool: "~x \<Longrightarrow> x=False"  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
153  | 
by simp  | 
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
154  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
155  | 
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
 | 
156  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
157  | 
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
 | 
158  | 
|
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
159  | 
lemmas powerarith = nat_numeral zpower_numeral_even  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
160  | 
zpower_numeral_odd zpower_Pls  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
161  | 
|
| 
45495
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
162  | 
definition float :: "(int \<times> int) \<Rightarrow> real" where  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
163  | 
"float = (\<lambda>(a, b). real a * 2 powr real b)"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
164  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
165  | 
lemma float_add_l0: "float (0, e) + x = x"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
166  | 
by (simp add: float_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
167  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
168  | 
lemma float_add_r0: "x + float (0, e) = x"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
169  | 
by (simp add: float_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
170  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
171  | 
lemma float_add:  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
172  | 
"float (a1, e1) + float (a2, e2) =  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
173  | 
(if e1<=e2 then float (a1+a2*2^(nat(e2-e1)), e1) else float (a1*2^(nat (e1-e2))+a2, e2))"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
174  | 
by (simp add: float_def algebra_simps powr_realpow[symmetric] powr_divide2[symmetric])  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
175  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
176  | 
lemma float_mult_l0: "float (0, e) * x = float (0, 0)"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
177  | 
by (simp add: float_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
178  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
179  | 
lemma float_mult_r0: "x * float (0, e) = float (0, 0)"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
180  | 
by (simp add: float_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
181  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
182  | 
lemma float_mult:  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
183  | 
"float (a1, e1) * float (a2, e2) = (float (a1 * a2, e1 + e2))"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
184  | 
by (simp add: float_def powr_add)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
185  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
186  | 
lemma float_minus:  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
187  | 
"- (float (a,b)) = float (-a, b)"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
188  | 
by (simp add: float_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
189  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
190  | 
lemma zero_le_float:  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
191  | 
"(0 <= float (a,b)) = (0 <= a)"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
192  | 
using powr_gt_zero[of 2 "real b", arith]  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
193  | 
by (simp add: float_def zero_le_mult_iff)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
194  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
195  | 
lemma float_le_zero:  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
196  | 
"(float (a,b) <= 0) = (a <= 0)"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
197  | 
using powr_gt_zero[of 2 "real b", arith]  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
198  | 
by (simp add: float_def mult_le_0_iff)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
199  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
200  | 
lemma float_abs:  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
201  | 
"abs (float (a,b)) = (if 0 <= a then (float (a,b)) else (float (-a,b)))"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
202  | 
using powr_gt_zero[of 2 "real b", arith]  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
203  | 
by (simp add: float_def abs_if mult_less_0_iff)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
204  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
205  | 
lemma float_zero:  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
206  | 
"float (0, b) = 0"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
207  | 
by (simp add: float_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
208  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
209  | 
lemma float_pprt:  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
210  | 
"pprt (float (a, b)) = (if 0 <= a then (float (a,b)) else (float (0, b)))"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
211  | 
by (auto simp add: zero_le_float float_le_zero float_zero)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
212  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
213  | 
lemma float_nprt:  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
214  | 
"nprt (float (a, b)) = (if 0 <= a then (float (0,b)) else (float (a, b)))"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
215  | 
by (auto simp add: zero_le_float float_le_zero float_zero)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
216  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
217  | 
definition lbound :: "real \<Rightarrow> real"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
218  | 
where "lbound x = min 0 x"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
219  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
220  | 
definition ubound :: "real \<Rightarrow> real"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
221  | 
where "ubound x = max 0 x"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
222  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
223  | 
lemma lbound: "lbound x \<le> x"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
224  | 
by (simp add: lbound_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
225  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
226  | 
lemma ubound: "x \<le> ubound x"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
227  | 
by (simp add: ubound_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
228  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
229  | 
lemma pprt_lbound: "pprt (lbound x) = float (0, 0)"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
230  | 
by (auto simp: float_def lbound_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
231  | 
|
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
232  | 
lemma nprt_ubound: "nprt (ubound x) = float (0, 0)"  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
233  | 
by (auto simp: float_def ubound_def)  | 
| 
 
c55a07526dbe
cleaned up float theories; removed duplicate definitions and theorems
 
hoelzl 
parents: 
42676 
diff
changeset
 | 
234  | 
|
| 24301 | 235  | 
lemmas floatarith[simplified norm_0_1] = float_add float_add_l0 float_add_r0 float_mult float_mult_l0 float_mult_r0  | 
| 24653 | 236  | 
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
 | 
237  | 
|
| 
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
238  | 
(* for use with the compute oracle *)  | 
| 
47108
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
239  | 
lemmas arith = arith_simps rel_simps diff_nat_numeral nat_0  | 
| 
 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 
huffman 
parents: 
46988 
diff
changeset
 | 
240  | 
nat_neg_numeral powerarith floatarith not_false_eq_true not_true_eq_false  | 
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
241  | 
|
| 48891 | 242  | 
ML_file "~~/src/HOL/Tools/float_arith.ML"  | 
| 20771 | 243  | 
|
| 
16782
 
b214f21ae396
- use TableFun instead of homebrew binary tree in am_interpreter.ML
 
obua 
parents:  
diff
changeset
 | 
244  | 
end  |