| author | wenzelm | 
| Mon, 02 Nov 2015 09:43:20 +0100 | |
| changeset 61536 | 346aa2c5447f | 
| parent 61424 | c3658c18b7bc | 
| child 63355 | 7b23053be254 | 
| permissions | -rw-r--r-- | 
| 45483 | 1 | (* Authors: Florian Haftmann, Johannes Hölzl, Tobias Nipkow *) | 
| 2 | ||
| 3 | theory Code_Real_Approx_By_Float | |
| 51542 | 4 | imports Complex_Main Code_Target_Int | 
| 45483 | 5 | begin | 
| 6 | ||
| 60500 | 7 | text\<open>\textbf{WARNING} This theory implements mathematical reals by machine
 | 
| 45483 | 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. | |
| 11 | ||
| 12 | The value command cannot display real results yet. | |
| 13 | ||
| 14 | The only legitimate use of this theory is as a tool for code generation | |
| 60500 | 15 | purposes.\<close> | 
| 45483 | 16 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 17 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 18 | type_constructor real \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 19 | (SML) "real" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 20 | and (OCaml) "float" | 
| 45483 | 21 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 22 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 23 | constant Ratreal \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 24 | (SML) "error/ \"Bad constant: Ratreal\"" | 
| 45483 | 25 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 26 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 27 | constant "0 :: real" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 28 | (SML) "0.0" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 29 | and (OCaml) "0.0" | 
| 45483 | 30 | declare zero_real_code[code_unfold del] | 
| 31 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 32 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 33 | constant "1 :: real" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 34 | (SML) "1.0" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 35 | and (OCaml) "1.0" | 
| 45483 | 36 | declare one_real_code[code_unfold del] | 
| 37 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 38 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 39 | constant "HOL.equal :: real \<Rightarrow> real \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 40 | (SML) "Real.== ((_), (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 41 | and (OCaml) "Pervasives.(=)" | 
| 45483 | 42 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 43 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 44 | constant "Orderings.less_eq :: real \<Rightarrow> real \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 45 | (SML) "Real.<= ((_), (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 46 | and (OCaml) "Pervasives.(<=)" | 
| 45483 | 47 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 48 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 49 | constant "Orderings.less :: real \<Rightarrow> real \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 50 | (SML) "Real.< ((_), (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 51 | and (OCaml) "Pervasives.(<)" | 
| 45483 | 52 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 53 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 54 | constant "op + :: real \<Rightarrow> real \<Rightarrow> real" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 55 | (SML) "Real.+ ((_), (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 56 | and (OCaml) "Pervasives.( +. )" | 
| 45483 | 57 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 58 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 59 | constant "op * :: real \<Rightarrow> real \<Rightarrow> real" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 60 | (SML) "Real.* ((_), (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 61 | and (OCaml) "Pervasives.( *. )" | 
| 45483 | 62 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 63 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 64 | constant "op - :: real \<Rightarrow> real \<Rightarrow> real" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 65 | (SML) "Real.- ((_), (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 66 | and (OCaml) "Pervasives.( -. )" | 
| 45483 | 67 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 68 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 69 | constant "uminus :: real \<Rightarrow> real" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 70 | (SML) "Real.~" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 71 | and (OCaml) "Pervasives.( ~-. )" | 
| 45483 | 72 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 73 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 74 | constant "op / :: real \<Rightarrow> real \<Rightarrow> real" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 75 | (SML) "Real.'/ ((_), (_))" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 76 | and (OCaml) "Pervasives.( '/. )" | 
| 45483 | 77 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 78 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 79 | constant "HOL.equal :: real \<Rightarrow> real \<Rightarrow> bool" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 80 | (SML) "Real.== ((_:real), (_))" | 
| 45483 | 81 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 82 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 83 | constant "sqrt :: real \<Rightarrow> real" \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 84 | (SML) "Math.sqrt" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 85 | and (OCaml) "Pervasives.sqrt" | 
| 45483 | 86 | declare sqrt_def[code del] | 
| 87 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 88 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 89 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 90 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 91 | qualified definition real_exp :: "real \<Rightarrow> real" where "real_exp = exp" | 
| 45483 | 92 | |
| 93 | lemma exp_eq_real_exp[code_unfold]: "exp = real_exp" | |
| 94 | unfolding real_exp_def .. | |
| 95 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 96 | end | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 97 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 98 | code_printing | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 99 | constant Code_Real_Approx_By_Float.real_exp \<rightharpoonup> | 
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 100 | (SML) "Math.exp" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 101 | and (OCaml) "Pervasives.exp" | 
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 102 | declare Code_Real_Approx_By_Float.real_exp_def[code del] | 
| 45483 | 103 | declare exp_def[code del] | 
| 104 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 105 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 106 | constant ln \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 107 | (SML) "Math.ln" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 108 | and (OCaml) "Pervasives.ln" | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
54489diff
changeset | 109 | declare ln_real_def[code del] | 
| 45483 | 110 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 111 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 112 | constant cos \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 113 | (SML) "Math.cos" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 114 | and (OCaml) "Pervasives.cos" | 
| 45483 | 115 | declare cos_def[code del] | 
| 116 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 117 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 118 | constant sin \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 119 | (SML) "Math.sin" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 120 | and (OCaml) "Pervasives.sin" | 
| 45483 | 121 | declare sin_def[code del] | 
| 122 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 123 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 124 | constant pi \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 125 | (SML) "Math.pi" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 126 | and (OCaml) "Pervasives.pi" | 
| 45483 | 127 | declare pi_def[code del] | 
| 128 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 129 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 130 | constant arctan \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 131 | (SML) "Math.atan" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 132 | and (OCaml) "Pervasives.atan" | 
| 45483 | 133 | declare arctan_def[code del] | 
| 134 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 135 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 136 | constant arccos \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 137 | (SML) "Math.scos" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 138 | and (OCaml) "Pervasives.acos" | 
| 45483 | 139 | declare arccos_def[code del] | 
| 140 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 141 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 142 | constant arcsin \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 143 | (SML) "Math.asin" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 144 | and (OCaml) "Pervasives.asin" | 
| 45483 | 145 | declare arcsin_def[code del] | 
| 146 | ||
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
47108diff
changeset | 147 | definition real_of_integer :: "integer \<Rightarrow> real" where | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
47108diff
changeset | 148 | "real_of_integer = of_int \<circ> int_of_integer" | 
| 45483 | 149 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 150 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 151 | constant real_of_integer \<rightharpoonup> | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 152 | (SML) "Real.fromInt" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 153 | and (OCaml) "Pervasives.float (Big'_int.int'_of'_big'_int (_))" | 
| 45483 | 154 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 155 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 156 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 157 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 158 | qualified definition real_of_int :: "int \<Rightarrow> real" where | 
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
47108diff
changeset | 159 | [code_abbrev]: "real_of_int = of_int" | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
47108diff
changeset | 160 | |
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
47108diff
changeset | 161 | lemma [code]: | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
47108diff
changeset | 162 | "real_of_int = real_of_integer \<circ> integer_of_int" | 
| 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
47108diff
changeset | 163 | by (simp add: fun_eq_iff real_of_integer_def real_of_int_def) | 
| 45483 | 164 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 165 | lemma [code_unfold del]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 166 | "0 \<equiv> (of_rat 0 :: real)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 167 | by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 168 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 169 | lemma [code_unfold del]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 170 | "1 \<equiv> (of_rat 1 :: real)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 171 | by simp | 
| 45483 | 172 | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 173 | lemma [code_unfold del]: | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 174 | "numeral k \<equiv> (of_rat (numeral k) :: real)" | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 175 | by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 176 | |
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 177 | lemma [code_unfold del]: | 
| 54489 
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
 haftmann parents: 
52435diff
changeset | 178 | "- numeral k \<equiv> (of_rat (- numeral k) :: real)" | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 179 | by simp | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46530diff
changeset | 180 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 181 | end | 
| 45483 | 182 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 183 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 184 | constant Ratreal \<rightharpoonup> (SML) | 
| 52403 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 185 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 186 | definition Realfract :: "int => int => real" | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 187 | where | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 188 | "Realfract p q = of_int p / of_int q" | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 189 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 190 | code_datatype Realfract | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 191 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 192 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
52403diff
changeset | 193 | constant Realfract \<rightharpoonup> (SML) "Real.fromInt _/ '// Real.fromInt _" | 
| 52403 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 194 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 195 | lemma [code]: | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61115diff
changeset | 196 | "Ratreal r = case_prod Realfract (quotient_of r)" | 
| 52403 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 197 | by (cases r) (simp add: Realfract_def quotient_of_Fract of_rat_rat) | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 198 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 199 | lemma [code, code del]: | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 200 | "(HOL.equal :: real=>real=>bool) = (HOL.equal :: real => real => bool) " | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 201 | .. | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 202 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 203 | lemma [code, code del]: | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 204 | "(plus :: real => real => real) = plus" | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 205 | .. | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 206 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 207 | lemma [code, code del]: | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 208 | "(uminus :: real => real) = uminus" | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 209 | .. | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 210 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 211 | lemma [code, code del]: | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 212 | "(minus :: real => real => real) = minus" | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 213 | .. | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 214 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 215 | lemma [code, code del]: | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 216 | "(times :: real => real => real) = times" | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 217 | .. | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 218 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 219 | lemma [code, code del]: | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 220 | "(divide :: real => real => real) = divide" | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 221 | .. | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 222 | |
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 223 | lemma [code]: | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 224 | fixes r :: real | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 225 | shows "inverse r = 1 / r" | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 226 | by (fact inverse_eq_divide) | 
| 
140ae824ea4a
Code_Real_Approx_By_Float: remove code equations using Ratreal
 hoelzl parents: 
51542diff
changeset | 227 | |
| 46530 | 228 | notepad | 
| 229 | begin | |
| 230 | have "cos (pi/2) = 0" by (rule cos_pi_half) | |
| 231 | moreover have "cos (pi/2) \<noteq> 0" by eval | |
| 232 | ultimately have "False" by blast | |
| 233 | end | |
| 45483 | 234 | |
| 235 | end | |
| 51143 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 haftmann parents: 
47108diff
changeset | 236 |