1 (* Authors: Florian Haftmann, Johannes Hölzl, Tobias Nipkow *)
3 theory Code_Real_Approx_By_Float
4 imports Complex_Main "~~/src/HOL/Library/Code_Integer"
7 text{* \textbf{WARNING} This theory implements mathematical reals by machine
8 reals (floats). This is inconsistent. See the proof of False at the end of
9 the theory, where an equality on mathematical reals is (incorrectly)
10 disproved by mapping it to machine reals.
12 The value command cannot display real results yet.
14 The only legitimate use of this theory is as a tool for code generation
22 (SML "error/ \"Bad constant: Ratreal\"")
24 code_const "0 :: real"
27 declare zero_real_code[code_unfold del]
29 code_const "1 :: real"
32 declare one_real_code[code_unfold del]
34 code_const "HOL.equal :: real \<Rightarrow> real \<Rightarrow> bool"
35 (SML "Real.== ((_), (_))")
36 (OCaml "Pervasives.(=)")
38 code_const "Orderings.less_eq :: real \<Rightarrow> real \<Rightarrow> bool"
39 (SML "Real.<= ((_), (_))")
40 (OCaml "Pervasives.(<=)")
42 code_const "Orderings.less :: real \<Rightarrow> real \<Rightarrow> bool"
43 (SML "Real.< ((_), (_))")
44 (OCaml "Pervasives.(<)")
46 code_const "op + :: real \<Rightarrow> real \<Rightarrow> real"
47 (SML "Real.+ ((_), (_))")
48 (OCaml "Pervasives.( +. )")
50 code_const "op * :: real \<Rightarrow> real \<Rightarrow> real"
51 (SML "Real.* ((_), (_))")
52 (OCaml "Pervasives.( *. )")
54 code_const "op - :: real \<Rightarrow> real \<Rightarrow> real"
55 (SML "Real.- ((_), (_))")
56 (OCaml "Pervasives.( -. )")
58 code_const "uminus :: real \<Rightarrow> real"
60 (OCaml "Pervasives.( ~-. )")
62 code_const "op / :: real \<Rightarrow> real \<Rightarrow> real"
63 (SML "Real.'/ ((_), (_))")
64 (OCaml "Pervasives.( '/. )")
66 code_const "HOL.equal :: real \<Rightarrow> real \<Rightarrow> bool"
67 (SML "Real.== ((_:real), (_))")
69 code_const "sqrt :: real \<Rightarrow> real"
71 (OCaml "Pervasives.sqrt")
72 declare sqrt_def[code del]
74 definition real_exp :: "real \<Rightarrow> real" where "real_exp = exp"
76 lemma exp_eq_real_exp[code_unfold]: "exp = real_exp"
77 unfolding real_exp_def ..
81 (OCaml "Pervasives.exp")
82 declare real_exp_def[code del]
83 declare exp_def[code del]
85 hide_const (open) real_exp
89 (OCaml "Pervasives.ln")
90 declare ln_def[code del]
94 (OCaml "Pervasives.cos")
95 declare cos_def[code del]
99 (OCaml "Pervasives.sin")
100 declare sin_def[code del]
104 (OCaml "Pervasives.pi")
105 declare pi_def[code del]
109 (OCaml "Pervasives.atan")
110 declare arctan_def[code del]
114 (OCaml "Pervasives.acos")
115 declare arccos_def[code del]
119 (OCaml "Pervasives.asin")
120 declare arcsin_def[code del]
122 definition real_of_int :: "int \<Rightarrow> real" where
123 "real_of_int \<equiv> of_int"
125 code_const real_of_int
127 (OCaml "Pervasives.float (Big'_int.int'_of'_big'_int (_))")
129 lemma of_int_eq_real_of_int[code_unfold]: "of_int = real_of_int"
130 unfolding real_of_int_def ..
132 hide_const (open) real_of_int
134 declare number_of_real_code [code_unfold del]
138 have "cos(pi/2) = 0" by(rule cos_pi_half)
139 moreover have "cos(pi/2) \<noteq> 0" by eval
140 ultimately show "False" by blast
144 sorry -- "Use quick_and_dirty to load this theory."