# HG changeset patch # User haftmann # Date 1660155982 0 # Node ID a21debbc707460655e858620d5f937a3dd019646 # Parent f1141438b4db25ed33c6f8870d251832aa5be40c Further streamlining of quick-and-dirty evaluation. diff -r f1141438b4db -r a21debbc7074 src/HOL/Library/Code_Real_Approx_By_Float.thy --- a/src/HOL/Library/Code_Real_Approx_By_Float.thy Mon Jul 25 08:24:21 2022 +0100 +++ b/src/HOL/Library/Code_Real_Approx_By_Float.thy Wed Aug 10 18:26:22 2022 +0000 @@ -1,4 +1,6 @@ (* Title: HOL/Library/Code_Real_Approx_By_Float.thy + Author: Jesús Aransay + Author: Jose Divasón Author: Florian Haftmann Author: Johannes Hölzl Author: Tobias Nipkow @@ -20,10 +22,21 @@ purposes. \ -definition Realfract :: \real \ real \ real\ - where \Realfract p q = p / q\ +context +begin + +qualified definition real_of_integer :: "integer \ real" + where [code_abbrev]: "real_of_integer = of_int \ int_of_integer" + +end -code_datatype Realfract +code_datatype Code_Real_Approx_By_Float.real_of_integer \(/) :: real \ real \ real\ + +lemma [code_unfold del]: "numeral k \ real_of_rat (numeral k)" + by simp + +lemma [code_unfold del]: "- numeral k \ real_of_rat (- numeral k)" + by simp context begin @@ -31,11 +44,8 @@ qualified definition real_of_int :: \int \ real\ where [code_abbrev]: \real_of_int = of_int\ -qualified definition real_of_integer :: "integer \ real" - where [code_abbrev]: "real_of_integer = of_int \ int_of_integer" - -lemma [code]: "real_of_int = real_of_integer \ integer_of_int" - by (simp add: fun_eq_iff real_of_integer_def real_of_int_def) +lemma [code]: "real_of_int = Code_Real_Approx_By_Float.real_of_integer \ integer_of_int" + by (simp add: fun_eq_iff Code_Real_Approx_By_Float.real_of_integer_def real_of_int_def) qualified definition exp_real :: \real \ real\ where [code_abbrev, code del]: \exp_real = exp\ @@ -46,33 +56,31 @@ qualified definition cos_real :: \real \ real\ where [code_abbrev, code del]: \cos_real = cos\ +qualified definition tan_real :: \real \ real\ + where [code_abbrev, code del]: \tan_real = tan\ + end -definition Quotreal :: \int \ int \ real\ - where \Quotreal p q = Realfract (of_int p) (of_int q)\ - -lemma [code]: "Ratreal r = case_prod Quotreal (quotient_of r)" - by (cases r) (simp add: Realfract_def Quotreal_def quotient_of_Fract of_rat_rat) +lemma [code]: \Ratreal r = (case quotient_of r of (p, q) \ real_of_int p / real_of_int q)\ + by (cases r) (simp add: quotient_of_Fract of_rat_rat) lemma [code]: \inverse r = 1 / r\ for r :: real by (fact inverse_eq_divide) -lemma [code_unfold del]: "numeral k \ real_of_rat (numeral k)" - by simp - -lemma [code_unfold del]: "- numeral k \ real_of_rat (- numeral k)" - by simp - declare [[code drop: \HOL.equal :: real \ real \ bool\ + \(\) :: real \ real \ bool\ + \(<) :: real \ real \ bool\ \plus :: real \ real \ real\ + \times :: real \ real \ real\ \uminus :: real \ real\ \minus :: real \ real \ real\ - \times :: real \ real \ real\ \divide :: real \ real \ real\ - \(<) :: real \ real \ bool\ - \(\) :: real \ real \ bool\ - sqrt \ln :: real \ real\ pi - arcsin arccos arctan]] + sqrt + \ln :: real \ real\ + pi + arcsin + arccos + arctan]] code_reserved SML Real @@ -80,109 +88,91 @@ type_constructor real \ (SML) "real" and (OCaml) "float" - -code_printing - constant Realfract \ (SML) "_/ '// _" - -code_printing - constant "0 :: real" \ + and (Haskell) "Prelude.Double" (*Double precision*) +| constant "0 :: real" \ (SML) "0.0" and (OCaml) "0.0" - -code_printing - constant "1 :: real" \ + and (Haskell) "0.0" +| constant "1 :: real" \ (SML) "1.0" and (OCaml) "1.0" - -code_printing - constant "HOL.equal :: real \ real \ bool" \ + and (Haskell) "1.0" +| constant "HOL.equal :: real \ real \ bool" \ (SML) "Real.== ((_), (_))" and (OCaml) "Pervasives.(=)" - -code_printing - constant "Orderings.less_eq :: real \ real \ bool" \ + and (Haskell) infix 4 "==" +| class_instance real :: "HOL.equal" => (Haskell) - (*This is necessary. See the tutorial on code generation, page 29*) +| constant "(\) :: real \ real \ bool" \ (SML) "Real.<= ((_), (_))" and (OCaml) "Pervasives.(<=)" - -code_printing - constant "Orderings.less :: real \ real \ bool" \ + and (Haskell) infix 4 "<=" +| constant "(<) :: real \ real \ bool" \ (SML) "Real.< ((_), (_))" and (OCaml) "Pervasives.(<)" - -code_printing - constant "(+) :: real \ real \ real" \ + and (Haskell) infix 4 "<" +| constant "(+) :: real \ real \ real" \ (SML) "Real.+ ((_), (_))" and (OCaml) "Pervasives.( +. )" - -code_printing - constant "(*) :: real \ real \ real" \ + and (Haskell) infixl 6 "+" +| constant "(*) :: real \ real \ real" \ (SML) "Real.* ((_), (_))" - and (OCaml) "Pervasives.( *. )" - -code_printing - constant "(-) :: real \ real \ real" \ - (SML) "Real.- ((_), (_))" - and (OCaml) "Pervasives.( -. )" - -code_printing - constant "uminus :: real \ real" \ + and (Haskell) infixl 7 "*" +| constant "uminus :: real \ real" \ (SML) "Real.~" and (OCaml) "Pervasives.( ~-. )" - -code_printing - constant "(/) :: real \ real \ real" \ + and (Haskell) "negate" +| constant "(-) :: real \ real \ real" \ + (SML) "Real.- ((_), (_))" + and (OCaml) "Pervasives.( -. )" + and (Haskell) infixl 6 "-" +| constant "(/) :: real \ real \ real" \ (SML) "Real.'/ ((_), (_))" and (OCaml) "Pervasives.( '/. )" - -code_printing - constant "sqrt :: real \ real" \ + and (Haskell) infixl 7 "/" +| constant "sqrt :: real \ real" \ (SML) "Math.sqrt" and (OCaml) "Pervasives.sqrt" - -code_printing - constant Code_Real_Approx_By_Float.exp_real \ + and (Haskell) "Prelude.sqrt" +| constant Code_Real_Approx_By_Float.exp_real \ (SML) "Math.exp" and (OCaml) "Pervasives.exp" - -code_printing - constant ln \ + and (Haskell) "Prelude.exp" +| constant ln \ (SML) "Math.ln" and (OCaml) "Pervasives.ln" - -code_printing - constant Code_Real_Approx_By_Float.cos_real \ + and (Haskell) "Prelude.log" +| constant Code_Real_Approx_By_Float.sin_real \ + (SML) "Math.sin" + and (OCaml) "Pervasives.sin" + and (Haskell) "Prelude.sin" +| constant Code_Real_Approx_By_Float.cos_real \ (SML) "Math.cos" and (OCaml) "Pervasives.cos" - -code_printing - constant Code_Real_Approx_By_Float.sin_real \ - (SML) "Math.sin" - and (OCaml) "Pervasives.sin" - -code_printing - constant pi \ + and (Haskell) "Prelude.cos" +| constant Code_Real_Approx_By_Float.tan_real \ + (SML) "Math.tan" + and (OCaml) "Pervasives.tan" + and (Haskell) "Prelude.tan" +| constant pi \ (SML) "Math.pi" - and (OCaml) "Pervasives.pi" - -code_printing - constant arctan \ + (*missing in OCaml*) + and (Haskell) "Prelude.pi" +| constant arcsin \ + (SML) "Math.asin" + and (OCaml) "Pervasives.asin" + and (Haskell) "Prelude.asin" +| constant arccos \ + (SML) "Math.scos" + and (OCaml) "Pervasives.acos" + and (Haskell) "Prelude.acos" +| constant arctan \ (SML) "Math.atan" and (OCaml) "Pervasives.atan" - -code_printing - constant arccos \ - (SML) "Math.scos" - and (OCaml) "Pervasives.acos" - -code_printing - constant arcsin \ - (SML) "Math.asin" - and (OCaml) "Pervasives.asin" - -code_printing - constant Code_Real_Approx_By_Float.real_of_integer \ + and (Haskell) "Prelude.atan" +| constant Code_Real_Approx_By_Float.real_of_integer \ (SML) "Real.fromInt" and (OCaml) "Pervasives.float/ (Big'_int.to'_int (_))" + and (Haskell) "Prelude.fromIntegral (_)" notepad begin