diff options
| author | Maxime Dénès | 2019-11-01 15:53:30 +0100 |
|---|---|---|
| committer | Maxime Dénès | 2019-11-01 15:53:30 +0100 |
| commit | fdabd4dbd6bfd60ad46fc8c945ed063860498e53 (patch) | |
| tree | 01edf91f8b536ad4acfbba39e114daa06b40f3f8 | |
| parent | d00c0b93ec4cb5ca48a9dc2ddf2cfd2038208ee2 (diff) | |
| parent | acdaab2a8c2ccb63df364bb75de8a515b2cef484 (diff) | |
Merge PR #9867: Add primitive floats (binary64 floating-point numbers)
Ack-by: SkySkimmer
Reviewed-by: Zimmi48
Ack-by: ejgallego
Reviewed-by: maximedenes
Ack-by: proux01
Ack-by: silene
Ack-by: vbgl
179 files changed, 4803 insertions, 265 deletions
diff --git a/.gitignore b/.gitignore index 7b3eadf33b..a1c0dc774e 100644 --- a/.gitignore +++ b/.gitignore @@ -152,6 +152,7 @@ plugins/ssr/ssrvernac.ml # other auto-generated files +kernel/byterun/dune.c_flags kernel/byterun/coq_instruct.h kernel/byterun/coq_jumptbl.h kernel/genOpcodeFiles.exe @@ -21,9 +21,19 @@ WHAT DO YOU NEED ? - a C compiler + - an IEEE-754 compliant architecture with rounding to nearest + ties to even as default rounding mode (most architectures + should work nowadays) + - for CoqIDE, the lablgtk development files (version >= 3.0.0), and the GTK 3.x libraries including gtksourceview3. + The IEEE-754 compliance is required by primitive floating-point + numbers (Require Import Floats). Common sources of incompatibility + are checked at configure time, preventing compilation. In the, + unlikely, event an incompatibility remains undetected, using Floats + would enable to prove False on this architecture. + Note that num and lablgtk should be properly registered with findlib/ocamlfind as Coq's makefile will use it to locate the libraries during the build. diff --git a/Makefile.build b/Makefile.build index ed4cde2b16..b63d582740 100644 --- a/Makefile.build +++ b/Makefile.build @@ -566,15 +566,15 @@ $(FAKEIDEBYTE): $(FAKEIDECMO) | $(IDETOPBYTE) # votour: a small vo explorer (based on the checker) -VOTOURCMO:=clib/cObj.cmo kernel/uint63.cmo checker/analyze.cmo checker/values.cmo checker/votour.cmo +VOTOURCMO:=clib/cObj.cmo kernel/uint63.cmo kernel/float64.cmo checker/analyze.cmo checker/values.cmo checker/votour.cmo -bin/votour: $(call bestobj, $(VOTOURCMO)) +bin/votour: $(call bestobj, $(VOTOURCMO)) $(LIBCOQRUN) $(SHOW)'OCAMLBEST -o $@' $(HIDE)$(call bestocaml, -I checker) -bin/votour.byte: $(VOTOURCMO) +bin/votour.byte: $(VOTOURCMO) $(LIBCOQRUN) $(SHOW)'OCAMLC -o $@' - $(HIDE)$(call ocamlbyte, -I checker) + $(HIDE)$(call ocamlbyte, -I checker $(VMBYTEFLAGS)) ########################################################################### # Csdp to micromega special targets diff --git a/Makefile.common b/Makefile.common index 1ad255d156..e392e51153 100644 --- a/Makefile.common +++ b/Makefile.common @@ -149,6 +149,7 @@ RTAUTOCMO:=plugins/rtauto/rtauto_plugin.cmo SYNTAXCMO:=$(addprefix plugins/syntax/, \ r_syntax_plugin.cmo \ int63_syntax_plugin.cmo \ + float_syntax_plugin.cmo \ numeral_notation_plugin.cmo \ string_notation_plugin.cmo) DERIVECMO:=plugins/derive/derive_plugin.cmo diff --git a/checker/analyze.ml b/checker/analyze.ml index e145988b4f..4c06f1e250 100644 --- a/checker/analyze.ml +++ b/checker/analyze.ml @@ -107,6 +107,7 @@ end type repr = | RInt of int | RInt63 of Uint63.t +| RFloat64 of Float64.t | RBlock of (int * int) (* tag × len *) | RString of string | RPointer of int @@ -121,6 +122,7 @@ type data = type obj = | Struct of int * data array (* tag × data *) | Int63 of Uint63.t (* Primitive integer *) +| Float64 of Float64.t (* Primitive float *) | String of string module type Input = @@ -279,6 +281,25 @@ let input_intL chan : int64 = (i lsl 56) lor (j lsl 48) lor (k lsl 40) lor (l lsl 32) lor (m lsl 24) lor (n lsl 16) lor (o lsl 8) lor (Int64.of_int p) +let input_double_big chan : float = + Int64.float_of_bits (input_intL chan) + +let input_double_little chan : float = + let i = input_byte chan in + let j = input_byte chan in + let k = input_byte chan in + let l = input_byte chan in + let m = input_byte chan in + let n = input_byte chan in + let o = input_byte chan in + let p = input_byte chan in + let ( lsl ) x y = Int64.(shift_left (of_int x) y) in + let ( lor ) = Int64.logor in + let bits = + (p lsl 56) lor (o lsl 48) lor (n lsl 40) lor (m lsl 32) lor + (l lsl 24) lor (k lsl 16) lor (j lsl 8) lor (Int64.of_int i) in + Int64.float_of_bits bits + let parse_object chan = let data = input_byte chan in if prefix_small_block <= data then @@ -326,9 +347,11 @@ let parse_object chan = | "_j" -> RInt63 (Uint63.of_int64 (input_intL chan)) | s -> Printf.eprintf "Unhandled custom code: %s" s; assert false end + | CODE_DOUBLE_BIG -> + RFloat64 (Float64.of_float (input_double_big chan)) + | CODE_DOUBLE_LITTLE -> + RFloat64 (Float64.of_float (input_double_little chan)) | CODE_DOUBLE_ARRAY32_LITTLE - | CODE_DOUBLE_BIG - | CODE_DOUBLE_LITTLE | CODE_DOUBLE_ARRAY8_BIG | CODE_DOUBLE_ARRAY8_LITTLE | CODE_DOUBLE_ARRAY32_BIG @@ -370,6 +393,11 @@ let parse chan = let () = LargeArray.set memory !current_object (Int63 i) in let () = incr current_object in data, None + | RFloat64 f -> + let data = Ptr !current_object in + let () = LargeArray.set memory !current_object (Float64 f) in + let () = incr current_object in + data, None in let rec fill block off accu = @@ -434,6 +462,7 @@ let instantiate (p, mem) = let obj = match LargeArray.get mem i with | Struct (tag, blk) -> Obj.new_block tag (Array.length blk) | Int63 i -> Obj.repr i + | Float64 f -> Obj.repr f | String str -> Obj.repr str in LargeArray.set ans i obj @@ -453,6 +482,7 @@ let instantiate (p, mem) = Obj.set_field obj k (get_data blk.(k)) done | Int63 _ + | Float64 _ | String _ -> () done; get_data p diff --git a/checker/analyze.mli b/checker/analyze.mli index 029f595959..e579f4896d 100644 --- a/checker/analyze.mli +++ b/checker/analyze.mli @@ -8,6 +8,7 @@ type data = type obj = | Struct of int * data array (* tag × data *) | Int63 of Uint63.t (* Primitive integer *) +| Float64 of Float64.t (* Primitive float *) | String of string module LargeArray : diff --git a/checker/validate.ml b/checker/validate.ml index 178bb4c527..678af9f2d5 100644 --- a/checker/validate.ml +++ b/checker/validate.ml @@ -87,6 +87,7 @@ let rec val_gen v ctx o = match v with | Dyn -> val_dyn ctx o | Proxy { contents = v } -> val_gen v ctx o | Uint63 -> val_uint63 ctx o + | Float64 -> val_float64 ctx o (* Check that an object is a tuple (or a record). vs is an array of value representation for each field. Its size corresponds to the @@ -138,6 +139,10 @@ and val_uint63 ctx o = if not (Uint63.is_uint63 o) then fail ctx o "not a 63-bit unsigned integer" +and val_float64 ctx o = + if not (Float64.is_float64 o) then + fail ctx o "not a 64-bit float" + let print_frame = function | CtxType t -> t | CtxAnnot t -> t diff --git a/checker/values.ml b/checker/values.ml index 9a2028a96b..3882f88391 100644 --- a/checker/values.ml +++ b/checker/values.ml @@ -35,6 +35,7 @@ type value = | Proxy of value ref | Uint63 + | Float64 let fix (f : value -> value) : value = let self = ref Any in @@ -147,7 +148,8 @@ let rec v_constr = [|v_fix|]; (* Fix *) [|v_cofix|]; (* CoFix *) [|v_proj;v_constr|]; (* Proj *) - [|Uint63|] (* Int *) + [|Uint63|]; (* Int *) + [|Float64|] (* Int *) |]) and v_prec = Tuple ("prec_declaration", @@ -226,7 +228,7 @@ let v_pol_arity = v_tuple "polymorphic_arity" [|List(Opt v_level);v_univ|] let v_primitive = - v_enum "primitive" 25 + v_enum "primitive" 44 (* Number of "Primitive" in Int63.v and PrimFloat.v *) let v_cst_def = v_sum "constant_def" 0 @@ -300,9 +302,11 @@ let v_ind_pack = v_tuple "mutual_inductive_body" Opt v_bool; v_typing_flags|] -let v_prim_ind = v_enum "prim_ind" 4 +let v_prim_ind = v_enum "prim_ind" 6 +(* Number of "Register ... as kernel.ind_..." in Int63.v and PrimFloat.v *) -let v_prim_type = v_enum "prim_type" 1 +let v_prim_type = v_enum "prim_type" 2 +(* Number of constructors of prim_type in "kernel/cPrimitives.ml" *) let v_retro_action = v_sum "retro_action" 0 [| diff --git a/checker/values.mli b/checker/values.mli index db6b0be250..ec3b91d5dd 100644 --- a/checker/values.mli +++ b/checker/values.mli @@ -39,6 +39,7 @@ type value = (** Same as the inner value, used to define recursive types *) | Uint63 + | Float64 (** NB: List and Opt have their own constructors to make it easy to define eg [let rec foo = List foo]. *) diff --git a/checker/votour.ml b/checker/votour.ml index 97584831e5..fe6c487496 100644 --- a/checker/votour.ml +++ b/checker/votour.ml @@ -101,6 +101,7 @@ struct in fold 0 1 (fun size -> let () = LargeArray.set !sizes p size in k size) | Int63 _ -> k 0 + | Float64 _ -> k 0 | String s -> let size = 2 + (String.length s / ws) in let () = LargeArray.set !sizes p size in @@ -118,6 +119,7 @@ struct match LargeArray.get !memory p with | Struct (tag, os) -> BLOCK (tag, os) | Int63 _ -> OTHER (* TODO: pretty-print int63 values *) + | Float64 _ -> OTHER (* TODO: pretty-print float64 values *) | String s -> STRING s let input ch = @@ -156,6 +158,7 @@ let rec get_name ?(extra=false) = function |Dyn -> "<dynamic>" | Proxy v -> get_name ~extra !v | Uint63 -> "Uint63" + | Float64 -> "Float64" (** For tuples, its quite handy to display the inner 1st string (if any). Cf. [structure_body] for instance *) @@ -261,6 +264,7 @@ let rec get_children v o pos = match v with |Fail s -> raise Forbidden | Proxy v -> get_children !v o pos | Uint63 -> raise Exit + | Float64 -> raise Exit let get_children v o pos = try get_children v o pos diff --git a/configure.ml b/configure.ml index 8e04dc46b2..a53292b4cf 100644 --- a/configure.ml +++ b/configure.ml @@ -456,8 +456,6 @@ let coq_bin_annot_flag = if !prefs.bin_annot then "-bin-annot" else "" let coq_safe_string = "-safe-string" let coq_strict_sequence = "-strict-sequence" -let cflags = "-Wall -Wno-unused -g -O2" - (** * Architecture *) let arch_progs = @@ -917,6 +915,40 @@ let configdir,configdirsuffix = let (_,_,d,s) = select "CONFIGDIR" in d,s let datadir,datadirsuffix = let (_,_,d,s) = select "DATADIR" in d,s +(** * CC runtime flags *) + +let cflags_dflt = "-Wall -Wno-unused -g -O2 -fexcess-precision=standard" + +let cflags_sse2 = ["-msse2"; "-mfpmath=sse"] + +let cflags, sse2_math = + let _, slurp = + (* Test SSE2_MATH support <https://stackoverflow.com/a/45667927> *) + tryrun "cc" (["-march=native"; "-dM"; "-E"] + @ cflags_sse2 + @ [coqtop/"kernel/byterun/coq_interp.h"] (* any file *)) in + if List.exists (fun line -> starts_with line "#define __SSE2_MATH__ 1") slurp + then (cflags_dflt ^ " " ^ String.concat " " cflags_sse2, true) + else (cflags_dflt, false) + +(** Test at configure time that no harmful double rounding seems to + be performed with an intermediate 80-bit representation (x87). + + If this test fails but SSE2_MATH is available, the build can go + further as Coq's primitive floats will use it through a dedicated + external C implementation (instead of relying on OCaml operations) + + If this test fails and SSE2_MATH is not available, abort. + *) +let () = + let add = (+.) in + let b = ldexp 1. 53 in + let s = add 1. (ldexp 1. (-52)) in + if (add b s <= b || add b 1. <> b || ldexp 1. (-1074) <= 0.) + && not sse2_math then + die "Detected non IEEE-754 compliant architecture (or wrong \ + rounding mode). Use of Float is thus unsafe." + (** * OCaml runtime flags *) (** Do we use -custom (yes by default on Windows and MacOS) *) @@ -1176,6 +1208,16 @@ let write_makefile f = let _ = write_makefile "config/Makefile" +let write_dune_c_flags f = + safe_remove f; + let o = open_out f in + let pr s = fprintf o s in + pr "(%s)\n" cflags; + close_out o; + Unix.chmod f 0o444 + +let _ = try write_dune_c_flags "kernel/byterun/dune.c_flags" with _ -> () + let write_macos_metadata exec = let f = "config/Info-"^exec^".plist" in let () = safe_remove f in diff --git a/dev/ci/user-overlays/09867-primitive-floats.sh b/dev/ci/user-overlays/09867-primitive-floats.sh new file mode 100644 index 0000000000..a0e9085afd --- /dev/null +++ b/dev/ci/user-overlays/09867-primitive-floats.sh @@ -0,0 +1,12 @@ +if [ "$CI_PULL_REQUEST" = "9867" ] || [ "$CI_BRANCH" = "primitive-floats" ]; then + + unicoq_CI_REF=primitive-floats + unicoq_CI_GITURL=https://github.com/validsdp/unicoq + + elpi_CI_REF=primitive-floats + elpi_CI_GITURL=https://github.com/validsdp/coq-elpi + + coqhammer_CI_REF=primitive-floats + coqhammer_CI_GITURL=https://github.com/validsdp/coqhammer + +fi diff --git a/dev/top_printers.ml b/dev/top_printers.ml index aa28bce018..ccb8658eee 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -307,6 +307,8 @@ let constr_display csr = ^(array_display bl)^")" | Int i -> "Int("^(Uint63.to_string i)^")" + | Float f -> + "Float("^(Float64.to_string f)^")" and array_display v = "[|"^ @@ -439,6 +441,8 @@ let print_pure_constr csr = in print_string"{"; print_fix (); print_string"}" | Int i -> print_string ("Int("^(Uint63.to_string i)^")") + | Float f -> + print_string ("Float("^(Float64.to_string f)^")") and box_display c = open_hovbox 1; term_display c; close_box() diff --git a/dev/vm_printers.ml b/dev/vm_printers.ml index 863d930968..11565b99eb 100644 --- a/dev/vm_printers.ml +++ b/dev/vm_printers.ml @@ -84,6 +84,7 @@ and ppwhd whd = | Vconstr_const i -> print_string "C(";print_int i;print_string")" | Vconstr_block b -> ppvblock b | Vint64 i -> printf "int64(%LiL)" i + | Vfloat64 f -> printf "float64(%.17g)" f | Vatom_stk(a,s) -> open_hbox();ppatom a;close_box(); print_string"@";ppstack s diff --git a/doc/changelog/01-kernel/09867-floats.rst b/doc/changelog/01-kernel/09867-floats.rst new file mode 100644 index 0000000000..56b5fc747a --- /dev/null +++ b/doc/changelog/01-kernel/09867-floats.rst @@ -0,0 +1,13 @@ +- A built-in support of floating-point arithmetic was added, allowing + one to devise efficient reflection tactics involving numerical + computation. Primitive floats are added in the language of terms, + following the binary64 format of the IEEE 754 standard, and the + related operations are implemented for the different reduction + engines of Coq by using the corresponding processor operators in + rounding-to-nearest-even. The properties of these operators are + axiomatized in the theory :g:`Coq.Floats.FloatAxioms` which is part + of the library :g:`Coq.Floats.Floats`. + See Section :ref:`primitive-floats` + (`#9867 <https://github.com/coq/coq/pull/9867>`_, + closes `#8276 <https://github.com/coq/coq/issues/8276>`_, + by Guillaume Bertholon, Erik Martin-Dorel, Pierre Roux). diff --git a/doc/sphinx/language/coq-library.rst b/doc/sphinx/language/coq-library.rst index ac75240cfb..cad5e4e67e 100644 --- a/doc/sphinx/language/coq-library.rst +++ b/doc/sphinx/language/coq-library.rst @@ -756,6 +756,7 @@ subdirectories: * **Sets** : Sets (classical, constructive, finite, infinite, power set, etc.) * **FSets** : Specification and implementations of finite sets and finite maps (by lists and by AVL trees) * **Reals** : Axiomatization of real numbers (classical, basic functions, integer part, fractional part, limit, derivative, Cauchy series, power series and results,...) + * **Floats** : Machine implementation of floating-point arithmetic (for the binary64 format) * **Relations** : Relations (definitions and basic results) * **Sorting** : Sorted list (basic definitions and heapsort correctness) * **Strings** : 8-bits characters and strings @@ -768,7 +769,7 @@ are directly accessible with the command ``Require`` (see Section :ref:`compiled-files`). The different modules of the |Coq| standard library are documented -online at http://coq.inria.fr/stdlib. +online at https://coq.inria.fr/stdlib. Peano’s arithmetic (nat) ~~~~~~~~~~~~~~~~~~~~~~~~ @@ -988,6 +989,106 @@ Notation Interpretation Precedence Associativity ``_ :: _`` ``cons`` 60 right ========== ============== ========== ============= +.. _floats_library: + +Floats library +~~~~~~~~~~~~~~ + +The library of primitive floating-point arithmetic can be loaded by +requiring module ``Floats``: + +.. coqtop:: in + + Require Import Floats. + +It exports the module ``PrimFloat`` that provides a primitive type +named ``float``, defined in the kernel (see section :ref:`primitive-floats`), +as well as two variant types ``float_comparison`` and ``float_class``: + + +.. coqtop:: all + + Print float. + Print float_comparison. + Print float_class. + +It then defines the primitive operators below, using the processor +floating-point operators for binary64 in rounding-to-nearest even: + +* ``abs`` +* ``opp`` +* ``sub`` +* ``add`` +* ``mul`` +* ``div`` +* ``sqrt`` +* ``compare`` : compare two floats and return a ``float_comparison`` +* ``classify`` : analyze a float and return a ``float_class`` +* ``of_int63`` : round a primitive integer and convert it into a float +* ``normfr_mantissa`` : take a float in ``[0.5; 1.0)`` and return its mantissa +* ``frshiftexp`` : convert a float to fractional part in ``[0.5; 1.0)`` and integer part +* ``ldshiftexp`` : multiply a float by an integral power of ``2`` +* ``next_up`` : return the next float towards positive infinity +* ``next_down`` : return the next float towards negative infinity + +For special floating-point values, the following constants are also +defined: + +* ``zero`` +* ``neg_zero`` +* ``one`` +* ``two`` +* ``infinity`` +* ``neg_infinity`` +* ``nan`` : Not a Number (assumed to be unique: the "payload" of NaNs is ignored) + +The following table shows the notations available when opening scope +``float_scope``. + +=========== ============== +Notation Interpretation +=========== ============== +``- _`` ``opp`` +``_ - _`` ``sub`` +``_ + _`` ``add`` +``_ * _`` ``mul`` +``_ / _`` ``div`` +``_ == _`` ``eqb`` +``_ < _`` ``ltb`` +``_ <= _`` ``leb`` +``_ ?= _`` ``compare`` +=========== ============== + +Floating-point constants are parsed and pretty-printed as (17-digit) +decimal constants. This ensures that the composition +:math:`\text{parse} \circ \text{print}` amounts to the identity. + +.. example:: + + .. coqtop:: all + + Open Scope float_scope. + Eval compute in 1 + 0.5. + Eval compute in 1 / 0. + Eval compute in 1 / -0. + Eval compute in 0 / 0. + Eval compute in 0 ?= -0. + Eval compute in nan ?= nan. + Eval compute in next_down (-1). + +The primitive operators are specified with respect to their Gallina +counterpart, using the variant type ``spec_float``, and the injection +``Prim2SF``: + +.. coqtop:: all + + Print spec_float. + Check Prim2SF. + Check mul_spec. + +For more details on the available definitions and lemmas, see the +online documentation of the ``Floats`` library. + .. _userscontributions: Users’ contributions diff --git a/doc/sphinx/language/gallina-extensions.rst b/doc/sphinx/language/gallina-extensions.rst index f50cf9340c..54669534c7 100644 --- a/doc/sphinx/language/gallina-extensions.rst +++ b/doc/sphinx/language/gallina-extensions.rst @@ -2411,7 +2411,7 @@ by means of the interactive proof engine. .. _primitive-integers: Primitive Integers --------------------------------- +------------------ The language of terms features 63-bit machine integers as values. The type of such a value is *axiomatized*; it is declared through the following sentence @@ -2464,6 +2464,55 @@ wrapped into the :g:`Uint63.of_int` (resp. :g:`Uint63.of_int64`) constructor on 64-bit (resp. 32-bit) platforms. Currently, this cannot be customized (see the function :g:`Uint63.compile` from the kernel). +.. _primitive-floats: + +Primitive Floats +---------------- + +The language of terms features Binary64 floating-point numbers as values. +The type of such a value is *axiomatized*; it is declared through the +following sentence (excerpt from the :g:`PrimFloat` module): + +.. coqdoc:: + + Primitive float := #float64_type. + +This type is equipped with a few operators, that must be similarly declared. +For instance, the product of two primitive floats can be computed using the +:g:`PrimFloat.mul` function, declared and specified as follows: + +.. coqdoc:: + + Primitive mul := #float64_mul. + Notation "x * y" := (mul x y) : float_scope. + + Axiom mul_spec : forall x y, Prim2SF (x * y)%float = SF64mul (Prim2SF x) (Prim2SF y). + +where :g:`Prim2SF` is defined in the :g:`FloatOps` module. + +The set of such operators is described in section :ref:`floats_library`. + +These primitive declarations are regular axioms. As such, they must be trusted, and are listed by the +:g:`Print Assumptions` command. + +The reduction machines (:tacn:`vm_compute`, :tacn:`native_compute`) implement +dedicated, efficient rules to reduce the applications of these primitive +operations, using the floating-point processor operators that are assumed +to comply with the IEEE 754 standard for floating-point arithmetic. + +The extraction of these primitives can be customized similarly to the extraction +of regular axioms (see :ref:`extraction`). Nonetheless, the :g:`ExtrOCamlFloats` +module can be used when extracting to OCaml: it maps the Coq primitives to types +and functions of a :g:`Float64` module. Said OCaml module is not produced by +extraction. Instead, it has to be provided by the user (if they want to compile +or execute the extracted code). For instance, an implementation of this module +can be taken from the kernel of Coq. + +Literal values (of type :g:`Float64.t`) are extracted to literal OCaml +values (of type :g:`float`) written in hexadecimal notation and +wrapped into the :g:`Float64.of_float` constructor, e.g.: +:g:`Float64.of_float (0x1p+0)`. + Bidirectionality hints ---------------------- diff --git a/doc/stdlib/hidden-files b/doc/stdlib/hidden-files index bb6df87970..5b243c8a9e 100644 --- a/doc/stdlib/hidden-files +++ b/doc/stdlib/hidden-files @@ -13,6 +13,7 @@ plugins/extraction/ExtrHaskellZNum.v plugins/extraction/ExtrOcamlBasic.v plugins/extraction/ExtrOcamlBigIntConv.v plugins/extraction/ExtrOCamlInt63.v +plugins/extraction/ExtrOCamlFloats.v plugins/extraction/ExtrOcamlIntConv.v plugins/extraction/ExtrOcamlNatBigInt.v plugins/extraction/ExtrOcamlNatInt.v diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template index f0ada745e7..eedd8a3d61 100644 --- a/doc/stdlib/index-list.html.template +++ b/doc/stdlib/index-list.html.template @@ -328,6 +328,19 @@ through the <tt>Require Import</tt> command.</p> theories/Numbers/Integer/Binary/ZBinary.v theories/Numbers/Integer/NatPairs/ZNatPairs.v </dd> + + <dt> <b> Floats</b>: + Floating-point arithmetic + </dt> + <dd> + theories/Floats/FloatClass.v + theories/Floats/PrimFloat.v + theories/Floats/SpecFloat.v + theories/Floats/FloatOps.v + theories/Floats/FloatAxioms.v + theories/Floats/FloatLemmas.v + (theories/Floats/Floats.v) + </dd> </dl> </dd> diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 23d066df58..46a80239cf 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -76,6 +76,7 @@ let mkProj (p, c) = of_kind (Proj (p, c)) let mkArrow t1 r t2 = of_kind (Prod (make_annot Anonymous r, t1, t2)) let mkArrowR t1 t2 = mkArrow t1 Sorts.Relevant t2 let mkInt i = of_kind (Int i) +let mkFloat f = of_kind (Float f) let mkRef (gr,u) = let open GlobRef in match gr with | ConstRef c -> mkConstU (c,u) @@ -334,7 +335,7 @@ let iter_with_full_binders sigma g f n c = let open Context.Rel.Declaration in match kind sigma c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> () + | Construct _ | Int _ | Float _) -> () | Cast (c,_,t) -> f n c; f n t | Prod (na,t,c) -> f n t; f (g (LocalAssum (na, t)) n) c | Lambda (na,t,c) -> f n t; f (g (LocalAssum (na, t)) n) c diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 2afce38db7..90f50b764c 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -127,6 +127,7 @@ val mkCoFix : (t, t) pcofixpoint -> t val mkArrow : t -> Sorts.relevance -> t -> t val mkArrowR : t -> t -> t val mkInt : Uint63.t -> t +val mkFloat : Float64.t -> t val mkRef : GlobRef.t * EInstance.t -> t diff --git a/engine/namegen.ml b/engine/namegen.ml index 89c2fade62..b850f38b4d 100644 --- a/engine/namegen.ml +++ b/engine/namegen.ml @@ -118,7 +118,7 @@ let head_name sigma c = (* Find the head constant of a constr if any *) Some (Nametab.basename_of_global (global_of_constr c)) | Fix ((_,i),(lna,_,_)) | CoFix (i,(lna,_,_)) -> Some (match lna.(i).binder_name with Name id -> id | _ -> assert false) - | Sort _ | Rel _ | Meta _|Evar _|Case (_, _, _, _) | Int _ -> None + | Sort _ | Rel _ | Meta _|Evar _|Case (_, _, _, _) | Int _ | Float _ -> None in hdrec c @@ -165,6 +165,7 @@ let hdchar env sigma c = | Evar _ (* We could do better... *) | Meta _ | Case (_, _, _, _) -> "y" | Int _ -> "i" + | Float _ -> "f" in hdrec 0 c diff --git a/engine/termops.ml b/engine/termops.ml index 2ab2f60421..90fa8546ce 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -600,7 +600,7 @@ let map_constr_with_binders_left_to_right sigma g f l c = let open EConstr in match EConstr.kind sigma c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> c + | Construct _ | Int _ | Float _) -> c | Cast (b,k,t) -> let b' = f l b in let t' = f l t in @@ -681,7 +681,7 @@ let map_constr_with_full_binders_gen userview sigma g f l cstr = let open EConstr in match EConstr.kind sigma cstr with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> cstr + | Construct _ | Int _ | Float _) -> cstr | Cast (c,k, t) -> let c' = f l c in let t' = f l t in diff --git a/interp/constrextern.ml b/interp/constrextern.ml index 217381d854..0a1371413a 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -752,6 +752,30 @@ let extended_glob_local_binder_of_decl loc = function let extended_glob_local_binder_of_decl ?loc u = DAst.make ?loc (extended_glob_local_binder_of_decl loc u) (**********************************************************************) +(* mapping special floats *) + +(* this helper function is copied from notation.ml as it's not exported *) +let qualid_of_ref n = + n |> Coqlib.lib_ref |> Nametab.shortest_qualid_of_global Id.Set.empty + +let q_infinity () = qualid_of_ref "num.float.infinity" +let q_neg_infinity () = qualid_of_ref "num.float.neg_infinity" +let q_nan () = qualid_of_ref "num.float.nan" + +let extern_float f scopes = + if Float64.is_nan f then CRef(q_nan (), None) + else if Float64.is_infinity f then CRef(q_infinity (), None) + else if Float64.is_neg_infinity f then CRef(q_neg_infinity (), None) + else + let sign = if Float64.sign f then SMinus else SPlus in + let s = Float64.(to_string (abs f)) in + match NumTok.of_string s with + | None -> assert false + | Some n -> + extern_prim_token_delimiter_if_required (Numeral (sign, n)) + "float" "float_scope" scopes + +(**********************************************************************) (* mapping glob_constr to constr_expr *) let extern_glob_sort = function @@ -972,6 +996,8 @@ let rec extern inctx scopes vars r = (Numeral (SPlus, NumTok.int (Uint63.to_string i))) "int63" "int63_scope" (snd scopes) + | GFloat f -> extern_float f (snd scopes) + in insert_coercion coercion (CAst.make ?loc c) and extern_typ (subentry,(_,scopes)) = @@ -1314,6 +1340,7 @@ let rec glob_of_pat avoid env sigma pat = DAst.make @@ match pat with | PSort Sorts.InSet -> GSort (UNamed [GSet,0]) | PSort Sorts.InType -> GSort (UAnonymous {rigid=true}) | PInt i -> GInt i + | PFloat f -> GFloat f let extern_constr_pattern env sigma pat = extern true (InConstrEntrySomeLevel,(None,[])) Id.Set.empty (glob_of_pat Id.Set.empty env sigma pat) diff --git a/interp/impargs.ml b/interp/impargs.ml index 5f41c2a366..0de4eb5fa1 100644 --- a/interp/impargs.ml +++ b/interp/impargs.ml @@ -216,7 +216,7 @@ let rec is_rigid_head sigma t = match kind sigma t with | Fix ((fi,i),_) -> is_rigid_head sigma (args.(fi.(i))) | _ -> is_rigid_head sigma f) | Lambda _ | LetIn _ | Construct _ | CoFix _ | Fix _ - | Prod _ | Meta _ | Cast _ | Int _ -> assert false + | Prod _ | Meta _ | Cast _ | Int _ | Float _ -> assert false let is_rigid env sigma t = let open Context.Rel.Declaration in diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index f30a874426..7e146754b2 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -90,9 +90,11 @@ let rec eq_notation_constr (vars1,vars2 as vars) t1 t2 = match t1, t2 with (eq_notation_constr vars) t1 t2 && cast_type_eq (eq_notation_constr vars) c1 c2 | NInt i1, NInt i2 -> Uint63.equal i1 i2 +| NFloat f1, NFloat f2 -> + Float64.equal f1 f2 | (NRef _ | NVar _ | NApp _ | NHole _ | NList _ | NLambda _ | NProd _ | NBinderList _ | NLetIn _ | NCases _ | NLetTuple _ | NIf _ - | NRec _ | NSort _ | NCast _ | NInt _), _ -> false + | NRec _ | NSort _ | NCast _ | NInt _ | NFloat _), _ -> false (**********************************************************************) (* Re-interpret a notation as a glob_constr, taking care of binders *) @@ -222,6 +224,7 @@ let glob_constr_of_notation_constr_with_binders ?loc g f e nc = | NHole (x, naming, arg) -> GHole (x, naming, arg) | NRef x -> GRef (x,None) | NInt i -> GInt i + | NFloat f -> GFloat f let glob_constr_of_notation_constr ?loc x = let rec aux () x = @@ -438,6 +441,7 @@ let notation_constr_and_vars_of_glob_constr recvars a = | GCast (c,k) -> NCast (aux c,map_cast_type aux k) | GSort s -> NSort s | GInt i -> NInt i + | GFloat f -> NFloat f | GHole (w,naming,arg) -> if arg != None then has_ltac := true; NHole (w, naming, arg) @@ -627,6 +631,7 @@ let rec subst_notation_constr subst bound raw = | NSort _ -> raw | NInt _ -> raw + | NFloat _ -> raw | NHole (knd, naming, solve) -> let nknd = match knd with @@ -1196,6 +1201,7 @@ let rec match_ inner u alp metas sigma a1 a2 = | GSort s1, NSort s2 when glob_sort_eq s1 s2 -> sigma | GInt i1, NInt i2 when Uint63.equal i1 i2 -> sigma + | GFloat f1, NFloat f2 when Float64.equal f1 f2 -> sigma | GPatVar _, NHole _ -> (*Don't hide Metas, they bind in ltac*) raise No_match | a, NHole _ -> sigma @@ -1223,7 +1229,7 @@ let rec match_ inner u alp metas sigma a1 a2 = | (GRef _ | GVar _ | GEvar _ | GPatVar _ | GApp _ | GLambda _ | GProd _ | GLetIn _ | GCases _ | GLetTuple _ | GIf _ | GRec _ | GSort _ | GHole _ - | GCast _ | GInt _ ), _ -> raise No_match + | GCast _ | GInt _ | GFloat _), _ -> raise No_match and match_in u = match_ true u diff --git a/interp/notation_term.ml b/interp/notation_term.ml index 908455bd05..c6ddd9ac95 100644 --- a/interp/notation_term.ml +++ b/interp/notation_term.ml @@ -44,6 +44,7 @@ type notation_constr = | NSort of glob_sort | NCast of notation_constr * notation_constr cast_type | NInt of Uint63.t + | NFloat of Float64.t (** Note concerning NList: first constr is iterator, second is terminator; first id is where each argument of the list has to be substituted diff --git a/kernel/byterun/coq_fix_code.c b/kernel/byterun/coq_fix_code.c index 0865487c98..931b509f48 100644 --- a/kernel/byterun/coq_fix_code.c +++ b/kernel/byterun/coq_fix_code.c @@ -44,6 +44,7 @@ void init_arity () { arity[PUSHCONST0]=arity[PUSHCONST1]=arity[PUSHCONST2]=arity[PUSHCONST3]= arity[ACCUMULATE]=arity[STOP]=arity[MAKEPROD]= arity[ADDINT63]=arity[SUBINT63]=arity[LTINT63]=arity[LEINT63]= + arity[LTFLOAT]=arity[LEFLOAT]= arity[ISINT]=arity[AREINT2]=0; /* instruction with one operand */ arity[ACC]=arity[PUSHACC]=arity[POP]=arity[ENVACC]=arity[PUSHENVACC]= @@ -63,7 +64,15 @@ void init_arity () { arity[CHECKLSLINT63]=arity[CHECKLSRINT63]=arity[CHECKADDMULDIVINT63]= arity[CHECKLSLINT63CONST1]=arity[CHECKLSRINT63CONST1]= arity[CHECKEQINT63]=arity[CHECKLTINT63]=arity[CHECKLEINT63]= - arity[CHECKCOMPAREINT63]=arity[CHECKHEAD0INT63]=arity[CHECKTAIL0INT63]=1; + arity[CHECKCOMPAREINT63]=arity[CHECKHEAD0INT63]=arity[CHECKTAIL0INT63]= + arity[CHECKEQFLOAT]=arity[CHECKLTFLOAT]=arity[CHECKLEFLOAT]= + arity[CHECKOPPFLOAT]=arity[CHECKABSFLOAT]=arity[CHECKCOMPAREFLOAT]= + arity[CHECKCLASSIFYFLOAT]= + arity[CHECKADDFLOAT]=arity[CHECKSUBFLOAT]=arity[CHECKMULFLOAT]= + arity[CHECKDIVFLOAT]=arity[CHECKSQRTFLOAT]= + arity[CHECKFLOATOFINT63]=arity[CHECKFLOATNORMFRMANTISSA]= + arity[CHECKFRSHIFTEXP]=arity[CHECKLDSHIFTEXP]= + arity[CHECKNEXTUPFLOAT]=arity[CHECKNEXTDOWNFLOAT]=1; /* instruction with two operands */ arity[APPTERM]=arity[MAKEBLOCK]=arity[CLOSURE]= arity[PROJ]=2; diff --git a/kernel/byterun/coq_float64.h b/kernel/byterun/coq_float64.h new file mode 100644 index 0000000000..c41079c8ff --- /dev/null +++ b/kernel/byterun/coq_float64.h @@ -0,0 +1,48 @@ +#ifndef _COQ_FLOAT64_ +#define _COQ_FLOAT64_ + +#include <math.h> + +#define DECLARE_FREL(name, e) \ + int coq_##name(double x, double y) { \ + return e; \ + } \ + \ + value coq_##name##_byte(value x, value y) { \ + return coq_##name(Double_val(x), Double_val(y)); \ + } + +#define DECLARE_FBINOP(name, e) \ + double coq_##name(double x, double y) { \ + return e; \ + } \ + \ + value coq_##name##_byte(value x, value y) { \ + return caml_copy_double(coq_##name(Double_val(x), Double_val(y))); \ + } + +#define DECLARE_FUNOP(name, e) \ + double coq_##name(double x) { \ + return e; \ + } \ + \ + value coq_##name##_byte(value x) { \ + return caml_copy_double(coq_##name(Double_val(x))); \ + } + +DECLARE_FREL(feq, x == y) +DECLARE_FREL(flt, x < y) +DECLARE_FREL(fle, x <= y) +DECLARE_FBINOP(fmul, x * y) +DECLARE_FBINOP(fadd, x + y) +DECLARE_FBINOP(fsub, x - y) +DECLARE_FBINOP(fdiv, x / y) +DECLARE_FUNOP(fsqrt, sqrt(x)) +DECLARE_FUNOP(next_up, nextafter(x, INFINITY)) +DECLARE_FUNOP(next_down, nextafter(x, -INFINITY)) + +value coq_is_double(value x) { + return Val_long(Is_double(x)); +} + +#endif /* _COQ_FLOAT64_ */ diff --git a/kernel/byterun/coq_interp.c b/kernel/byterun/coq_interp.c index 4b45608ae5..ca1308696c 100644 --- a/kernel/byterun/coq_interp.c +++ b/kernel/byterun/coq_interp.c @@ -17,11 +17,13 @@ #include <signal.h> #include <stdint.h> #include <caml/memory.h> +#include <math.h> #include "coq_gc.h" #include "coq_instruct.h" #include "coq_fix_code.h" #include "coq_memory.h" #include "coq_values.h" +#include "coq_float64.h" #ifdef ARCH_SIXTYFOUR #include "coq_uint63_native.h" @@ -167,38 +169,34 @@ if (sp - num_args < coq_stack_threshold) { \ #endif #endif -#define CheckInt1() do{ \ - if (Is_uint63(accu)) pc++; \ +#define CheckPrimArgs(cond, apply_lbl) do{ \ + if (cond) pc++; \ else{ \ *--sp=accu; \ accu = Field(coq_global_data, *pc++); \ - goto apply1; \ - } \ - }while(0) - -#define CheckInt2() do{ \ - if (Is_uint63(accu) && Is_uint63(sp[0])) pc++; \ - else{ \ - *--sp=accu; \ - accu = Field(coq_global_data, *pc++); \ - goto apply2; \ + goto apply_lbl; \ } \ }while(0) - - -#define CheckInt3() do{ \ - if (Is_uint63(accu) && Is_uint63(sp[0]) && Is_uint63(sp[1]) ) pc++; \ - else{ \ - *--sp=accu; \ - accu = Field(coq_global_data, *pc++); \ - goto apply3; \ - } \ - }while(0) +#define CheckInt1() CheckPrimArgs(Is_uint63(accu), apply1) +#define CheckInt2() CheckPrimArgs(Is_uint63(accu) && Is_uint63(sp[0]), apply2) +#define CheckInt3() CheckPrimArgs(Is_uint63(accu) && Is_uint63(sp[0]) \ + && Is_uint63(sp[1]), apply3) +#define CheckFloat1() CheckPrimArgs(Is_double(accu), apply1) +#define CheckFloat2() CheckPrimArgs(Is_double(accu) && Is_double(sp[0]), apply2) #define AllocCarry(cond) Alloc_small(accu, 1, (cond)? coq_tag_C1 : coq_tag_C0) #define AllocPair() Alloc_small(accu, 2, coq_tag_pair) +/* Beware: we cannot use caml_copy_double here as it doesn't use + Alloc_small, hence doesn't protect the stack via + Setup_for_gc/Restore_after_gc. */ +#define Coq_copy_double(val) do{ \ + double Coq_copy_double_f__ = (val); \ + Alloc_small(accu, Double_wosize, Double_tag); \ + Store_double_val(accu, Coq_copy_double_f__); \ + }while(0); + #define Swap_accu_sp do{ \ value swap_accu_sp_tmp__ = accu; \ accu = *sp; \ @@ -1533,6 +1531,206 @@ value coq_interprete } + Instruct (CHECKOPPFLOAT) { + print_instr("CHECKOPPFLOAT"); + CheckFloat1(); + Coq_copy_double(-Double_val(accu)); + Next; + } + + Instruct (CHECKABSFLOAT) { + print_instr("CHECKABSFLOAT"); + CheckFloat1(); + Coq_copy_double(fabs(Double_val(accu))); + Next; + } + + Instruct (CHECKEQFLOAT) { + print_instr("CHECKEQFLOAT"); + CheckFloat2(); + accu = coq_feq(Double_val(accu), Double_val(*sp++)) ? coq_true : coq_false; + Next; + } + + Instruct (CHECKLTFLOAT) { + print_instr("CHECKLTFLOAT"); + CheckFloat2(); + } + Instruct (LTFLOAT) { + print_instr("LTFLOAT"); + accu = coq_flt(Double_val(accu), Double_val(*sp++)) ? coq_true : coq_false; + Next; + } + + Instruct (CHECKLEFLOAT) { + print_instr("CHECKLEFLOAT"); + CheckFloat2(); + } + Instruct (LEFLOAT) { + print_instr("LEFLOAT"); + accu = coq_fle(Double_val(accu), Double_val(*sp++)) ? coq_true : coq_false; + Next; + } + + Instruct (CHECKCOMPAREFLOAT) { + double x, y; + print_instr("CHECKCOMPAREFLOAT"); + CheckFloat2(); + x = Double_val(accu); + y = Double_val(*sp++); + if(x < y) { + accu = coq_FLt; + } + else if(x > y) { + accu = coq_FGt; + } + else if(x == y) { + accu = coq_FEq; + } + else { // nan value + accu = coq_FNotComparable; + } + Next; + } + + Instruct (CHECKCLASSIFYFLOAT) { + double x; + print_instr("CHECKCLASSIFYFLOAT"); + CheckFloat1(); + x = Double_val(accu); + switch (fpclassify(x)) { + case FP_NORMAL: + accu = signbit(x) ? coq_NNormal : coq_PNormal; + break; + case FP_SUBNORMAL: + accu = signbit(x) ? coq_NSubn : coq_PSubn; + break; + case FP_ZERO: + accu = signbit(x) ? coq_NZero : coq_PZero; + break; + case FP_INFINITE: + accu = signbit(x) ? coq_NInf : coq_PInf; + break; + default: /* FP_NAN */ + accu = coq_NaN; + break; + } + Next; + } + + Instruct (CHECKADDFLOAT) { + print_instr("CHECKADDFLOAT"); + CheckFloat2(); + Coq_copy_double(coq_fadd(Double_val(accu), Double_val(*sp++))); + Next; + } + + Instruct (CHECKSUBFLOAT) { + print_instr("CHECKSUBFLOAT"); + CheckFloat2(); + Coq_copy_double(coq_fsub(Double_val(accu), Double_val(*sp++))); + Next; + } + + Instruct (CHECKMULFLOAT) { + print_instr("CHECKMULFLOAT"); + CheckFloat2(); + Coq_copy_double(coq_fmul(Double_val(accu), Double_val(*sp++))); + Next; + } + + Instruct (CHECKDIVFLOAT) { + print_instr("CHECKDIVFLOAT"); + CheckFloat2(); + Coq_copy_double(coq_fdiv(Double_val(accu), Double_val(*sp++))); + Next; + } + + Instruct (CHECKSQRTFLOAT) { + print_instr("CHECKSQRTFLOAT"); + CheckFloat1(); + Coq_copy_double(coq_fsqrt(Double_val(accu))); + Next; + } + + Instruct (CHECKFLOATOFINT63) { + print_instr("CHECKFLOATOFINT63"); + CheckInt1(); + Uint63_to_double(accu); + Next; + } + + Instruct (CHECKFLOATNORMFRMANTISSA) { + double f; + print_instr("CHECKFLOATNORMFRMANTISSA"); + CheckFloat1(); + f = fabs(Double_val(accu)); + if (f >= 0.5 && f < 1) { + Uint63_of_double(ldexp(f, DBL_MANT_DIG)); + } + else { + Uint63_of_int(Val_int(0)); + } + Next; + } + + Instruct (CHECKFRSHIFTEXP) { + int exp; + double f; + print_instr("CHECKFRSHIFTEXP"); + CheckFloat1(); + /* frexp(infinity) incorrectly returns nan on mingw */ +#if defined(__MINGW32__) || defined(__MINGW64__) + if (fpclassify(Double_val(accu)) == FP_INFINITE) { + f = Double_val(accu); + } else +#endif + f = frexp(Double_val(accu), &exp); + if (fpclassify(f) == FP_NORMAL) { + exp += FLOAT_EXP_SHIFT; + } + else { + exp = 0; + } + Coq_copy_double(f); + *--sp = accu; +#ifdef ARCH_SIXTYFOUR + Alloc_small(accu, 2, coq_tag_pair); + Field(accu, 1) = Val_int(exp); +#else + Uint63_of_int(Val_int(exp)); + *--sp = accu; + Alloc_small(accu, 2, coq_tag_pair); + Field(accu, 1) = *sp++; +#endif + Field(accu, 0) = *sp++; + Next; + } + + Instruct (CHECKLDSHIFTEXP) { + print_instr("CHECKLDSHIFTEXP"); + CheckPrimArgs(Is_double(accu) && Is_uint63(sp[0]), apply2); + Swap_accu_sp; + Uint63_to_int_min(accu, Val_int(2 * FLOAT_EXP_SHIFT)); + accu = Int_val(accu); + Coq_copy_double(ldexp(Double_val(*sp++), accu - FLOAT_EXP_SHIFT)); + Next; + } + + Instruct (CHECKNEXTUPFLOAT) { + print_instr("CHECKNEXTUPFLOAT"); + CheckFloat1(); + Coq_copy_double(coq_next_up(Double_val(accu))); + Next; + } + + Instruct (CHECKNEXTDOWNFLOAT) { + print_instr("CHECKNEXTDOWNFLOAT"); + CheckFloat1(); + Coq_copy_double(coq_next_down(Double_val(accu))); + Next; + } + /* Debugging and machine control */ Instruct(STOP){ diff --git a/kernel/byterun/coq_uint63_emul.h b/kernel/byterun/coq_uint63_emul.h index 528cc6fc1f..143a6d098c 100644 --- a/kernel/byterun/coq_uint63_emul.h +++ b/kernel/byterun/coq_uint63_emul.h @@ -156,3 +156,18 @@ DECLARE_BINOP(mulc_ml) *(h) = Field(uint63_return_value__, 0); \ accu = Field(uint63_return_value__, 1); \ }while(0) + +DECLARE_UNOP(to_float) +#define Uint63_to_double(x) CALL_UNOP(to_float, x) + +DECLARE_UNOP(of_float) +#define Uint63_of_double(f) do{ \ + Coq_copy_double(f); \ + CALL_UNOP(of_float, accu); \ + }while(0) + +DECLARE_UNOP(of_int) +#define Uint63_of_int(x) CALL_UNOP(of_int, x) + +DECLARE_BINOP(to_int_min) +#define Uint63_to_int_min(n, m) CALL_BINOP(to_int_min, n, m) diff --git a/kernel/byterun/coq_uint63_native.h b/kernel/byterun/coq_uint63_native.h index 9fbd3f83d8..5be7587091 100644 --- a/kernel/byterun/coq_uint63_native.h +++ b/kernel/byterun/coq_uint63_native.h @@ -138,3 +138,26 @@ value uint63_div21(value xh, value xl, value y, value* ql) { } } #define Uint63_div21(xh, xl, y, q) (accu = uint63_div21(xh, xl, y, q)) + +#define Uint63_to_double(x) Coq_copy_double((double) uint63_of_value(x)) + +double coq_uint63_to_float(value x) { + return (double) uint63_of_value(x); +} + +value coq_uint63_to_float_byte(value x) { + return caml_copy_double(coq_uint63_to_float(x)); +} + +#define Uint63_of_double(f) do{ \ + accu = Val_long((uint64_t)(f)); \ + }while(0) + +#define Uint63_of_int(x) (accu = (x)) + +#define Uint63_to_int_min(n, m) do { \ + if (uint63_lt((n),(m))) \ + accu = (n); \ + else \ + accu = (m); \ + }while(0) diff --git a/kernel/byterun/coq_values.h b/kernel/byterun/coq_values.h index 0cf6ccf532..b027673ac7 100644 --- a/kernel/byterun/coq_values.h +++ b/kernel/byterun/coq_values.h @@ -14,6 +14,8 @@ #include <caml/alloc.h> #include <caml/mlvalues.h> +#include <float.h> + #define Default_tag 0 #define Accu_tag 0 @@ -29,8 +31,9 @@ /* Les blocs accumulate */ #define Is_accu(v) (Is_block(v) && (Tag_val(v) == Accu_tag)) #define IS_EVALUATED_COFIX(v) (Is_accu(v) && Is_block(Field(v,1)) && (Tag_val(Field(v,1)) == ATOM_COFIXEVALUATED_TAG)) +#define Is_double(v) (Tag_val(v) == Double_tag) -/* */ +/* coq values for primitive operations */ #define coq_tag_C1 2 #define coq_tag_C0 1 #define coq_tag_pair 1 @@ -39,5 +42,20 @@ #define coq_Eq Val_int(0) #define coq_Lt Val_int(1) #define coq_Gt Val_int(2) +#define coq_FEq Val_int(0) +#define coq_FLt Val_int(1) +#define coq_FGt Val_int(2) +#define coq_FNotComparable Val_int(3) +#define coq_PNormal Val_int(0) +#define coq_NNormal Val_int(1) +#define coq_PSubn Val_int(2) +#define coq_NSubn Val_int(3) +#define coq_PZero Val_int(4) +#define coq_NZero Val_int(5) +#define coq_PInf Val_int(6) +#define coq_NInf Val_int(7) +#define coq_NaN Val_int(8) + +#define FLOAT_EXP_SHIFT (2101) /* 2*emax + prec */ #endif /* _COQ_VALUES_ */ diff --git a/kernel/byterun/dune b/kernel/byterun/dune index 20bdf28e54..d0145176ea 100644 --- a/kernel/byterun/dune +++ b/kernel/byterun/dune @@ -1,3 +1,16 @@ +; Dune doesn't use configure's output, but it is still necessary for +; some Coq files to work; will be fixed in the future. +(rule + (targets dune.c_flags) + (mode fallback) + (deps %{project_root}/configure.ml %{project_root}/dev/ocamldebug-coq.run (env_var COQ_CONFIGURE_PREFIX)) + (action (chdir %{project_root} (run %{ocaml} configure.ml -no-ask -native-compiler no)))) + +(env + (dev (c_flags (:include dune.c_flags))) + (release (c_flags (:include dune.c_flags))) + (ireport (c_flags (:include dune.c_flags)))) + (library (name byterun) (synopsis "Coq's Kernel Abstract Reduction Machine [C implementation]") diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml index 3fd613e905..72585e5014 100644 --- a/kernel/cClosure.ml +++ b/kernel/cClosure.ml @@ -348,6 +348,7 @@ and fterm = | FLetIn of Name.t Context.binder_annot * fconstr * fconstr * constr * fconstr subs | FEvar of existential * fconstr subs | FInt of Uint63.t + | FFloat of Float64.t | FLIFT of int * fconstr | FCLOS of constr * fconstr subs | FLOCKED @@ -428,7 +429,7 @@ let rec stack_args_size = function let rec lft_fconstr n ft = let r = Mark.relevance ft.mark in match ft.term with - | (FInd _|FConstruct _|FFlex(ConstKey _|VarKey _)|FInt _) -> ft + | (FInd _|FConstruct _|FFlex(ConstKey _|VarKey _)|FInt _|FFloat _) -> ft | FRel i -> {mark=mark Norm r;term=FRel(i+n)} | FLambda(k,tys,f,e) -> {mark=mark Cstr r; term=FLambda(k,tys,f,subs_shft(n,e))} | FFix(fx,e) -> @@ -499,6 +500,7 @@ let mk_clos e t = | Ind kn -> {mark = mark Norm KnownR; term = FInd kn } | Construct kn -> {mark = mark Cstr Unknown; term = FConstruct kn } | Int i -> {mark = mark Cstr Unknown; term = FInt i} + | Float f -> {mark = mark Cstr Unknown; term = FFloat f} | (CoFix _|Lambda _|Fix _|Prod _|Evar _|App _|Case _|Cast _|LetIn _|Proj _) -> {mark = mark Red Unknown; term = FCLOS(t,e)} @@ -616,6 +618,8 @@ let rec to_constr lfts v = | FInt i -> Constr.mkInt i + | FFloat f -> + Constr.mkFloat f | FCLOS (t,env) -> if is_subs_id env && is_lift_id lfts then t @@ -926,7 +930,7 @@ let rec knh info m stk = (* cases where knh stops *) | (FFlex _|FLetIn _|FConstruct _|FEvar _| - FCoFix _|FLambda _|FRel _|FAtom _|FInd _|FProd _|FInt _) -> + FCoFix _|FLambda _|FRel _|FAtom _|FInd _|FProd _|FInt _|FFloat _) -> (m, stk) (* The same for pure terms *) @@ -940,7 +944,7 @@ and knht info e t stk = | Cast(a,_,_) -> knht info e a stk | Rel n -> knh info (clos_rel e n) stk | Proj (p, c) -> knh info { mark = mark Red Unknown; term = FProj (p, mk_clos e c) } stk - | (Ind _|Const _|Construct _|Var _|Meta _ | Sort _ | Int _) -> (mk_clos e t, stk) + | (Ind _|Const _|Construct _|Var _|Meta _ | Sort _ | Int _|Float _) -> (mk_clos e t, stk) | CoFix cfx -> { mark = mark Cstr Unknown; term = FCoFix (cfx,e) }, stk | Lambda _ -> { mark = mark Cstr Unknown; term = mk_lambda e t }, stk | Prod (n, t, c) -> @@ -969,6 +973,11 @@ module FNativeEntries = | FInt i -> i | _ -> raise Primred.NativeDestKO + let get_float () e = + match [@ocaml.warning "-4"] e.term with + | FFloat f -> f + | _ -> raise Primred.NativeDestKO + let dummy = {mark = mark Norm KnownR; term = FRel 0} let current_retro = ref Retroknowledge.empty @@ -982,6 +991,16 @@ module FNativeEntries = fint := { mark = mark Norm KnownR; term = FFlex (ConstKey (Univ.in_punivs c)) } | None -> defined_int := false + let defined_float = ref false + let ffloat = ref dummy + + let init_float retro = + match retro.Retroknowledge.retro_float64 with + | Some c -> + defined_float := true; + ffloat := { mark = mark Norm KnownR; term = FFlex (ConstKey (Univ.in_punivs c)) } + | None -> defined_float := false + let defined_bool = ref false let ftrue = ref dummy let ffalse = ref dummy @@ -1020,6 +1039,7 @@ module FNativeEntries = let fEq = ref dummy let fLt = ref dummy let fGt = ref dummy + let fcmp = ref dummy let init_cmp retro = match retro.Retroknowledge.retro_cmp with @@ -1027,9 +1047,54 @@ module FNativeEntries = defined_cmp := true; fEq := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cEq) }; fLt := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cLt) }; - fGt := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cGt) } + fGt := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cGt) }; + let (icmp, _) = cEq in + fcmp := { mark = mark Norm KnownR; term = FInd (Univ.in_punivs icmp) } | None -> defined_cmp := false + let defined_f_cmp = ref false + let fFEq = ref dummy + let fFLt = ref dummy + let fFGt = ref dummy + let fFNotComparable = ref dummy + + let init_f_cmp retro = + match retro.Retroknowledge.retro_f_cmp with + | Some (cFEq, cFLt, cFGt, cFNotComparable) -> + defined_f_cmp := true; + fFEq := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cFEq) }; + fFLt := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cFLt) }; + fFGt := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cFGt) }; + fFNotComparable := + { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cFNotComparable) }; + | None -> defined_f_cmp := false + + let defined_f_class = ref false + let fPNormal = ref dummy + let fNNormal = ref dummy + let fPSubn = ref dummy + let fNSubn = ref dummy + let fPZero = ref dummy + let fNZero = ref dummy + let fPInf = ref dummy + let fNInf = ref dummy + let fNaN = ref dummy + + let init_f_class retro = + match retro.Retroknowledge.retro_f_class with + | Some (cPNormal, cNNormal, cPSubn, cNSubn, cPZero, cNZero, + cPInf, cNInf, cNaN) -> + defined_f_class := true; + fPNormal := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cPNormal) }; + fNNormal := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNNormal) }; + fPSubn := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cPSubn) }; + fNSubn := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNSubn) }; + fPZero := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cPZero) }; + fNZero := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNZero) }; + fPInf := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cPInf) }; + fNInf := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNInf) }; + fNaN := { mark = mark Cstr KnownR; term = FConstruct (Univ.in_punivs cNaN) }; + | None -> defined_f_class := false let defined_refl = ref false let frefl = ref dummy @@ -1044,10 +1109,13 @@ module FNativeEntries = let init env = current_retro := env.retroknowledge; init_int !current_retro; + init_float !current_retro; init_bool !current_retro; init_carry !current_retro; init_pair !current_retro; init_cmp !current_retro; + init_f_cmp !current_retro; + init_f_class !current_retro; init_refl !current_retro let check_env env = @@ -1057,6 +1125,10 @@ module FNativeEntries = check_env env; assert (!defined_int) + let check_float env = + check_env env; + assert (!defined_float) + let check_bool env = check_env env; assert (!defined_bool) @@ -1073,10 +1145,22 @@ module FNativeEntries = check_env env; assert (!defined_cmp) + let check_f_cmp env = + check_env env; + assert (!defined_f_cmp) + + let check_f_class env = + check_env env; + assert (!defined_f_class) + let mkInt env i = check_int env; { mark = mark Cstr KnownR; term = FInt i } + let mkFloat env f = + check_float env; + { mark = mark Norm KnownR; term = FFloat f } + let mkBool env b = check_bool env; if b then !ftrue else !ffalse @@ -1090,6 +1174,11 @@ module FNativeEntries = check_pair env; { mark = mark Cstr KnownR; term = FApp(!fPair, [|!fint;!fint;e1;e2|]) } + let mkFloatIntPair env f i = + check_pair env; + check_float env; + { mark = mark Cstr KnownR; term = FApp(!fPair, [|!ffloat;!fint;f;i|]) } + let mkLt env = check_cmp env; !fLt @@ -1102,6 +1191,57 @@ module FNativeEntries = check_cmp env; !fGt + let mkFLt env = + check_f_cmp env; + !fFLt + + let mkFEq env = + check_f_cmp env; + !fFEq + + let mkFGt env = + check_f_cmp env; + !fFGt + + let mkFNotComparable env = + check_f_cmp env; + !fFNotComparable + + let mkPNormal env = + check_f_class env; + !fPNormal + + let mkNNormal env = + check_f_class env; + !fNNormal + + let mkPSubn env = + check_f_class env; + !fPSubn + + let mkNSubn env = + check_f_class env; + !fNSubn + + let mkPZero env = + check_f_class env; + !fPZero + + let mkNZero env = + check_f_class env; + !fNZero + + let mkPInf env = + check_f_class env; + !fPInf + + let mkNInf env = + check_f_class env; + !fNInf + + let mkNaN env = + check_f_class env; + !fNaN end module FredNative = RedNative(FNativeEntries) @@ -1164,7 +1304,7 @@ let rec knr info tab m stk = (match info.i_cache.i_sigma ev with Some c -> knit info tab env c stk | None -> (m,stk)) - | FInt _ -> + | FInt _ | FFloat _ -> (match [@ocaml.warning "-4"] strip_update_shift_app m stk with | (_, _, Zprimitive(op,c,rargs,nargs)::s) -> let (rargs, nargs) = skip_native_args (m::rargs) nargs in @@ -1270,7 +1410,7 @@ and norm_head info tab m = | FProj (p,c) -> mkProj (p, kl info tab c) | FLOCKED | FRel _ | FAtom _ | FFlex _ | FInd _ | FConstruct _ - | FApp _ | FCaseT _ | FLIFT _ | FCLOS _ | FInt _ -> term_of_fconstr m + | FApp _ | FCaseT _ | FLIFT _ | FCLOS _ | FInt _ | FFloat _ -> term_of_fconstr m (* Initialization and then normalization *) diff --git a/kernel/cClosure.mli b/kernel/cClosure.mli index cd1de4c834..720f11b8f2 100644 --- a/kernel/cClosure.mli +++ b/kernel/cClosure.mli @@ -115,6 +115,7 @@ type fterm = | FLetIn of Name.t Context.binder_annot * fconstr * fconstr * constr * fconstr subs | FEvar of existential * fconstr subs | FInt of Uint63.t + | FFloat of Float64.t | FLIFT of int * fconstr | FCLOS of constr * fconstr subs | FLOCKED diff --git a/kernel/cPrimitives.ml b/kernel/cPrimitives.ml index d854cadd15..9ff7f69203 100644 --- a/kernel/cPrimitives.ml +++ b/kernel/cPrimitives.ml @@ -33,6 +33,24 @@ type t = | Int63lt | Int63le | Int63compare + | Float64opp + | Float64abs + | Float64eq + | Float64lt + | Float64le + | Float64compare + | Float64classify + | Float64add + | Float64sub + | Float64mul + | Float64div + | Float64sqrt + | Float64ofInt63 + | Float64normfr_mantissa + | Float64frshiftexp + | Float64ldshiftexp + | Float64next_up + | Float64next_down let equal (p1 : t) (p2 : t) = p1 == p2 @@ -62,6 +80,24 @@ let hash = function | Int63lt -> 22 | Int63le -> 23 | Int63compare -> 24 + | Float64opp -> 25 + | Float64abs -> 26 + | Float64compare -> 27 + | Float64classify -> 28 + | Float64add -> 29 + | Float64sub -> 30 + | Float64mul -> 31 + | Float64div -> 32 + | Float64sqrt -> 33 + | Float64ofInt63 -> 34 + | Float64normfr_mantissa -> 35 + | Float64frshiftexp -> 36 + | Float64ldshiftexp -> 37 + | Float64next_up -> 38 + | Float64next_down -> 39 + | Float64eq -> 40 + | Float64lt -> 41 + | Float64le -> 42 (* Should match names in nativevalues.ml *) let to_string = function @@ -89,6 +125,72 @@ let to_string = function | Int63lt -> "lt" | Int63le -> "le" | Int63compare -> "compare" + | Float64opp -> "fopp" + | Float64abs -> "fabs" + | Float64eq -> "feq" + | Float64lt -> "flt" + | Float64le -> "fle" + | Float64compare -> "fcompare" + | Float64classify -> "fclassify" + | Float64add -> "fadd" + | Float64sub -> "fsub" + | Float64mul -> "fmul" + | Float64div -> "fdiv" + | Float64sqrt -> "fsqrt" + | Float64ofInt63 -> "float_of_int" + | Float64normfr_mantissa -> "normfr_mantissa" + | Float64frshiftexp -> "frshiftexp" + | Float64ldshiftexp -> "ldshiftexp" + | Float64next_up -> "next_up" + | Float64next_down -> "next_down" + +type prim_type = + | PT_int63 + | PT_float64 + +type 'a prim_ind = + | PIT_bool : unit prim_ind + | PIT_carry : prim_type prim_ind + | PIT_pair : (prim_type * prim_type) prim_ind + | PIT_cmp : unit prim_ind + | PIT_f_cmp : unit prim_ind + | PIT_f_class : unit prim_ind + +type prim_ind_ex = PIE : 'a prim_ind -> prim_ind_ex + +type ind_or_type = + | PITT_ind : 'a prim_ind * 'a -> ind_or_type + | PITT_type : prim_type -> ind_or_type + +let types = + let int_ty = PITT_type PT_int63 in + let float_ty = PITT_type PT_float64 in + function + | Int63head0 | Int63tail0 -> [int_ty; int_ty] + | Int63add | Int63sub | Int63mul + | Int63div | Int63mod + | Int63lsr | Int63lsl + | Int63land | Int63lor | Int63lxor -> [int_ty; int_ty; int_ty] + | Int63addc | Int63subc | Int63addCarryC | Int63subCarryC -> + [int_ty; int_ty; PITT_ind (PIT_carry, PT_int63)] + | Int63mulc | Int63diveucl -> + [int_ty; int_ty; PITT_ind (PIT_pair, (PT_int63, PT_int63))] + | Int63eq | Int63lt | Int63le -> [int_ty; int_ty; PITT_ind (PIT_bool, ())] + | Int63compare -> [int_ty; int_ty; PITT_ind (PIT_cmp, ())] + | Int63div21 -> + [int_ty; int_ty; int_ty; PITT_ind (PIT_pair, (PT_int63, PT_int63))] + | Int63addMulDiv -> [int_ty; int_ty; int_ty; int_ty] + | Float64opp | Float64abs | Float64sqrt + | Float64next_up | Float64next_down -> [float_ty; float_ty] + | Float64ofInt63 -> [int_ty; float_ty] + | Float64normfr_mantissa -> [float_ty; int_ty] + | Float64frshiftexp -> [float_ty; PITT_ind (PIT_pair, (PT_float64, PT_int63))] + | Float64eq | Float64lt | Float64le -> [float_ty; float_ty; PITT_ind (PIT_bool, ())] + | Float64compare -> [float_ty; float_ty; PITT_ind (PIT_f_cmp, ())] + | Float64classify -> [float_ty; PITT_ind (PIT_f_class, ())] + | Float64add | Float64sub | Float64mul + | Float64div -> [float_ty; float_ty; float_ty] + | Float64ldshiftexp -> [float_ty; int_ty; float_ty] type arg_kind = | Kparam (* not needed for the evaluation of the primitive when it reduces *) @@ -97,58 +199,32 @@ type arg_kind = type args_red = arg_kind list -(* Invariant only argument of type int63 or an inductive can +(* Invariant only argument of type int63, float or an inductive can have kind Kwhnf *) -let kind = function - | Int63head0 | Int63tail0 -> [Kwhnf] - - | Int63add | Int63sub | Int63mul - | Int63div | Int63mod - | Int63lsr | Int63lsl - | Int63land | Int63lor | Int63lxor - | Int63addc | Int63subc - | Int63addCarryC | Int63subCarryC | Int63mulc | Int63diveucl - | Int63eq | Int63lt | Int63le | Int63compare -> [Kwhnf; Kwhnf] +let arity t = List.length (types t) - 1 - | Int63div21 | Int63addMulDiv -> [Kwhnf; Kwhnf; Kwhnf] - -let arity = function - | Int63head0 | Int63tail0 -> 1 - | Int63add | Int63sub | Int63mul - | Int63div | Int63mod - | Int63lsr | Int63lsl - | Int63land | Int63lor | Int63lxor - | Int63addc | Int63subc - | Int63addCarryC | Int63subCarryC | Int63mulc | Int63diveucl - | Int63eq | Int63lt | Int63le - | Int63compare -> 2 - - | Int63div21 | Int63addMulDiv -> 3 +let kind t = + let rec aux n = if n <= 0 then [] else Kwhnf :: aux (n - 1) in + aux (arity t) (** Special Entries for Register **) -type prim_ind = - | PIT_bool - | PIT_carry - | PIT_pair - | PIT_cmp - -type prim_type = - | PT_int63 - type op_or_type = | OT_op of t | OT_type of prim_type -let prim_ind_to_string = function +let prim_ind_to_string (type a) (p : a prim_ind) = match p with | PIT_bool -> "bool" | PIT_carry -> "carry" | PIT_pair -> "pair" | PIT_cmp -> "cmp" + | PIT_f_cmp -> "f_cmp" + | PIT_f_class -> "f_class" let prim_type_to_string = function | PT_int63 -> "int63_type" + | PT_float64 -> "float64_type" let op_or_type_to_string = function | OT_op op -> to_string op diff --git a/kernel/cPrimitives.mli b/kernel/cPrimitives.mli index 6913371caf..be65ba5305 100644 --- a/kernel/cPrimitives.mli +++ b/kernel/cPrimitives.mli @@ -33,6 +33,24 @@ type t = | Int63lt | Int63le | Int63compare + | Float64opp + | Float64abs + | Float64eq + | Float64lt + | Float64le + | Float64compare + | Float64classify + | Float64add + | Float64sub + | Float64mul + | Float64div + | Float64sqrt + | Float64ofInt63 + | Float64normfr_mantissa + | Float64frshiftexp + | Float64ldshiftexp + | Float64next_up + | Float64next_down val equal : t -> t -> bool @@ -53,18 +71,29 @@ val kind : t -> args_red (** Special Entries for Register **) -type prim_ind = - | PIT_bool - | PIT_carry - | PIT_pair - | PIT_cmp - type prim_type = | PT_int63 + | PT_float64 + +type 'a prim_ind = + | PIT_bool : unit prim_ind + | PIT_carry : prim_type prim_ind + | PIT_pair : (prim_type * prim_type) prim_ind + | PIT_cmp : unit prim_ind + | PIT_f_cmp : unit prim_ind + | PIT_f_class : unit prim_ind + +type prim_ind_ex = PIE : 'a prim_ind -> prim_ind_ex type op_or_type = | OT_op of t | OT_type of prim_type -val prim_ind_to_string : prim_ind -> string +val prim_ind_to_string : 'a prim_ind -> string val op_or_type_to_string : op_or_type -> string + +type ind_or_type = + | PITT_ind : 'a prim_ind * 'a -> ind_or_type + | PITT_type : prim_type -> ind_or_type + +val types : t -> ind_or_type list diff --git a/kernel/cbytegen.ml b/kernel/cbytegen.ml index 83d2a58d83..13cc6f7ea4 100644 --- a/kernel/cbytegen.ml +++ b/kernel/cbytegen.ml @@ -528,6 +528,8 @@ let rec compile_lam env cenv lam sz cont = | Luint i -> compile_structured_constant cenv (Const_uint i) sz cont + | Lfloat f -> compile_structured_constant cenv (Const_float f) sz cont + | Lproj (p,arg) -> compile_lam env cenv arg sz (Kproj p :: cont) diff --git a/kernel/cemitcodes.ml b/kernel/cemitcodes.ml index 76e2515ea7..5e82cef810 100644 --- a/kernel/cemitcodes.ml +++ b/kernel/cemitcodes.ml @@ -234,6 +234,24 @@ let check_prim_op = function | Int63lt -> opCHECKLTINT63 | Int63le -> opCHECKLEINT63 | Int63compare -> opCHECKCOMPAREINT63 + | Float64opp -> opCHECKOPPFLOAT + | Float64abs -> opCHECKABSFLOAT + | Float64eq -> opCHECKEQFLOAT + | Float64lt -> opCHECKLTFLOAT + | Float64le -> opCHECKLEFLOAT + | Float64compare -> opCHECKCOMPAREFLOAT + | Float64classify -> opCHECKCLASSIFYFLOAT + | Float64add -> opCHECKADDFLOAT + | Float64sub -> opCHECKSUBFLOAT + | Float64mul -> opCHECKMULFLOAT + | Float64div -> opCHECKDIVFLOAT + | Float64sqrt -> opCHECKSQRTFLOAT + | Float64ofInt63 -> opCHECKFLOATOFINT63 + | Float64normfr_mantissa -> opCHECKFLOATNORMFRMANTISSA + | Float64frshiftexp -> opCHECKFRSHIFTEXP + | Float64ldshiftexp -> opCHECKLDSHIFTEXP + | Float64next_up -> opCHECKNEXTUPFLOAT + | Float64next_down -> opCHECKNEXTDOWNFLOAT let emit_instr env = function | Klabel lbl -> define_label env lbl @@ -384,7 +402,8 @@ type to_patch = emitcodes * patches * fv (* Substitution *) let subst_strcst s sc = match sc with - | Const_sort _ | Const_b0 _ | Const_univ_level _ | Const_val _ | Const_uint _ -> sc + | Const_sort _ | Const_b0 _ | Const_univ_level _ | Const_val _ | Const_uint _ + | Const_float _ -> sc | Const_ind ind -> let kn,i = ind in Const_ind (subst_mind s kn, i) let subst_reloc s ri = diff --git a/kernel/clambda.ml b/kernel/clambda.ml index a764cca354..8c7aa6b17a 100644 --- a/kernel/clambda.ml +++ b/kernel/clambda.ml @@ -28,6 +28,7 @@ type lambda = | Lint of int | Lmakeblock of int * lambda array | Luint of Uint63.t + | Lfloat of Float64.t | Lval of structured_values | Lsort of Sorts.t | Lind of pinductive @@ -143,6 +144,7 @@ let rec pp_lam lam = prlist_with_sep spc pp_lam (Array.to_list args) ++ str")") | Luint i -> str (Uint63.to_string i) + | Lfloat f -> str (Float64.to_string f) | Lval _ -> str "values" | Lsort s -> pp_sort s | Lind ((mind,i), _) -> MutInd.print mind ++ str"#" ++ int i @@ -195,7 +197,8 @@ let shift subst = subs_shft (1, subst) let map_lam_with_binders g f n lam = match lam with - | Lrel _ | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Lint _ | Luint _ -> lam + | Lrel _ | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Lint _ | Luint _ + | Lfloat _ -> lam | Levar (evk, args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Levar (evk, args') @@ -416,7 +419,8 @@ let rec occurrence k kind lam = if n = k then if kind then false else raise Not_found else kind - | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Lint _ | Luint _ -> kind + | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Lint _ | Luint _ + | Lfloat _ -> kind | Levar (_, args) -> occurrence_args k kind args | Lprod(dom, codom) -> @@ -763,6 +767,7 @@ let rec lambda_of_constr env c = Lproj (Projection.repr p,lc) | Int i -> Luint i + | Float f -> Lfloat f and lambda_of_app env f args = match Constr.kind f with diff --git a/kernel/clambda.mli b/kernel/clambda.mli index 1476bb6e45..bd11c2667f 100644 --- a/kernel/clambda.mli +++ b/kernel/clambda.mli @@ -21,6 +21,7 @@ type lambda = | Lint of int | Lmakeblock of int * lambda array | Luint of Uint63.t + | Lfloat of Float64.t | Lval of structured_values | Lsort of Sorts.t | Lind of pinductive diff --git a/kernel/constr.ml b/kernel/constr.ml index 8375316003..b60b2d6d04 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -104,6 +104,7 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term = | CoFix of ('constr, 'types) pcofixpoint | Proj of Projection.t * 'constr | Int of Uint63.t + | Float of Float64.t (* constr is the fixpoint of the previous type. Requires option -rectypes of the Caml compiler to be set *) type t = (t, t, Sorts.t, Instance.t) kind_of_term @@ -241,6 +242,9 @@ let mkRef (gr,u) = let open GlobRef in match gr with (* Constructs a primitive integer *) let mkInt i = Int i +(* Constructs a primitive float number *) +let mkFloat f = Float f + (************************************************************************) (* kind_of_term = constructions as seen by the user *) (************************************************************************) @@ -446,7 +450,7 @@ let decompose_appvect c = let fold f acc c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> acc + | Construct _ | Int _ | Float _) -> acc | Cast (c,_,t) -> f (f acc c) t | Prod (_,t,c) -> f (f acc t) c | Lambda (_,t,c) -> f (f acc t) c @@ -466,7 +470,7 @@ let fold f acc c = match kind c with let iter f c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> () + | Construct _ | Int _ | Float _) -> () | Cast (c,_,t) -> f c; f t | Prod (_,t,c) -> f t; f c | Lambda (_,t,c) -> f t; f c @@ -486,7 +490,7 @@ let iter f c = match kind c with let iter_with_binders g f n c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> () + | Construct _ | Int _ | Float _) -> () | Cast (c,_,t) -> f n c; f n t | Prod (_,t,c) -> f n t; f (g n) c | Lambda (_,t,c) -> f n t; f (g n) c @@ -512,7 +516,7 @@ let iter_with_binders g f n c = match kind c with let fold_constr_with_binders g f n acc c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> acc + | Construct _ | Int _ | Float _) -> acc | Cast (c,_, t) -> f n (f n acc c) t | Prod (_na,t,c) -> f (g n) (f n acc t) c | Lambda (_na,t,c) -> f (g n) (f n acc t) c @@ -608,7 +612,7 @@ let map_return_predicate_with_full_binders g f l ci p = let map_gen userview f c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> c + | Construct _ | Int _ | Float _) -> c | Cast (b,k,t) -> let b' = f b in let t' = f t in @@ -673,7 +677,7 @@ let map = map_gen false let fold_map f accu c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> accu, c + | Construct _ | Int _ | Float _) -> accu, c | Cast (b,k,t) -> let accu, b' = f accu b in let accu, t' = f accu t in @@ -733,7 +737,7 @@ let fold_map f accu c = match kind c with let map_with_binders g f l c0 = match kind c0 with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _) -> c0 + | Construct _ | Int _ | Float _) -> c0 | Cast (c, k, t) -> let c' = f l c in let t' = f l t in @@ -810,7 +814,7 @@ let lift n = liftn n 1 let fold_with_full_binders g f n acc c = let open Context.Rel.Declaration in match kind c with - | Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ -> acc + | Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ | Float _ -> acc | Cast (c,_, t) -> f n (f n acc c) t | Prod (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c | Lambda (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c @@ -852,6 +856,7 @@ let compare_head_gen_leq_with kind1 kind2 leq_universes leq_sorts eq leq nargs t | Meta m1, Meta m2 -> Int.equal m1 m2 | Var id1, Var id2 -> Id.equal id1 id2 | Int i1, Int i2 -> Uint63.equal i1 i2 + | Float f1, Float f2 -> Float64.equal f1 f2 | Sort s1, Sort s2 -> leq_sorts s1 s2 | Prod (_,t1,c1), Prod (_,t2,c2) -> eq 0 t1 t2 && leq 0 c1 c2 | Lambda (_,t1,c1), Lambda (_,t2,c2) -> eq 0 t1 t2 && eq 0 c1 c2 @@ -878,7 +883,7 @@ let compare_head_gen_leq_with kind1 kind2 leq_universes leq_sorts eq leq nargs t Int.equal ln1 ln2 && Array.equal_norefl (eq 0) tl1 tl2 && Array.equal_norefl (eq 0) bl1 bl2 | (Rel _ | Meta _ | Var _ | Sort _ | Prod _ | Lambda _ | LetIn _ | App _ | Proj _ | Evar _ | Const _ | Ind _ | Construct _ | Case _ | Fix _ - | CoFix _ | Int _), _ -> false + | CoFix _ | Int _ | Float _), _ -> false (* [compare_head_gen_leq u s eq leq c1 c2] compare [c1] and [c2] using [eq] to compare the immediate subterms of [c1] of [c2] for conversion if needed, [leq] for cumulativity, @@ -1055,6 +1060,8 @@ let constr_ord_int f t1 t2 = | Proj (p1,c1), Proj (p2,c2) -> (Projection.compare =? f) p1 p2 c1 c2 | Proj _, _ -> -1 | _, Proj _ -> 1 | Int i1, Int i2 -> Uint63.compare i1 i2 + | Int _, _ -> -1 | _, Int _ -> 1 + | Float f1, Float f2 -> Float64.total_compare f1 f2 let rec compare m n= constr_ord_int compare m n @@ -1139,9 +1146,10 @@ let hasheq t1 t2 = && array_eqeq tl1 tl2 && array_eqeq bl1 bl2 | Int i1, Int i2 -> i1 == i2 + | Float f1, Float f2 -> Float64.equal f1 f2 | (Rel _ | Meta _ | Var _ | Sort _ | Cast _ | Prod _ | Lambda _ | LetIn _ | App _ | Proj _ | Evar _ | Const _ | Ind _ | Construct _ | Case _ - | Fix _ | CoFix _ | Int _), _ -> false + | Fix _ | CoFix _ | Int _ | Float _), _ -> false (** Note that the following Make has the side effect of creating once and for all the table we'll use for hash-consing all constr *) @@ -1247,6 +1255,7 @@ let hashcons (sh_sort,sh_ci,sh_construct,sh_ind,sh_con,sh_na,sh_id) = | Int i -> let (h,l) = Uint63.to_int2 i in (t, combinesmall 18 (combine h l)) + | Float f -> (t, combinesmall 19 (Float64.hash f)) and sh_rec t = let (y, h) = hash_term t in @@ -1311,6 +1320,7 @@ let rec hash t = | Proj (p,c) -> combinesmall 17 (combine (Projection.hash p) (hash c)) | Int i -> combinesmall 18 (Uint63.hash i) + | Float f -> combinesmall 19 (Float64.hash f) and hash_term_array t = Array.fold_left (fun acc t -> combine (hash t) acc) 0 t @@ -1455,3 +1465,4 @@ let rec debug_print c = cut() ++ str":=" ++ debug_print bd) (Array.to_list fixl)) ++ str"}") | Int i -> str"Int("++str (Uint63.to_string i) ++ str")" + | Float i -> str"Float("++str (Float64.to_string i) ++ str")" diff --git a/kernel/constr.mli b/kernel/constr.mli index 45ec8a7e64..4f8d682e42 100644 --- a/kernel/constr.mli +++ b/kernel/constr.mli @@ -76,6 +76,9 @@ val mkVar : Id.t -> constr (** Constructs a machine integer *) val mkInt : Uint63.t -> constr +(** Constructs a machine float number *) +val mkFloat : Float64.t -> constr + (** Constructs an patvar named "?n" *) val mkMeta : metavariable -> constr @@ -234,6 +237,7 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term = | CoFix of ('constr, 'types) pcofixpoint | Proj of Projection.t * 'constr | Int of Uint63.t + | Float of Float64.t (** User view of [constr]. For [App], it is ensured there is at least one argument and the function is not itself an applicative diff --git a/kernel/csymtable.ml b/kernel/csymtable.ml index 6c9e73b50d..cbffdc731e 100644 --- a/kernel/csymtable.ml +++ b/kernel/csymtable.ml @@ -184,7 +184,16 @@ and eval_to_patch env (buff,pl,fv) = | Reloc_proj_name p -> slot_for_proj_name p in let tc = patch buff pl slots in - let vm_env = Array.map (slot_for_fv env) fv in + let vm_env = + (* Beware, this may look like a call to [Array.map], but it's not. + Calling [Array.map f] when the first argument returned by [f] + is a float would lead to [vm_env] being an unboxed Double_array + (Tag_val = Double_array_tag) whereas eval_tcode expects a + regular array (Tag_val = 0). + See test-suite/primitive/float/coq_env_double_array.v + for an actual instance. *) + let a = Array.make (Array.length fv) crazy_val in + Array.iteri (fun i v -> a.(i) <- slot_for_fv env v) fv; a in eval_tcode tc (get_atom_rel ()) (vm_global global_data.glob_val) vm_env and val_of_constr env c = diff --git a/kernel/float64.ml b/kernel/float64.ml new file mode 100644 index 0000000000..3e36373b77 --- /dev/null +++ b/kernel/float64.ml @@ -0,0 +1,159 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(* * (see LICENSE file for the text of the license) *) +(************************************************************************) + +(* OCaml's float type follows the IEEE 754 Binary64 (double precision) + format *) +type t = float + +let is_nan f = f <> f +let is_infinity f = f = infinity +let is_neg_infinity f = f = neg_infinity + +(* Printing a binary64 float in 17 decimal places and parsing it again + will yield the same float. We assume [to_string_raw] is not given a + [nan] as input. *) +let to_string_raw f = Printf.sprintf "%.17g" f + +(* OCaml gives a sign to nan values which should not be displayed as + all NaNs are considered equal here *) +let to_string f = if is_nan f then "nan" else to_string_raw f +let of_string = float_of_string + +(* Compiles a float to OCaml code *) +let compile f = + let s = + if is_nan f then "nan" else if is_neg_infinity f then "neg_infinity" + else Printf.sprintf "%h" f in + Printf.sprintf "Float64.of_float (%s)" s + +let of_float f = f + +let sign f = copysign 1. f < 0. + +let opp = ( ~-. ) +let abs = abs_float + +type float_comparison = FEq | FLt | FGt | FNotComparable + +let eq x y = x = y +[@@ocaml.inline always] + +let lt x y = x < y +[@@ocaml.inline always] + +let le x y = x <= y +[@@ocaml.inline always] + +(* inspired by lib/util.ml; see also #10471 *) +let pervasives_compare = compare + +let compare x y = + if x < y then FLt + else + ( + if x > y then FGt + else + ( + if x = y then FEq + else FNotComparable (* NaN case *) + ) + ) +[@@ocaml.inline always] + +type float_class = + | PNormal | NNormal | PSubn | NSubn | PZero | NZero | PInf | NInf | NaN + +let classify x = + match classify_float x with + | FP_normal -> if 0. < x then PNormal else NNormal + | FP_subnormal -> if 0. < x then PSubn else NSubn + | FP_zero -> if 0. < 1. /. x then PZero else NZero + | FP_infinite -> if 0. < x then PInf else NInf + | FP_nan -> NaN +[@@ocaml.inline always] + +external mul : float -> float -> float = "coq_fmul_byte" "coq_fmul" +[@@unboxed] [@@noalloc] + +external add : float -> float -> float = "coq_fadd_byte" "coq_fadd" +[@@unboxed] [@@noalloc] + +external sub : float -> float -> float = "coq_fsub_byte" "coq_fsub" +[@@unboxed] [@@noalloc] + +external div : float -> float -> float = "coq_fdiv_byte" "coq_fdiv" +[@@unboxed] [@@noalloc] + +external sqrt : float -> float = "coq_fsqrt_byte" "coq_fsqrt" +[@@unboxed] [@@noalloc] + +let of_int63 x = Uint63.to_float x +[@@ocaml.inline always] + +let prec = 53 +let normfr_mantissa f = + let f = abs f in + if f >= 0.5 && f < 1. then Uint63.of_float (ldexp f prec) + else Uint63.zero +[@@ocaml.inline always] + +let eshift = 2101 (* 2*emax + prec *) + +(* When calling frexp on a nan or an infinity, the returned value inside + the exponent is undefined. + Therefore we must always set it to a fixed value (here 0). *) +let frshiftexp f = + match classify_float f with + | FP_zero | FP_infinite | FP_nan -> (f, Uint63.zero) + | FP_normal | FP_subnormal -> + let (m, e) = frexp f in + m, Uint63.of_int (e + eshift) +[@@ocaml.inline always] + +let ldshiftexp f e = ldexp f (Uint63.to_int_min e (2 * eshift) - eshift) +[@@ocaml.inline always] + +external next_up : float -> float = "coq_next_up_byte" "coq_next_up" +[@@unboxed] [@@noalloc] + +external next_down : float -> float = "coq_next_down_byte" "coq_next_down" +[@@unboxed] [@@noalloc] + +let equal f1 f2 = + match classify_float f1 with + | FP_normal | FP_subnormal | FP_infinite -> (f1 = f2) + | FP_nan -> is_nan f2 + | FP_zero -> f1 = f2 && 1. /. f1 = 1. /. f2 (* OCaml consider 0. = -0. *) +[@@ocaml.inline always] + +let hash = + (* Hashtbl.hash already considers all NaNs as equal, + cf. https://github.com/ocaml/ocaml/commit/aea227fdebe0b5361fd3e1d0aaa42cf929052269 + and http://caml.inria.fr/pub/docs/manual-ocaml/libref/Hashtbl.html *) + Hashtbl.hash + +let total_compare f1 f2 = + (* pervasives_compare considers all NaNs as equal, which is fine here, + but also considers -0. and +0. as equal *) + if f1 = 0. && f2 = 0. then pervasives_compare (1. /. f1) (1. /. f2) + else pervasives_compare f1 f2 + +let is_float64 t = + Obj.tag t = Obj.double_tag +[@@ocaml.inline always] + +(*** Test at runtime that no harmful double rounding seems to + be performed with an intermediate 80 bits representation (x87). *) +let () = + let b = ldexp 1. 53 in + let s = add 1. (ldexp 1. (-52)) in + if add b s <= b || add b 1. <> b || ldexp 1. (-1074) <= 0. then + failwith "Detected non IEEE-754 compliant architecture (or wrong \ + rounding mode). Use of Float is thus unsafe." diff --git a/kernel/float64.mli b/kernel/float64.mli new file mode 100644 index 0000000000..2aa9796526 --- /dev/null +++ b/kernel/float64.mli @@ -0,0 +1,95 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(* * (see LICENSE file for the text of the license) *) +(************************************************************************) + +(** [t] is currently implemented by OCaml's [float] type. + +Beware: NaNs have a sign and a payload, while they should be +indistinguishable from Coq's perspective. *) +type t + +(** Test functions for special values to avoid calling [classify] *) +val is_nan : t -> bool +val is_infinity : t -> bool +val is_neg_infinity : t -> bool + +val to_string : t -> string +val of_string : string -> t + +val compile : t -> string + +val of_float : float -> t + +(** Return [true] for "-", [false] for "+". *) +val sign : t -> bool + +val opp : t -> t +val abs : t -> t + +type float_comparison = FEq | FLt | FGt | FNotComparable + +val eq : t -> t -> bool + +val lt : t -> t -> bool + +val le : t -> t -> bool + +(** The IEEE 754 float comparison. + * NotComparable is returned if there is a NaN in the arguments *) +val compare : t -> t -> float_comparison +[@@ocaml.inline always] + +type float_class = + | PNormal | NNormal | PSubn | NSubn | PZero | NZero | PInf | NInf | NaN + +val classify : t -> float_class +[@@ocaml.inline always] + +val mul : t -> t -> t + +val add : t -> t -> t + +val sub : t -> t -> t + +val div : t -> t -> t + +val sqrt : t -> t + +(** Link with integers *) +val of_int63 : Uint63.t -> t +[@@ocaml.inline always] + +val normfr_mantissa : t -> Uint63.t +[@@ocaml.inline always] + +(** Shifted exponent extraction *) +val eshift : int + +val frshiftexp : t -> t * Uint63.t (* float remainder, shifted exponent *) +[@@ocaml.inline always] + +val ldshiftexp : t -> Uint63.t -> t +[@@ocaml.inline always] + +val next_up : t -> t + +val next_down : t -> t + +(** Return true if two floats are equal. + * All NaN values are considered equal. *) +val equal : t -> t -> bool +[@@ocaml.inline always] + +val hash : t -> int + +(** Total order relation over float values. Behaves like [Pervasives.compare].*) +val total_compare : t -> t -> int + +val is_float64 : Obj.t -> bool +[@@ocaml.inline always] diff --git a/kernel/genOpcodeFiles.ml b/kernel/genOpcodeFiles.ml index a8a4ffce9c..82bb2b584d 100644 --- a/kernel/genOpcodeFiles.ml +++ b/kernel/genOpcodeFiles.ml @@ -137,6 +137,26 @@ let opcodes = "CHECKTAIL0INT63"; "ISINT"; "AREINT2"; + "CHECKOPPFLOAT"; + "CHECKABSFLOAT"; + "CHECKEQFLOAT"; + "CHECKLTFLOAT"; + "LTFLOAT"; + "CHECKLEFLOAT"; + "LEFLOAT"; + "CHECKCOMPAREFLOAT"; + "CHECKCLASSIFYFLOAT"; + "CHECKADDFLOAT"; + "CHECKSUBFLOAT"; + "CHECKMULFLOAT"; + "CHECKDIVFLOAT"; + "CHECKSQRTFLOAT"; + "CHECKFLOATOFINT63"; + "CHECKFLOATNORMFRMANTISSA"; + "CHECKFRSHIFTEXP"; + "CHECKLDSHIFTEXP"; + "CHECKNEXTUPFLOAT"; + "CHECKNEXTDOWNFLOAT"; "STOP" |] diff --git a/kernel/inductive.ml b/kernel/inductive.ml index cd969ea457..320bc6a1cd 100644 --- a/kernel/inductive.ml +++ b/kernel/inductive.ml @@ -812,7 +812,7 @@ let rec subterm_specif renv stack t = | Not_subterm -> Not_subterm) | Var _ | Sort _ | Cast _ | Prod _ | LetIn _ | App _ | Const _ | Ind _ - | Construct _ | CoFix _ | Int _ -> Not_subterm + | Construct _ | CoFix _ | Int _ | Float _ -> Not_subterm (* Other terms are not subterms *) @@ -1057,7 +1057,7 @@ let check_one_fix renv recpos trees def = check_rec_call renv stack (Term.applist(c,l)) end - | Sort _ | Int _ -> + | Sort _ | Int _ | Float _ -> assert (List.is_empty l) (* l is not checked because it is considered as the meta's context *) @@ -1254,7 +1254,7 @@ let check_one_cofix env nbfix def deftype = | Evar _ -> List.iter (check_rec_call env alreadygrd n tree vlra) args | Rel _ | Var _ | Sort _ | Cast _ | Prod _ | LetIn _ | App _ | Const _ - | Ind _ | Fix _ | Proj _ | Int _ -> + | Ind _ | Fix _ | Proj _ | Int _ | Float _ -> raise (CoFixGuardError (env,NotGuardedForm t)) in let ((mind, _),_) = codomain_is_coind env deftype in diff --git a/kernel/inferCumulativity.ml b/kernel/inferCumulativity.ml index 3b8c2cd788..550c81ed82 100644 --- a/kernel/inferCumulativity.ml +++ b/kernel/inferCumulativity.ml @@ -102,6 +102,7 @@ let rec infer_fterm cv_pb infos variances hd stk = infer_vect infos variances (Array.map (mk_clos e) args) | FRel _ -> infer_stack infos variances stk | FInt _ -> infer_stack infos variances stk + | FFloat _ -> infer_stack infos variances stk | FFlex fl -> let variances = infer_table_key variances fl in infer_stack infos variances stk diff --git a/kernel/kernel.mllib b/kernel/kernel.mllib index 20e742d7f8..2b83c2d868 100644 --- a/kernel/kernel.mllib +++ b/kernel/kernel.mllib @@ -1,6 +1,7 @@ Names TransparentState Uint63 +Float64 CPrimitives Univ UGraph diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml index 1a5455cf3a..63dc49ba57 100644 --- a/kernel/nativecode.ml +++ b/kernel/nativecode.ml @@ -258,16 +258,19 @@ type primitive = | Mk_var of Id.t | Mk_proj | Is_int + | Is_float | Cast_accu | Upd_cofix | Force_cofix | Mk_uint + | Mk_float | Mk_int | Mk_bool | Val_to_int | Mk_meta | Mk_evar | MLand + | MLnot | MLle | MLlt | MLinteq @@ -349,6 +352,9 @@ let primitive_hash = function | Mk_proj -> 36 | MLarrayget -> 37 | Mk_empty_instance -> 38 + | Mk_float -> 39 + | Is_float -> 40 + | MLnot -> 41 type mllambda = | MLlocal of lname @@ -365,6 +371,7 @@ type mllambda = (* prefix, inductive name, tag, arguments *) | MLint of int | MLuint of Uint63.t + | MLfloat of Float64.t | MLsetref of string * mllambda | MLsequence of mllambda * mllambda | MLarray of mllambda array @@ -436,6 +443,8 @@ let rec eq_mllambda gn1 gn2 n env1 env2 t1 t2 = Int.equal i1 i2 | MLuint i1, MLuint i2 -> Uint63.equal i1 i2 + | MLfloat f1, MLfloat f2 -> + Float64.equal f1 f2 | MLsetref (id1, ml1), MLsetref (id2, ml2) -> String.equal id1 id2 && eq_mllambda gn1 gn2 n env1 env2 ml1 ml2 @@ -450,7 +459,7 @@ let rec eq_mllambda gn1 gn2 n env1 env2 t1 t2 = eq_mllambda gn1 gn2 n env1 env2 ml1 ml2 | (MLlocal _ | MLglobal _ | MLprimitive _ | MLlam _ | MLletrec _ | MLlet _ | MLapp _ | MLif _ | MLmatch _ | MLconstruct _ | MLint _ | MLuint _ | - MLsetref _ | MLsequence _ | MLarray _ | MLisaccu _), _ -> false + MLfloat _ | MLsetref _ | MLsequence _ | MLarray _ | MLisaccu _), _ -> false and eq_letrec gn1 gn2 n env1 env2 defs1 defs2 = let eq_def (_,args1,ml1) (_,args2,ml2) = @@ -535,6 +544,8 @@ let rec hash_mllambda gn n env t = combinesmall 15 (hash_mllambda_array gn n env 1 arr) | MLisaccu (s, ind, c) -> combinesmall 16 (combine (String.hash s) (combine (ind_hash ind) (hash_mllambda gn n env c))) + | MLfloat f -> + combinesmall 17 (Float64.hash f) and hash_mllambda_letrec gn n env init defs = let hash_def (_,args,ml) = @@ -568,7 +579,7 @@ let fv_lam l = match l with | MLlocal l -> if LNset.mem l bind then fv else LNset.add l fv - | MLglobal _ | MLprimitive _ | MLint _ | MLuint _ -> fv + | MLglobal _ | MLprimitive _ | MLint _ | MLuint _ | MLfloat _ -> fv | MLlam (ln,body) -> let bind = Array.fold_right LNset.add ln bind in aux body bind fv @@ -757,7 +768,7 @@ type env = env_named : (Id.t * mllambda) list ref; env_univ : lname option} -let empty_env univ () = +let empty_env univ = { env_rel = []; env_bound = 0; env_urel = ref []; @@ -958,25 +969,29 @@ type prim_aux = | PAprim of string * pconstant * CPrimitives.t * prim_aux array | PAml of mllambda -let add_check cond args = - let aux cond a = +let add_check cond targs args = + let aux cond t a = match a with | PAml(MLint _) -> cond | PAml ml -> (* FIXME: use explicit equality function *) - if List.mem ml cond then cond else ml::cond + if List.mem (t, ml) cond then cond else (t, ml)::cond | _ -> cond in - Array.fold_left aux cond args + Array.fold_left2 aux cond targs args let extract_prim ml_of l = let decl = ref [] in let cond = ref [] in + let type_args p = + let rec aux = function [] | [_] -> [] | h :: t -> h :: aux t in + Array.of_list (aux (CPrimitives.types p)) in let rec aux l = match l with | Lprim(prefix,kn,p,args) -> + let targs = type_args p in let args = Array.map aux args in - cond := add_check !cond args; + cond := add_check !cond targs args; PAprim(prefix,kn,p,args) | Lrel _ | Lvar _ | Luint _ | Lval _ | Lconst _ -> PAml (ml_of l) | _ -> @@ -1010,15 +1025,35 @@ let compile_prim decl cond paux = let compile_cond cond paux = match cond with | [] -> opt_prim_aux paux - | [c1] -> + | [CPrimitives.(PITT_type PT_int63), c1] -> MLif(app_prim Is_int [|c1|], opt_prim_aux paux, naive_prim_aux paux) - | c1::cond -> - let cond = - List.fold_left - (fun ml c -> app_prim MLland [| ml; cast_to_int c|]) - (app_prim MLland [| cast_to_int c1; MLint 0 |]) cond in - let cond = app_prim MLmagic [|cond|] in - MLif(cond, naive_prim_aux paux, opt_prim_aux paux) in + | _ -> + let ci, cf = + let is_int = + function CPrimitives.(PITT_type PT_int63), _ -> true | _ -> false in + List.partition is_int cond in + let condi = + let cond = + List.fold_left + (fun ml (_, c) -> app_prim MLland [| ml; cast_to_int c|]) + (MLint 0) ci in + app_prim MLmagic [|cond|] in + let condf = match cf with + | [] -> MLint 0 + | [_, c1] -> app_prim Is_float [|c1|] + | (_, c1) :: condf -> + List.fold_left + (fun ml (_, c) -> app_prim MLand [| ml; app_prim Is_float [|c|]|]) + (app_prim Is_float [|c1|]) condf in + match ci, cf with + | [], [] -> opt_prim_aux paux + | _ :: _, [] -> + MLif(condi, naive_prim_aux paux, opt_prim_aux paux) + | [], _ :: _ -> + MLif(condf, opt_prim_aux paux, naive_prim_aux paux) + | _ :: _, _ :: _ -> + let cond = app_prim MLand [|condf; app_prim MLnot [|condi|]|] in + MLif(cond, opt_prim_aux paux, naive_prim_aux paux) in let add_decl decl body = List.fold_left (fun body (x,d) -> MLlet(x,d,body)) body decl in @@ -1095,14 +1130,14 @@ let ml_of_instance instance u = (* Remark: if we do not want to compile the predicate we should a least compute the fv, then store the lambda representation of the predicate (not the mllambda) *) - let env_p = empty_env env.env_univ () in + let env_p = empty_env env.env_univ in let pn = fresh_gpred l in let mlp = ml_of_lam env_p l p in let mlp = generalize_fv env_p mlp in let (pfvn,pfvr) = !(env_p.env_named), !(env_p.env_urel) in let pn = push_global_let pn mlp in (* Compilation of the case *) - let env_c = empty_env env.env_univ () in + let env_c = empty_env env.env_univ in let a_uid = fresh_lname Anonymous in let la_uid = MLlocal a_uid in (* compilation of branches *) @@ -1158,7 +1193,7 @@ let ml_of_instance instance u = start *) (* Compilation of type *) - let env_t = empty_env env.env_univ () in + let env_t = empty_env env.env_univ in let ml_t = Array.map (ml_of_lam env_t l) tt in let params_t = fv_params env_t in let args_t = fv_args env !(env_t.env_named) !(env_t.env_urel) in @@ -1167,7 +1202,7 @@ let ml_of_instance instance u = let mk_type = MLapp(MLglobal gft, args_t) in (* Compilation of norm_i *) let ndef = Array.length ids in - let lf,env_n = push_rels (empty_env env.env_univ ()) ids in + let lf,env_n = push_rels (empty_env env.env_univ) ids in let t_params = Array.make ndef [||] in let t_norm_f = Array.make ndef (Gnorm (l,-1)) in let mk_let _envi (id,def) t = MLlet (id,def,t) in @@ -1224,7 +1259,7 @@ let ml_of_instance instance u = MLletrec(Array.mapi mkrec lf, lf_args.(start)) | Lcofix (start, (ids, tt, tb)) -> (* Compilation of type *) - let env_t = empty_env env.env_univ () in + let env_t = empty_env env.env_univ in let ml_t = Array.map (ml_of_lam env_t l) tt in let params_t = fv_params env_t in let args_t = fv_args env !(env_t.env_named) !(env_t.env_urel) in @@ -1233,7 +1268,7 @@ let ml_of_instance instance u = let mk_type = MLapp(MLglobal gft, args_t) in (* Compilation of norm_i *) let ndef = Array.length ids in - let lf,env_n = push_rels (empty_env env.env_univ ()) ids in + let lf,env_n = push_rels (empty_env env.env_univ) ids in let t_params = Array.make ndef [||] in let t_norm_f = Array.make ndef (Gnorm (l,-1)) in let ml_of_fix i body = @@ -1297,6 +1332,7 @@ let ml_of_instance instance u = let args = Array.map (ml_of_lam env l) args in MLconstruct(prefix,cn,tag,args) | Luint i -> MLapp(MLprimitive Mk_uint, [|MLuint i|]) + | Lfloat f -> MLapp(MLprimitive Mk_float, [|MLfloat f|]) | Lval v -> let i = push_symbol (SymbValue v) in get_value_code i | Lsort s -> @@ -1314,7 +1350,7 @@ let ml_of_instance instance u = | Lforce -> MLglobal (Ginternal "Lazy.force") let mllambda_of_lambda univ auxdefs l t = - let env = empty_env univ () in + let env = empty_env univ in global_stack := auxdefs; let ml = ml_of_lam env l t in let fv_rel = !(env.env_urel) in @@ -1347,7 +1383,7 @@ let subst s l = let rec aux l = match l with | MLlocal id -> (try LNmap.find id s with Not_found -> l) - | MLglobal _ | MLprimitive _ | MLint _ | MLuint _ -> l + | MLglobal _ | MLprimitive _ | MLint _ | MLuint _ | MLfloat _ -> l | MLlam(params,body) -> MLlam(params, aux body) | MLletrec(defs,body) -> let arec (f,params,body) = (f,params,aux body) in @@ -1417,7 +1453,7 @@ let optimize gdef l = let rec optimize s l = match l with | MLlocal id -> (try LNmap.find id s with Not_found -> l) - | MLglobal _ | MLprimitive _ | MLint _ | MLuint _ -> l + | MLglobal _ | MLprimitive _ | MLint _ | MLuint _ | MLfloat _ -> l | MLlam(params,body) -> MLlam(params, optimize s body) | MLletrec(decls,body) -> @@ -1623,6 +1659,7 @@ let pp_mllam fmt l = (string_of_construct prefix ~constant:false ind tag) pp_cargs args | MLint i -> pp_int fmt i | MLuint i -> Format.fprintf fmt "(%s)" (Uint63.compile i) + | MLfloat f -> Format.fprintf fmt "(%s)" (Float64.compile f) | MLsetref (s, body) -> Format.fprintf fmt "@[%s@ :=@\n %a@]" s pp_mllam body | MLsequence(l1,l2) -> @@ -1739,16 +1776,19 @@ let pp_mllam fmt l = Format.fprintf fmt "mk_var_accu (Names.Id.of_string \"%s\")" (string_of_id id) | Mk_proj -> Format.fprintf fmt "mk_proj_accu" | Is_int -> Format.fprintf fmt "is_int" + | Is_float -> Format.fprintf fmt "is_float" | Cast_accu -> Format.fprintf fmt "cast_accu" | Upd_cofix -> Format.fprintf fmt "upd_cofix" | Force_cofix -> Format.fprintf fmt "force_cofix" | Mk_uint -> Format.fprintf fmt "mk_uint" + | Mk_float -> Format.fprintf fmt "mk_float" | Mk_int -> Format.fprintf fmt "mk_int" | Mk_bool -> Format.fprintf fmt "mk_bool" | Val_to_int -> Format.fprintf fmt "val_to_int" | Mk_meta -> Format.fprintf fmt "mk_meta_accu" | Mk_evar -> Format.fprintf fmt "mk_evar_accu" | MLand -> Format.fprintf fmt "(&&)" + | MLnot -> Format.fprintf fmt "not" | MLle -> Format.fprintf fmt "(<=)" | MLlt -> Format.fprintf fmt "(<)" | MLinteq -> Format.fprintf fmt "(==)" diff --git a/kernel/nativeconv.ml b/kernel/nativeconv.ml index dd010e5cad..ef610ce7e9 100644 --- a/kernel/nativeconv.ml +++ b/kernel/nativeconv.ml @@ -35,6 +35,9 @@ let rec conv_val env pb lvl v1 v2 cu = if Int.equal i1 i2 then cu else raise NotConvertible | Vint64 i1, Vint64 i2 -> if Int64.equal i1 i2 then cu else raise NotConvertible + | Vfloat64 f1, Vfloat64 f2 -> + if Float64.(equal (of_float f1) (of_float f2)) then cu + else raise NotConvertible | Vblock b1, Vblock b2 -> let n1 = block_size b1 in let n2 = block_size b2 in @@ -48,7 +51,7 @@ let rec conv_val env pb lvl v1 v2 cu = aux lvl max b1 b2 (i+1) cu in aux lvl (n1-1) b1 b2 0 cu - | Vaccu _, _ | Vconst _, _ | Vint64 _, _ | Vblock _, _ -> raise NotConvertible + | (Vaccu _ | Vconst _ | Vint64 _ | Vfloat64 _ | Vblock _), _ -> raise NotConvertible and conv_accu env pb lvl k1 k2 cu = let n1 = accu_nargs k1 in diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml index 70b3beb2dc..7a4e62cdfe 100644 --- a/kernel/nativelambda.ml +++ b/kernel/nativelambda.ml @@ -44,6 +44,7 @@ type lambda = (* prefix, inductive name, constructor tag, arguments *) (* A fully applied non-constant constructor *) | Luint of Uint63.t + | Lfloat of Float64.t | Lval of Nativevalues.t | Lsort of Sorts.t | Lind of prefix * pinductive @@ -123,7 +124,7 @@ let get_const_prefix env c = let map_lam_with_binders g f n lam = match lam with | Lrel _ | Lvar _ | Lconst _ | Lproj _ | Lval _ | Lsort _ | Lind _ | Luint _ - | Llazy | Lforce | Lmeta _ | Lint _ -> lam + | Llazy | Lforce | Lmeta _ | Lint _ | Lfloat _ -> lam | Lprod(dom,codom) -> let dom' = f n dom in let codom' = f n codom in @@ -331,7 +332,7 @@ and reduce_lapp substf lids body substa largs = let is_value lc = match lc with - | Lval _ | Lint _ | Luint _ -> true + | Lval _ | Lint _ | Luint _ | Lfloat _ -> true | _ -> false let get_value lc = @@ -339,6 +340,7 @@ let get_value lc = | Lval v -> v | Lint tag -> Nativevalues.mk_int tag | Luint i -> Nativevalues.mk_uint i + | Lfloat f -> Nativevalues.mk_float f | _ -> raise Not_found let make_args start _end = @@ -364,7 +366,12 @@ let makeblock env ind tag nparams arity args = if Int.equal arity 0 then Lint tag else if Array.for_all is_value args then - let args = Array.map get_value args in + let dummy_val = Obj.magic 0 in + let args = + (* Don't simplify this to Array.map, cf. the related comment in + function eval_to_patch, file kernel/csymtable.ml *) + let a = Array.make (Array.length args) dummy_val in + Array.iteri (fun i v -> a.(i) <- get_value v) args; a in Lval (Nativevalues.mk_block tag args) else let prefix = get_mind_prefix env (fst ind) in @@ -580,6 +587,8 @@ let rec lambda_of_constr cache env sigma c = | Int i -> Luint i + | Float f -> Lfloat f + and lambda_of_app cache env sigma f args = match kind f with | Const (_kn,_u as c) -> diff --git a/kernel/nativelambda.mli b/kernel/nativelambda.mli index f17339f84d..1d7bf5343a 100644 --- a/kernel/nativelambda.mli +++ b/kernel/nativelambda.mli @@ -38,6 +38,7 @@ type lambda = (* prefix, inductive name, constructor tag, arguments *) (* A fully applied non-constant constructor *) | Luint of Uint63.t + | Lfloat of Float64.t | Lval of Nativevalues.t | Lsort of Sorts.t | Lind of prefix * pinductive diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml index f788832d5b..e4a8344eaf 100644 --- a/kernel/nativevalues.ml +++ b/kernel/nativevalues.ml @@ -225,6 +225,9 @@ let mk_bool (b : bool) = (Obj.magic (not b) : t) let mk_uint (x : Uint63.t) = (Obj.magic x : t) [@@ocaml.inline always] +let mk_float (x : Float64.t) = (Obj.magic x : t) +[@@ocaml.inline always] + type block let block_size (b:block) = @@ -240,16 +243,19 @@ type kind_of_value = | Vfun of (t -> t) | Vconst of int | Vint64 of int64 + | Vfloat64 of float | Vblock of block let kind_of_value (v:t) = let o = Obj.repr v in if Obj.is_int o then Vconst (Obj.magic v) + else if Obj.tag o == Obj.double_tag then Vfloat64 (Obj.magic v) else let tag = Obj.tag o in if Int.equal tag accumulate_tag then Vaccu (Obj.magic v) else if Int.equal tag Obj.custom_tag then Vint64 (Obj.magic v) + else if Int.equal tag Obj.double_tag then Vfloat64 (Obj.magic v) else if (tag < Obj.lazy_tag) then Vblock (Obj.magic v) else (* assert (tag = Obj.closure_tag || tag = Obj.infix_tag); @@ -261,6 +267,7 @@ let kind_of_value (v:t) = let is_int (x:t) = let o = Obj.repr x in Obj.is_int o || Int.equal (Obj.tag o) Obj.custom_tag +[@@ocaml.inline always] let val_to_int (x:t) = (Obj.magic x : int) [@@ocaml.inline always] @@ -508,6 +515,177 @@ let print x = flush stderr; x +(** Support for machine floating point values *) + +external is_float : t -> bool = "coq_is_double" +[@@noalloc] + +let to_float (x:t) = (Obj.magic x : Float64.t) +[@@ocaml.inline always] + +let no_check_fopp x = + mk_float (Float64.opp (to_float x)) +[@@ocaml.inline always] + +let fopp accu x = + if is_float x then no_check_fopp x + else accu x + +let no_check_fabs x = + mk_float (Float64.abs (to_float x)) +[@@ocaml.inline always] + +let fabs accu x = + if is_float x then no_check_fabs x + else accu x + +let no_check_feq x y = + mk_bool (Float64.eq (to_float x) (to_float y)) + +let feq accu x y = + if is_float x && is_float y then no_check_feq x y + else accu x y + +let no_check_flt x y = + mk_bool (Float64.lt (to_float x) (to_float y)) + +let flt accu x y = + if is_float x && is_float y then no_check_flt x y + else accu x y + +let no_check_fle x y = + mk_bool (Float64.le (to_float x) (to_float y)) + +let fle accu x y = + if is_float x && is_float y then no_check_fle x y + else accu x y + +type coq_fcmp = + | CFcmpAccu of t + | CFcmpEq + | CFcmpLt + | CFcmpGt + | CFcmpNotComparable + +let no_check_fcompare x y = + let c = Float64.compare (to_float x) (to_float y) in + (Obj.magic c:t) +[@@ocaml.inline always] + +let fcompare accu x y = + if is_float x && is_float y then no_check_fcompare x y + else accu x y + +type coq_fclass = + | CFclassAccu of t + | CFclassPNormal + | CFclassNNormal + | CFclassPSubn + | CFclassNSubn + | CFclassPZero + | CFclassNZero + | CFclassPInf + | CFclassNInf + | CFclassNaN + +let no_check_fclassify x = + let c = Float64.classify (to_float x) in + (Obj.magic c:t) +[@@ocaml.inline always] + +let fclassify accu x = + if is_float x then no_check_fclassify x + else accu x + +let no_check_fadd x y = + mk_float (Float64.add (to_float x) (to_float y)) +[@@ocaml.inline always] + +let fadd accu x y = + if is_float x && is_float y then no_check_fadd x y + else accu x y + +let no_check_fsub x y = + mk_float (Float64.sub (to_float x) (to_float y)) +[@@ocaml.inline always] + +let fsub accu x y = + if is_float x && is_float y then no_check_fsub x y + else accu x y + +let no_check_fmul x y = + mk_float (Float64.mul (to_float x) (to_float y)) +[@@ocaml.inline always] + +let fmul accu x y = + if is_float x && is_float y then no_check_fmul x y + else accu x y + +let no_check_fdiv x y = + mk_float (Float64.div (to_float x) (to_float y)) +[@@ocaml.inline always] + +let fdiv accu x y = + if is_float x && is_float y then no_check_fdiv x y + else accu x y + +let no_check_fsqrt x = + mk_float (Float64.sqrt (to_float x)) +[@@ocaml.inline always] + +let fsqrt accu x = + if is_float x then no_check_fsqrt x + else accu x + +let no_check_float_of_int x = + mk_float (Float64.of_int63 (to_uint x)) +[@@ocaml.inline always] + +let float_of_int accu x = + if is_int x then no_check_float_of_int x + else accu x + +let no_check_normfr_mantissa x = + mk_uint (Float64.normfr_mantissa (to_float x)) +[@@ocaml.inline always] + +let normfr_mantissa accu x = + if is_float x then no_check_normfr_mantissa x + else accu x + +let no_check_frshiftexp x = + let f, e = Float64.frshiftexp (to_float x) in + (Obj.magic (PPair(mk_float f, mk_uint e)):t) +[@@ocaml.inline always] + +let frshiftexp accu x = + if is_float x then no_check_frshiftexp x + else accu x + +let no_check_ldshiftexp x e = + mk_float (Float64.ldshiftexp (to_float x) (to_uint e)) +[@@ocaml.inline always] + +let ldshiftexp accu x e = + if is_float x && is_int e then no_check_ldshiftexp x e + else accu x e + +let no_check_next_up x = + mk_float (Float64.next_up (to_float x)) +[@@ocaml.inline always] + +let next_up accu x = + if is_float x then no_check_next_up x + else accu x + +let no_check_next_down x = + mk_float (Float64.next_down (to_float x)) +[@@ocaml.inline always] + +let next_down accu x = + if is_float x then no_check_next_down x + else accu x + let hobcnv = Array.init 256 (fun i -> Printf.sprintf "%02x" i) let bohcnv = Array.init 256 (fun i -> i - (if 0x30 <= i then 0x30 else 0) - diff --git a/kernel/nativevalues.mli b/kernel/nativevalues.mli index b54f437e73..815ef3e98e 100644 --- a/kernel/nativevalues.mli +++ b/kernel/nativevalues.mli @@ -102,6 +102,9 @@ val mk_int : int -> t val mk_uint : Uint63.t -> t [@@ocaml.inline always] +val mk_float : Float64.t -> t +[@@ocaml.inline always] + val napply : t -> t array -> t (* Functions over accumulators *) @@ -130,6 +133,7 @@ type kind_of_value = | Vfun of (t -> t) | Vconst of int | Vint64 of int64 + | Vfloat64 of float | Vblock of block val kind_of_value : t -> kind_of_value @@ -140,7 +144,9 @@ val str_decode : string -> 'a (** Support for machine integers *) val val_to_int : t -> int + val is_int : t -> bool +[@@ocaml.inline always] (* function with check *) val head0 : t -> t -> t @@ -247,3 +253,82 @@ val no_check_le : t -> t -> t [@@ocaml.inline always] val no_check_compare : t -> t -> t + +(** Support for machine floating point values *) + +val is_float : t -> bool +[@@ocaml.inline always] + +val fopp : t -> t -> t +val fabs : t -> t -> t +val feq : t -> t -> t -> t +val flt : t -> t -> t -> t +val fle : t -> t -> t -> t +val fcompare : t -> t -> t -> t +val fclassify : t -> t -> t +val fadd : t -> t -> t -> t +val fsub : t -> t -> t -> t +val fmul : t -> t -> t -> t +val fdiv : t -> t -> t -> t +val fsqrt : t -> t -> t +val float_of_int : t -> t -> t +val normfr_mantissa : t -> t -> t +val frshiftexp : t -> t -> t +val ldshiftexp : t -> t -> t -> t +val next_up : t -> t -> t +val next_down : t -> t -> t + +(* Function without check *) +val no_check_fopp : t -> t +[@@ocaml.inline always] + +val no_check_fabs : t -> t +[@@ocaml.inline always] + +val no_check_feq : t -> t -> t +[@@ocaml.inline always] + +val no_check_flt : t -> t -> t +[@@ocaml.inline always] + +val no_check_fle : t -> t -> t +[@@ocaml.inline always] + +val no_check_fcompare : t -> t -> t +[@@ocaml.inline always] + +val no_check_fclassify : t -> t +[@@ocaml.inline always] + +val no_check_fadd : t -> t -> t +[@@ocaml.inline always] + +val no_check_fsub : t -> t -> t +[@@ocaml.inline always] + +val no_check_fmul : t -> t -> t +[@@ocaml.inline always] + +val no_check_fdiv : t -> t -> t +[@@ocaml.inline always] + +val no_check_fsqrt : t -> t +[@@ocaml.inline always] + +val no_check_float_of_int : t -> t +[@@ocaml.inline always] + +val no_check_normfr_mantissa : t -> t +[@@ocaml.inline always] + +val no_check_frshiftexp : t -> t +[@@ocaml.inline always] + +val no_check_ldshiftexp : t -> t -> t +[@@ocaml.inline always] + +val no_check_next_up : t -> t +[@@ocaml.inline always] + +val no_check_next_down : t -> t +[@@ocaml.inline always] diff --git a/kernel/primred.ml b/kernel/primred.ml index d6d0a6143a..c475828cb3 100644 --- a/kernel/primred.ml +++ b/kernel/primred.ml @@ -14,6 +14,13 @@ let add_retroknowledge env action = | None -> { retro with retro_int63 = Some c } | Some c' -> assert (Constant.equal c c'); retro in set_retroknowledge env retro + | Register_type(PT_float64,c) -> + let retro = env.retroknowledge in + let retro = + match retro.retro_float64 with + | None -> { retro with retro_float64 = Some c } + | Some c' -> assert (Constant.equal c c'); retro in + set_retroknowledge env retro | Register_ind(pit,ind) -> let retro = env.retroknowledge in let retro = @@ -42,6 +49,21 @@ let add_retroknowledge env action = | None -> ((ind,1), (ind,2), (ind,3)) | Some (((ind',_),_,_) as t) -> assert (eq_ind ind ind'); t in { retro with retro_cmp = Some r } + | PIT_f_cmp -> + let r = + match retro.retro_f_cmp with + | None -> ((ind,1), (ind,2), (ind,3), (ind,4)) + | Some (((ind',_),_,_,_) as t) -> assert (eq_ind ind ind'); t in + { retro with retro_f_cmp = Some r } + | PIT_f_class -> + let r = + match retro.retro_f_class with + | None -> ((ind,1), (ind,2), (ind,3), (ind,4), + (ind,5), (ind,6), (ind,7), (ind,8), + (ind,9)) + | Some (((ind',_),_,_,_,_,_,_,_,_) as t) -> + assert (eq_ind ind ind'); t in + { retro with retro_f_class = Some r } in set_retroknowledge env retro @@ -50,6 +72,17 @@ let get_int_type env = | Some c -> c | None -> anomaly Pp.(str"Reduction of primitive: int63 not registered") +let get_float_type env = + match env.retroknowledge.retro_float64 with + | Some c -> c + | None -> anomaly Pp.(str"Reduction of primitive: float64 not registered") + +let get_cmp_type env = + match env.retroknowledge.retro_cmp with + | Some (((mindcmp,_),_),_,_) -> + Constant.make (MutInd.user mindcmp) (MutInd.canonical mindcmp) + | None -> anomaly Pp.(str"Reduction of primitive: comparison not registered") + let get_bool_constructors env = match env.retroknowledge.retro_bool with | Some r -> r @@ -70,6 +103,16 @@ let get_cmp_constructors env = | Some r -> r | None -> anomaly Pp.(str"Reduction of primitive: cmp not registered") +let get_f_cmp_constructors env = + match env.retroknowledge.retro_f_cmp with + | Some r -> r + | None -> anomaly Pp.(str"Reduction of primitive: fcmp not registered") + +let get_f_class_constructors env = + match env.retroknowledge.retro_f_class with + | Some r -> r + | None -> anomaly Pp.(str"Reduction of primitive: fclass not registered") + exception NativeDestKO module type RedNativeEntries = @@ -80,14 +123,29 @@ module type RedNativeEntries = val get : args -> int -> elem val get_int : evd -> elem -> Uint63.t + val get_float : evd -> elem -> Float64.t val mkInt : env -> Uint63.t -> elem + val mkFloat : env -> Float64.t -> elem val mkBool : env -> bool -> elem val mkCarry : env -> bool -> elem -> elem (* true if carry *) val mkIntPair : env -> elem -> elem -> elem + val mkFloatIntPair : env -> elem -> elem -> elem val mkLt : env -> elem val mkEq : env -> elem val mkGt : env -> elem - + val mkFLt : env -> elem + val mkFEq : env -> elem + val mkFGt : env -> elem + val mkFNotComparable : env -> elem + val mkPNormal : env -> elem + val mkNNormal : env -> elem + val mkPSubn : env -> elem + val mkNSubn : env -> elem + val mkPZero : env -> elem + val mkNZero : env -> elem + val mkPInf : env -> elem + val mkNInf : env -> elem + val mkNaN : env -> elem end module type RedNative = @@ -116,6 +174,12 @@ struct let get_int3 evd args = get_int evd args 0, get_int evd args 1, get_int evd args 2 + let get_float evd args i = E.get_float evd (E.get args i) + + let get_float1 evd args = get_float evd args 0 + + let get_float2 evd args = get_float evd args 0, get_float evd args 1 + let red_prim_aux env evd op args = let open CPrimitives in match op with @@ -193,6 +257,64 @@ struct | 0 -> E.mkEq env | _ -> E.mkGt env end + | Float64opp -> + let f = get_float1 evd args in E.mkFloat env (Float64.opp f) + | Float64abs -> + let f = get_float1 evd args in E.mkFloat env (Float64.abs f) + | Float64eq -> + let i1, i2 = get_float2 evd args in + E.mkBool env (Float64.eq i1 i2) + | Float64lt -> + let i1, i2 = get_float2 evd args in + E.mkBool env (Float64.lt i1 i2) + | Float64le -> + let i1, i2 = get_float2 evd args in + E.mkBool env (Float64.le i1 i2) + | Float64compare -> + let f1, f2 = get_float2 evd args in + (match Float64.compare f1 f2 with + | Float64.FEq -> E.mkFEq env + | Float64.FLt -> E.mkFLt env + | Float64.FGt -> E.mkFGt env + | Float64.FNotComparable -> E.mkFNotComparable env) + | Float64classify -> + let f = get_float1 evd args in + (match Float64.classify f with + | Float64.PNormal -> E.mkPNormal env + | Float64.NNormal -> E.mkNNormal env + | Float64.PSubn -> E.mkPSubn env + | Float64.NSubn -> E.mkNSubn env + | Float64.PZero -> E.mkPZero env + | Float64.NZero -> E.mkNZero env + | Float64.PInf -> E.mkPInf env + | Float64.NInf -> E.mkNInf env + | Float64.NaN -> E.mkNaN env) + | Float64add -> + let f1, f2 = get_float2 evd args in E.mkFloat env (Float64.add f1 f2) + | Float64sub -> + let f1, f2 = get_float2 evd args in E.mkFloat env (Float64.sub f1 f2) + | Float64mul -> + let f1, f2 = get_float2 evd args in E.mkFloat env (Float64.mul f1 f2) + | Float64div -> + let f1, f2 = get_float2 evd args in E.mkFloat env (Float64.div f1 f2) + | Float64sqrt -> + let f = get_float1 evd args in E.mkFloat env (Float64.sqrt f) + | Float64ofInt63 -> + let i = get_int1 evd args in E.mkFloat env (Float64.of_int63 i) + | Float64normfr_mantissa -> + let f = get_float1 evd args in E.mkInt env (Float64.normfr_mantissa f) + | Float64frshiftexp -> + let f = get_float1 evd args in + let (m,e) = Float64.frshiftexp f in + E.mkFloatIntPair env (E.mkFloat env m) (E.mkInt env e) + | Float64ldshiftexp -> + let f = get_float evd args 0 in + let e = get_int evd args 1 in + E.mkFloat env (Float64.ldshiftexp f e) + | Float64next_up -> + let f = get_float1 evd args in E.mkFloat env (Float64.next_up f) + | Float64next_down -> + let f = get_float1 evd args in E.mkFloat env (Float64.next_down f) let red_prim env evd p args = try diff --git a/kernel/primred.mli b/kernel/primred.mli index f5998982d7..bbe564d8e7 100644 --- a/kernel/primred.mli +++ b/kernel/primred.mli @@ -5,10 +5,17 @@ open Environ val add_retroknowledge : env -> Retroknowledge.action -> env val get_int_type : env -> Constant.t +val get_float_type : env -> Constant.t +val get_cmp_type : env -> Constant.t val get_bool_constructors : env -> constructor * constructor val get_carry_constructors : env -> constructor * constructor val get_pair_constructor : env -> constructor val get_cmp_constructors : env -> constructor * constructor * constructor +val get_f_cmp_constructors : env -> constructor * constructor * constructor * constructor +val get_f_class_constructors : + env -> constructor * constructor * constructor * constructor + * constructor * constructor * constructor * constructor + * constructor exception NativeDestKO (* Should be raised by get_* functions on failure *) @@ -20,13 +27,29 @@ module type RedNativeEntries = val get : args -> int -> elem val get_int : evd -> elem -> Uint63.t + val get_float : evd -> elem -> Float64.t val mkInt : env -> Uint63.t -> elem + val mkFloat : env -> Float64.t -> elem val mkBool : env -> bool -> elem val mkCarry : env -> bool -> elem -> elem (* true if carry *) val mkIntPair : env -> elem -> elem -> elem + val mkFloatIntPair : env -> elem -> elem -> elem val mkLt : env -> elem val mkEq : env -> elem val mkGt : env -> elem + val mkFLt : env -> elem + val mkFEq : env -> elem + val mkFGt : env -> elem + val mkFNotComparable : env -> elem + val mkPNormal : env -> elem + val mkNNormal : env -> elem + val mkPSubn : env -> elem + val mkNSubn : env -> elem + val mkPZero : env -> elem + val mkNZero : env -> elem + val mkPInf : env -> elem + val mkNInf : env -> elem + val mkNaN : env -> elem end module type RedNative = diff --git a/kernel/reduction.ml b/kernel/reduction.ml index 327cb2efeb..0cc7692fcf 100644 --- a/kernel/reduction.ml +++ b/kernel/reduction.ml @@ -138,10 +138,10 @@ let nf_betaiota env t = let whd_betaiotazeta env x = match kind x with | (Sort _|Var _|Meta _|Evar _|Const _|Ind _|Construct _| - Prod _|Lambda _|Fix _|CoFix _|Int _) -> x + Prod _|Lambda _|Fix _|CoFix _|Int _|Float _) -> x | App (c, _) -> begin match kind c with - | Ind _ | Construct _ | Evar _ | Meta _ | Const _ | Int _ -> x + | Ind _ | Construct _ | Evar _ | Meta _ | Const _ | Int _ | Float _ -> x | Sort _ | Rel _ | Var _ | Cast _ | Prod _ | Lambda _ | LetIn _ | App _ | Case _ | Fix _ | CoFix _ | Proj _ -> whd_val (create_clos_infos betaiotazeta env) (create_tab ()) (inject x) @@ -152,10 +152,10 @@ let whd_betaiotazeta env x = let whd_all env t = match kind t with | (Sort _|Meta _|Evar _|Ind _|Construct _| - Prod _|Lambda _|Fix _|CoFix _|Int _) -> t + Prod _|Lambda _|Fix _|CoFix _|Int _|Float _) -> t | App (c, _) -> begin match kind c with - | Ind _ | Construct _ | Evar _ | Meta _ | Int _ -> t + | Ind _ | Construct _ | Evar _ | Meta _ | Int _ | Float _ -> t | Sort _ | Rel _ | Var _ | Cast _ | Prod _ | Lambda _ | LetIn _ | App _ | Const _ |Case _ | Fix _ | CoFix _ | Proj _ -> whd_val (create_clos_infos all env) (create_tab ()) (inject t) @@ -166,10 +166,10 @@ let whd_all env t = let whd_allnolet env t = match kind t with | (Sort _|Meta _|Evar _|Ind _|Construct _| - Prod _|Lambda _|Fix _|CoFix _|LetIn _|Int _) -> t + Prod _|Lambda _|Fix _|CoFix _|LetIn _|Int _|Float _) -> t | App (c, _) -> begin match kind c with - | Ind _ | Construct _ | Evar _ | Meta _ | LetIn _ | Int _ -> t + | Ind _ | Construct _ | Evar _ | Meta _ | LetIn _ | Int _ | Float _ -> t | Sort _ | Rel _ | Var _ | Cast _ | Prod _ | Lambda _ | App _ | Const _ | Case _ | Fix _ | CoFix _ | Proj _ -> whd_val (create_clos_infos allnolet env) (create_tab ()) (inject t) @@ -627,13 +627,17 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = if Uint63.equal i1 i2 then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible + | FFloat f1, FFloat f2 -> + if Float64.equal f1 f2 then convert_stacks l2r infos lft1 lft2 v1 v2 cuniv + else raise NotConvertible + (* Should not happen because both (hd1,v1) and (hd2,v2) are in whnf *) | ( (FLetIn _, _) | (FCaseT _,_) | (FApp _,_) | (FCLOS _,_) | (FLIFT _,_) | (_, FLetIn _) | (_,FCaseT _) | (_,FApp _) | (_,FCLOS _) | (_,FLIFT _) | (FLOCKED,_) | (_,FLOCKED) ) -> assert false | (FRel _ | FAtom _ | FInd _ | FFix _ | FCoFix _ - | FProd _ | FEvar _ | FInt _), _ -> raise NotConvertible + | FProd _ | FEvar _ | FInt _ | FFloat _), _ -> raise NotConvertible and convert_stacks l2r infos lft1 lft2 stk1 stk2 cuniv = let f (l1, t1) (l2, t2) cuniv = ccnv CONV l2r infos l1 l2 t1 t2 cuniv in diff --git a/kernel/retroknowledge.ml b/kernel/retroknowledge.ml index 873c6af93d..479fe02295 100644 --- a/kernel/retroknowledge.ml +++ b/kernel/retroknowledge.ml @@ -18,23 +18,37 @@ open Names type retroknowledge = { retro_int63 : Constant.t option; + retro_float64 : Constant.t option; retro_bool : (constructor * constructor) option; (* true, false *) retro_carry : (constructor * constructor) option; (* C0, C1 *) retro_pair : constructor option; retro_cmp : (constructor * constructor * constructor) option; (* Eq, Lt, Gt *) - retro_refl : constructor option; + retro_f_cmp : (constructor * constructor * constructor * constructor) + option; + (* FEq, FLt, FGt, FNotComparable *) + retro_f_class : (constructor * constructor * constructor * constructor + * constructor * constructor * constructor * constructor + * constructor) + option; + (* PNormal, NNormal, PSubn, NSubn, + PZero, NZero, PInf, NInf, + NaN *) + retro_refl : constructor option } let empty = { retro_int63 = None; + retro_float64 = None; retro_bool = None; retro_carry = None; retro_pair = None; retro_cmp = None; + retro_f_cmp = None; + retro_f_class = None; retro_refl = None; } type action = - | Register_ind of CPrimitives.prim_ind * inductive - | Register_type of CPrimitives.prim_type * Constant.t + | Register_ind : 'a CPrimitives.prim_ind * inductive -> action + | Register_type : CPrimitives.prim_type * Constant.t -> action diff --git a/kernel/retroknowledge.mli b/kernel/retroknowledge.mli index 2a7b390951..2df8a00465 100644 --- a/kernel/retroknowledge.mli +++ b/kernel/retroknowledge.mli @@ -12,16 +12,27 @@ open Names type retroknowledge = { retro_int63 : Constant.t option; + retro_float64 : Constant.t option; retro_bool : (constructor * constructor) option; (* true, false *) retro_carry : (constructor * constructor) option; (* C0, C1 *) retro_pair : constructor option; retro_cmp : (constructor * constructor * constructor) option; (* Eq, Lt, Gt *) - retro_refl : constructor option; + retro_f_cmp : (constructor * constructor * constructor * constructor) + option; + (* FEq, FLt, FGt, FNotComparable *) + retro_f_class : (constructor * constructor * constructor * constructor + * constructor * constructor * constructor * constructor + * constructor) + option; + (* PNormal, NNormal, PSubn, NSubn, + PZero, NZero, PInf, NInf, + NaN *) + retro_refl : constructor option } val empty : retroknowledge type action = - | Register_ind of CPrimitives.prim_ind * inductive - | Register_type of CPrimitives.prim_type * Constant.t + | Register_ind : 'a CPrimitives.prim_ind * inductive -> action + | Register_type : CPrimitives.prim_type * Constant.t -> action diff --git a/kernel/retypeops.ml b/kernel/retypeops.ml index f398e6a5da..5c15257511 100644 --- a/kernel/retypeops.ml +++ b/kernel/retypeops.ml @@ -60,7 +60,7 @@ let rec relevance_of_fterm env extra lft f = | FRel n -> relevance_of_rel_extra env extra (Esubst.reloc_rel n lft) | FAtom c -> relevance_of_term_extra env extra lft (Esubst.subs_id 0) c | FFlex key -> relevance_of_flex env extra lft key - | FInt _ -> Sorts.Relevant + | FInt _ | FFloat _ -> Sorts.Relevant | FInd _ | FProd _ -> Sorts.Relevant (* types are always relevant *) | FConstruct (c,_) -> relevance_of_constructor env c | FApp (f, _) -> relevance_of_fterm env extra lft f @@ -105,7 +105,7 @@ and relevance_of_term_extra env extra lft subs c = | Fix ((_,i),(lna,_,_)) -> (lna.(i)).binder_relevance | CoFix (i,(lna,_,_)) -> (lna.(i)).binder_relevance | Proj (p, _) -> relevance_of_projection env p - | Int _ -> Sorts.Relevant + | Int _ | Float _ -> Sorts.Relevant | Meta _ | Evar _ -> Sorts.Relevant (* let's assume metas and evars are relevant for now *) diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml index e846b17aa0..d3cffd1546 100644 --- a/kernel/safe_typing.ml +++ b/kernel/safe_typing.ml @@ -1327,7 +1327,7 @@ let register_inline kn senv = let cb = {cb with const_inline_code = true} in let env = add_constant kn cb env in { senv with env} -let check_register_ind ind r env = +let check_register_ind (type t) ind (r : t CPrimitives.prim_ind) env = let (mb,ob as spec) = Inductive.lookup_mind_specif env ind in let check_if b msg = if not b then @@ -1403,6 +1403,36 @@ let check_register_ind ind r env = check_type_cte 1; check_name 2 "Gt"; check_type_cte 2 + | CPrimitives.PIT_f_cmp -> + check_nconstr 4; + check_name 0 "FEq"; + check_type_cte 0; + check_name 1 "FLt"; + check_type_cte 1; + check_name 2 "FGt"; + check_type_cte 2; + check_name 3 "FNotComparable"; + check_type_cte 3 + | CPrimitives.PIT_f_class -> + check_nconstr 9; + check_name 0 "PNormal"; + check_type_cte 0; + check_name 1 "NNormal"; + check_type_cte 1; + check_name 2 "PSubn"; + check_type_cte 2; + check_name 3 "NSubn"; + check_type_cte 3; + check_name 4 "PZero"; + check_type_cte 4; + check_name 5 "NZero"; + check_type_cte 5; + check_name 6 "PInf"; + check_type_cte 6; + check_name 7 "NInf"; + check_type_cte 7; + check_name 8 "NaN"; + check_type_cte 8 let register_inductive ind prim senv = check_register_ind ind prim senv.env; diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli index b2f6668577..ae6993b0e2 100644 --- a/kernel/safe_typing.mli +++ b/kernel/safe_typing.mli @@ -226,7 +226,7 @@ val mind_of_delta_kn_senv : safe_environment -> KerName.t -> MutInd.t (** {6 Retroknowledge / Native compiler } *) val register_inline : Constant.t -> safe_transformer0 -val register_inductive : inductive -> CPrimitives.prim_ind -> safe_transformer0 +val register_inductive : inductive -> 'a CPrimitives.prim_ind -> safe_transformer0 val set_strategy : Names.Constant.t Names.tableKey -> Conv_oracle.level -> safe_transformer0 diff --git a/kernel/term.ml b/kernel/term.ml index 38c0d043cf..7343507838 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -383,4 +383,4 @@ let kind_of_type t = match kind t with | (Rel _ | Meta _ | Var _ | Evar _ | Const _ | Proj _ | Case _ | Fix _ | CoFix _ | Ind _) -> AtomicType (t,[||]) - | (Lambda _ | Construct _ | Int _) -> failwith "Not a type" + | (Lambda _ | Construct _ | Int _ | Float _) -> failwith "Not a type" diff --git a/kernel/typeops.ml b/kernel/typeops.ml index b87384d228..1cc40a6707 100644 --- a/kernel/typeops.ml +++ b/kernel/typeops.ml @@ -215,14 +215,22 @@ let type_of_apply env func funt argsv argstv = (* Type of primitive constructs *) let type_of_prim_type _env = function | CPrimitives.PT_int63 -> Constr.mkSet + | CPrimitives.PT_float64 -> Constr.mkSet let type_of_int env = match env.retroknowledge.Retroknowledge.retro_int63 with | Some c -> mkConst c | None -> CErrors.user_err Pp.(str"The type int must be registered before this construction can be typechecked.") +let type_of_float env = + match env.retroknowledge.Retroknowledge.retro_float64 with + | Some c -> mkConst c + | None -> raise + (Invalid_argument "Typeops.type_of_float: float64 not_defined") + let type_of_prim env t = - let int_ty = type_of_int env in + let int_ty () = type_of_int env in + let float_ty () = type_of_float env in let bool_ty () = match env.retroknowledge.Retroknowledge.retro_bool with | Some ((ind,_),_) -> Constr.mkInd ind @@ -233,6 +241,16 @@ let type_of_prim env t = | Some ((ind,_),_,_) -> Constr.mkInd ind | None -> CErrors.user_err Pp.(str"The type compare must be registered before this primitive.") in + let f_compare_ty () = + match env.retroknowledge.Retroknowledge.retro_f_cmp with + | Some ((ind,_),_,_,_) -> Constr.mkInd ind + | None -> CErrors.user_err Pp.(str"The type float_comparison must be registered before this primitive.") + in + let f_class_ty () = + match env.retroknowledge.Retroknowledge.retro_f_class with + | Some ((ind,_),_,_,_,_,_,_,_,_) -> Constr.mkInd ind + | None -> CErrors.user_err Pp.(str"The type float_class must be registered before this primitive.") + in let pair_ty fst_ty snd_ty = match env.retroknowledge.Retroknowledge.retro_pair with | Some (ind,_) -> Constr.mkApp(Constr.mkInd ind, [|fst_ty;snd_ty|]) @@ -243,39 +261,27 @@ let type_of_prim env t = | Some ((ind,_),_) -> Constr.mkApp(Constr.mkInd ind, [|int_ty|]) | None -> CErrors.user_err Pp.(str"The type carry must be registered before this primitive.") in - let rec nary_int63_op arity ty = - if Int.equal arity 0 then ty - else Constr.mkProd(Context.nameR (Id.of_string "x"), int_ty, nary_int63_op (arity-1) ty) - in - let return_ty = - let open CPrimitives in - match t with - | Int63head0 - | Int63tail0 - | Int63add - | Int63sub - | Int63mul - | Int63div - | Int63mod - | Int63lsr - | Int63lsl - | Int63land - | Int63lor - | Int63lxor - | Int63addMulDiv -> int_ty - | Int63eq - | Int63lt - | Int63le -> bool_ty () - | Int63mulc - | Int63div21 - | Int63diveucl -> pair_ty int_ty int_ty - | Int63addc - | Int63subc - | Int63addCarryC - | Int63subCarryC -> carry_ty int_ty - | Int63compare -> compare_ty () - in - nary_int63_op (CPrimitives.arity t) return_ty + let open CPrimitives in + let tr_prim_type = function + | PT_int63 -> int_ty () + | PT_float64 -> float_ty () in + let tr_ind (type t) (i : t prim_ind) (a : t) = match i, a with + | PIT_bool, () -> bool_ty () + | PIT_carry, t -> carry_ty (tr_prim_type t) + | PIT_pair, (t1, t2) -> pair_ty (tr_prim_type t1) (tr_prim_type t2) + | PIT_cmp, () -> compare_ty () + | PIT_f_cmp, () -> f_compare_ty () + | PIT_f_class, () -> f_class_ty () in + let tr_type = function + | PITT_ind (i, a) -> tr_ind i a + | PITT_type t -> tr_prim_type t in + let rec nary_op = function + | [] -> assert false + | [ret_ty] -> tr_type ret_ty + | arg_ty :: r -> + let arg_ty = tr_type arg_ty in + Constr.mkProd(Context.nameR (Id.of_string "x"), arg_ty, nary_op r) in + nary_op (types t) let type_of_prim_or_type env = let open CPrimitives in function @@ -285,6 +291,9 @@ let type_of_prim_or_type env = let open CPrimitives in let judge_of_int env i = make_judge (Constr.mkInt i) (type_of_int env) +let judge_of_float env f = + make_judge (Constr.mkFloat f) (type_of_float env) + (* Type of product *) let sort_of_product env domsort rangsort = @@ -583,6 +592,7 @@ let rec execute env cstr = (* Primitive types *) | Int _ -> cstr, type_of_int env + | Float _ -> cstr, type_of_float env (* Partial proofs: unsupported by the kernel *) | Meta _ -> diff --git a/kernel/typeops.mli b/kernel/typeops.mli index c71a0e0ca4..ae816fe26e 100644 --- a/kernel/typeops.mli +++ b/kernel/typeops.mli @@ -120,6 +120,9 @@ val check_primitive_type : env -> CPrimitives.op_or_type -> types -> unit val type_of_int : env -> types val judge_of_int : env -> Uint63.t -> unsafe_judgment +val type_of_float : env -> types +val judge_of_float : env -> Float64.t -> unsafe_judgment + val type_of_prim_type : env -> CPrimitives.prim_type -> types val type_of_prim : env -> CPrimitives.t -> types diff --git a/kernel/uint63.mli b/kernel/uint63.mli index d22ba3468f..e0bf44da35 100644 --- a/kernel/uint63.mli +++ b/kernel/uint63.mli @@ -19,7 +19,14 @@ val to_int2 : t -> int * int (* msb, lsb *) val of_int64 : Int64.t -> t (* val of_uint : int -> t -*) + *) +(** [int_min n m] returns the minimum of [n] and [m], + [m] must be in [0, 2^30-1]. *) +val to_int_min : t -> int -> int + + (* conversion to float *) +val of_float : float -> t +val to_float : t -> float val hash : t -> int diff --git a/kernel/uint63_31.ml b/kernel/uint63_31.ml index b8eccd19fb..e38389ca13 100644 --- a/kernel/uint63_31.ml +++ b/kernel/uint63_31.ml @@ -26,6 +26,13 @@ let mask63 i = Int64.logand i maxuint63 let of_int i = Int64.of_int i let to_int2 i = (Int64.to_int (Int64.shift_right_logical i 31), Int64.to_int i) let of_int64 i = i + +let to_int_min n m = + if Int64.(compare n (of_int m)) < 0 then Int64.to_int n else m + +let of_float f = mask63 (Int64.of_float f) +let to_float = Int64.to_float + let hash i = let (h,l) = to_int2 i in (*Hashset.combine h l*) @@ -213,4 +220,8 @@ let () = Callback.register "uint63 one" one; Callback.register "uint63 sub" sub; Callback.register "uint63 subcarry" subcarry; - Callback.register "uint63 tail0" tail0 + Callback.register "uint63 tail0" tail0; + Callback.register "uint63 of_float" of_float; + Callback.register "uint63 to_float" to_float; + Callback.register "uint63 of_int" of_int; + Callback.register "uint63 to_int_min" to_int_min diff --git a/kernel/uint63_63.ml b/kernel/uint63_63.ml index 5c4028e1c8..85b44528a7 100644 --- a/kernel/uint63_63.ml +++ b/kernel/uint63_63.ml @@ -27,6 +27,12 @@ let to_int2 i = (0,i) let of_int64 _i = assert false +let of_float = int_of_float + +external to_float : int -> (float [@unboxed]) + = "coq_uint63_to_float_byte" "coq_uint63_to_float" +[@@noalloc] + let hash i = i [@@ocaml.inline always] @@ -96,6 +102,10 @@ let le (x : int) (y : int) = (x lxor 0x4000000000000000) <= (y lxor 0x4000000000000000) [@@ocaml.inline always] +let to_int_min n m = + if lt n m then n else m +[@@ocaml.inline always] + (* division of two numbers by one *) (* precondition: xh < y *) (* outputs: q, r s.t. x = q * y + r, r < y *) diff --git a/kernel/vconv.ml b/kernel/vconv.ml index 414c443c4e..5d36ad54a2 100644 --- a/kernel/vconv.ml +++ b/kernel/vconv.ml @@ -73,6 +73,9 @@ and conv_whd env pb k whd1 whd2 cu = else raise NotConvertible | Vint64 i1, Vint64 i2 -> if Int64.equal i1 i2 then cu else raise NotConvertible + | Vfloat64 f1, Vfloat64 f2 -> + if Float64.(equal (of_float f1) (of_float f2)) then cu + else raise NotConvertible | Vatom_stk(a1,stk1), Vatom_stk(a2,stk2) -> conv_atom env pb k a1 stk1 a2 stk2 cu | Vfun _, _ | _, Vfun _ -> @@ -80,7 +83,7 @@ and conv_whd env pb k whd1 whd2 cu = conv_val env CONV (k+1) (apply_whd k whd1) (apply_whd k whd2) cu | Vprod _, _ | Vfix _, _ | Vcofix _, _ | Vconstr_const _, _ | Vint64 _, _ - | Vconstr_block _, _ | Vatom_stk _, _ -> raise NotConvertible + | Vfloat64 _, _ | Vconstr_block _, _ | Vatom_stk _, _ -> raise NotConvertible and conv_atom env pb k a1 stk1 a2 stk2 cu = diff --git a/kernel/vm.ml b/kernel/vm.ml index 319a26d824..5f08720f77 100644 --- a/kernel/vm.ml +++ b/kernel/vm.ml @@ -169,7 +169,8 @@ let rec apply_stack a stk v = let apply_whd k whd = let v = val_of_rel k in match whd with - | Vprod _ | Vconstr_const _ | Vconstr_block _ | Vint64 _ -> assert false + | Vprod _ | Vconstr_const _ | Vconstr_block _ | Vint64 _ | Vfloat64 _ -> + assert false | Vfun f -> reduce_fun k f | Vfix(f, None) -> push_ra stop; diff --git a/kernel/vmvalues.ml b/kernel/vmvalues.ml index c8f5020d71..5acdd964b1 100644 --- a/kernel/vmvalues.ml +++ b/kernel/vmvalues.ml @@ -57,6 +57,7 @@ type structured_constant = | Const_univ_level of Univ.Level.t | Const_val of structured_values | Const_uint of Uint63.t + | Const_float of Float64.t type reloc_table = (tag * int) array @@ -75,6 +76,8 @@ let rec eq_structured_values v1 v2 = Int.equal (Obj.size o1) (Obj.size o2) then if Int.equal t1 Obj.custom_tag then Int64.equal (Obj.magic v1 : int64) (Obj.magic v2 : int64) + else if Int.equal t1 Obj.double_tag + then Float64.(equal (of_float (Obj.magic v1)) (of_float (Obj.magic v2))) else begin assert (t1 <= Obj.last_non_constant_constructor_tag && t2 <= Obj.last_non_constant_constructor_tag); @@ -105,6 +108,8 @@ let eq_structured_constant c1 c2 = match c1, c2 with | Const_val _, _ -> false | Const_uint i1, Const_uint i2 -> Uint63.equal i1 i2 | Const_uint _, _ -> false +| Const_float f1, Const_float f2 -> Float64.equal f1 f2 +| Const_float _, _ -> false let hash_structured_constant c = let open Hashset.Combine in @@ -115,6 +120,7 @@ let hash_structured_constant c = | Const_univ_level l -> combinesmall 4 (Univ.Level.hash l) | Const_val v -> combinesmall 5 (hash_structured_values v) | Const_uint i -> combinesmall 6 (Uint63.hash i) + | Const_float f -> combinesmall 7 (Float64.hash f) let eq_annot_switch asw1 asw2 = let eq_ci ci1 ci2 = @@ -149,6 +155,7 @@ let pp_struct_const = function | Const_univ_level l -> Univ.Level.pr l | Const_val _ -> Pp.str "(value)" | Const_uint i -> Pp.str (Uint63.to_string i) + | Const_float f -> Pp.str (Float64.to_string f) (* Abstract data *) type vprod @@ -284,6 +291,7 @@ type whd = | Vconstr_const of int | Vconstr_block of vblock | Vint64 of int64 + | Vfloat64 of float | Vatom_stk of atom * stack | Vuniv_level of Univ.Level.t @@ -315,6 +323,7 @@ let uni_lvl_val (v : values) : Univ.Level.t = | Vconstr_const _i -> str "Vconstr_const" | Vconstr_block _b -> str "Vconstr_block" | Vint64 _ -> str "Vint64" + | Vfloat64 _ -> str "Vfloat64" | Vatom_stk (_a,_stk) -> str "Vatom_stk" | Vuniv_level _ -> assert false in @@ -374,6 +383,8 @@ let rec whd_accu a stk = end | i when Int.equal i Obj.custom_tag -> Vint64 (Obj.magic i) + | i when Int.equal i Obj.double_tag -> + Vfloat64 (Obj.magic i) | tg -> CErrors.anomaly Pp.(strbrk "Failed to parse VM value. Tag = " ++ int tg ++ str ".") @@ -403,6 +414,7 @@ let whd_val : values -> whd = | 3 -> Vatom_stk(Aid(RelKey(int_tcode (fun_code o) 1)), []) | _ -> CErrors.anomaly ~label:"Vm.whd " (Pp.str "kind_of_closure does not work.")) else if Int.equal tag Obj.custom_tag then Vint64 (Obj.magic v) + else if Int.equal tag Obj.double_tag then Vfloat64 (Obj.magic v) else Vconstr_block(Obj.obj o) @@ -426,6 +438,7 @@ let obj_of_str_const str = | Const_univ_level l -> Obj.repr (Vuniv_level l) | Const_val v -> Obj.repr v | Const_uint i -> Obj.repr i + | Const_float f -> Obj.repr f let val_of_block tag (args : structured_values array) = let nargs = Array.length args in @@ -675,6 +688,7 @@ and pr_whd w = | Vconstr_const i -> str "Vconstr_const(" ++ int i ++ str ")" | Vconstr_block _b -> str "Vconstr_block" | Vint64 i -> i |> Format.sprintf "Vint64(%LiL)" |> str + | Vfloat64 f -> str "Vfloat64(" ++ str (Float64.(to_string (of_float f))) ++ str ")" | Vatom_stk (a,stk) -> str "Vatom_stk(" ++ pr_atom a ++ str ", " ++ pr_stack stk ++ str ")" | Vuniv_level _ -> assert false) and pr_stack stk = diff --git a/kernel/vmvalues.mli b/kernel/vmvalues.mli index d289e7db9a..9c24006ff0 100644 --- a/kernel/vmvalues.mli +++ b/kernel/vmvalues.mli @@ -45,6 +45,7 @@ type structured_constant = | Const_univ_level of Univ.Level.t | Const_val of structured_values | Const_uint of Uint63.t + | Const_float of Float64.t val pp_struct_const : structured_constant -> Pp.t @@ -127,6 +128,7 @@ type whd = | Vconstr_const of int | Vconstr_block of vblock | Vint64 of int64 + | Vfloat64 of float | Vatom_stk of atom * stack | Vuniv_level of Univ.Level.t diff --git a/library/global.mli b/library/global.mli index f8b1f35f4d..0570ad0102 100644 --- a/library/global.mli +++ b/library/global.mli @@ -157,7 +157,7 @@ val is_type_in_type : GlobRef.t -> bool (** {6 Retroknowledge } *) val register_inline : Constant.t -> unit -val register_inductive : inductive -> CPrimitives.prim_ind -> unit +val register_inductive : inductive -> 'a CPrimitives.prim_ind -> unit (** {6 Oracle } *) diff --git a/plugins/extraction/ExtrOCamlFloats.v b/plugins/extraction/ExtrOCamlFloats.v new file mode 100644 index 0000000000..1891772cc2 --- /dev/null +++ b/plugins/extraction/ExtrOCamlFloats.v @@ -0,0 +1,61 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(* * (see LICENSE file for the text of the license) *) +(************************************************************************) + +(** Extraction to OCaml of native binary64 floating-point numbers. + +Note: the extraction of primitive floats relies on Coq's internal file +kernel/float64.ml, so make sure the corresponding binary is available +when linking the extracted OCaml code. + +For example, if you build a (_CoqProject + coq_makefile)-based project +and if you created an empty subfolder "extracted" and a file "test.v" +containing [Cd "extracted". Separate Extraction function_to_extract.], +you will just need to add in the _CoqProject: [test.v], [-I extracted] +and the list of [extracted/*.ml] and [extracted/*.mli] files, then add +[CAMLFLAGS += -w -33] in the Makefile.local file. *) + +From Coq Require Floats Extraction. + +(** Basic data types used by some primitive operators. *) + +Extract Inductive bool => bool [ true false ]. +Extract Inductive prod => "( * )" [ "" ]. + +Extract Inductive FloatClass.float_class => + "Float64.float_class" + [ "PNormal" "NNormal" "PSubn" "NSubn" "PZero" "NZero" "PInf" "NInf" "NaN" ]. +Extract Inductive PrimFloat.float_comparison => + "Float64.float_comparison" + [ "FEq" "FLt" "FGt" "FNotComparable" ]. + +(** Primitive types and operators. *) + +Extract Constant PrimFloat.float => "Float64.t". +Extraction Inline PrimFloat.float. +(* Otherwise, the name conflicts with the primitive OCaml type [float] *) + +Extract Constant PrimFloat.classify => "Float64.classify". +Extract Constant PrimFloat.abs => "Float64.abs". +Extract Constant PrimFloat.sqrt => "Float64.sqrt". +Extract Constant PrimFloat.opp => "Float64.opp". +Extract Constant PrimFloat.eqb => "Float64.eq". +Extract Constant PrimFloat.ltb => "Float64.lt". +Extract Constant PrimFloat.leb => "Float64.le". +Extract Constant PrimFloat.compare => "Float64.compare". +Extract Constant PrimFloat.mul => "Float64.mul". +Extract Constant PrimFloat.add => "Float64.add". +Extract Constant PrimFloat.sub => "Float64.sub". +Extract Constant PrimFloat.div => "Float64.div". +Extract Constant PrimFloat.of_int63 => "Float64.of_int63". +Extract Constant PrimFloat.normfr_mantissa => "Float64.normfr_mantissa". +Extract Constant PrimFloat.frshiftexp => "Float64.frshiftexp". +Extract Constant PrimFloat.ldshiftexp => "Float64.ldshiftexp". +Extract Constant PrimFloat.next_up => "Float64.next_up". +Extract Constant PrimFloat.next_down => "Float64.next_down". diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index cca212f332..04f5b66241 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -351,7 +351,7 @@ let rec extract_type env sg db j c args = | (Info, TypeScheme) -> extract_type_app env sg db (r, type_sign env sg ty) args | (Info, Default) -> Tunknown)) - | Cast _ | LetIn _ | Construct _ | Int _ -> assert false + | Cast _ | LetIn _ | Construct _ | Int _ | Float _ -> assert false (*s Auxiliary function dealing with type application. Precondition: [r] is a type scheme represented by the signature [s], @@ -690,6 +690,7 @@ let rec extract_term env sg mle mlt c args = let extract_var mlt = put_magic (mlt,vty) (MLglob (GlobRef.VarRef v)) in extract_app env sg mle mlt extract_var args | Int i -> assert (args = []); MLuint i + | Float f -> assert (args = []); MLfloat f | Ind _ | Prod _ | Sort _ -> assert false (*s [extract_maybe_term] is [extract_term] for usual terms, else [MLdummy] *) diff --git a/plugins/extraction/haskell.ml b/plugins/extraction/haskell.ml index e4efbcff0c..4769bef475 100644 --- a/plugins/extraction/haskell.ml +++ b/plugins/extraction/haskell.ml @@ -215,6 +215,8 @@ let rec pp_expr par env args = | MLaxiom -> pp_par par (str "Prelude.error \"AXIOM TO BE REALIZED\"") | MLuint _ -> pp_par par (str "Prelude.error \"EXTRACTION OF UINT NOT IMPLEMENTED\"") + | MLfloat _ -> + pp_par par (str "Prelude.error \"EXTRACTION OF FLOAT NOT IMPLEMENTED\"") and pp_cons_pat par r ppl = pp_par par diff --git a/plugins/extraction/json.ml b/plugins/extraction/json.ml index 912a20f389..81b3e1bcdc 100644 --- a/plugins/extraction/json.ml +++ b/plugins/extraction/json.ml @@ -161,6 +161,10 @@ let rec json_expr env = function ("what", json_str "expr:int"); ("int", json_str (Uint63.to_string i)) ] + | MLfloat f -> json_dict [ + ("what", json_str "expr:float"); + ("float", json_str (Float64.to_string f)) + ] and json_one_pat env (ids,p,t) = let ids', env' = push_vars (List.rev_map id_of_mlid ids) env in json_dict [ diff --git a/plugins/extraction/miniml.ml b/plugins/extraction/miniml.ml index 8b69edbe4c..32e0d3c05d 100644 --- a/plugins/extraction/miniml.ml +++ b/plugins/extraction/miniml.ml @@ -126,7 +126,8 @@ and ml_ast = | MLdummy of kill_reason | MLaxiom | MLmagic of ml_ast - | MLuint of Uint63.t + | MLuint of Uint63.t + | MLfloat of Float64.t and ml_pattern = | Pcons of GlobRef.t * ml_pattern list diff --git a/plugins/extraction/miniml.mli b/plugins/extraction/miniml.mli index e3c9635c55..32e0d3c05d 100644 --- a/plugins/extraction/miniml.mli +++ b/plugins/extraction/miniml.mli @@ -127,6 +127,7 @@ and ml_ast = | MLaxiom | MLmagic of ml_ast | MLuint of Uint63.t + | MLfloat of Float64.t and ml_pattern = | Pcons of GlobRef.t * ml_pattern list diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml index 000df26858..44b95ae4c1 100644 --- a/plugins/extraction/mlutil.ml +++ b/plugins/extraction/mlutil.ml @@ -398,6 +398,7 @@ let rec eq_ml_ast t1 t2 = match t1, t2 with | MLaxiom, MLaxiom -> true | MLmagic t1, MLmagic t2 -> eq_ml_ast t1 t2 | MLuint i1, MLuint i2 -> Uint63.equal i1 i2 +| MLfloat f1, MLfloat f2 -> Float64.equal f1 f2 | _, _ -> false and eq_ml_pattern p1 p2 = match p1, p2 with @@ -430,7 +431,7 @@ let ast_iter_rel f = | MLapp (a,l) -> iter n a; List.iter (iter n) l | MLcons (_,_,l) | MLtuple l -> List.iter (iter n) l | MLmagic a -> iter n a - | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> () + | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ -> () in iter 0 (*s Map over asts. *) @@ -449,7 +450,8 @@ let ast_map f = function | MLcons (typ,c,l) -> MLcons (typ,c, List.map f l) | MLtuple l -> MLtuple (List.map f l) | MLmagic a -> MLmagic (f a) - | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ as a -> a + | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom + | MLuint _ | MLfloat _ as a -> a (*s Map over asts, with binding depth as parameter. *) @@ -467,7 +469,8 @@ let ast_map_lift f n = function | MLcons (typ,c,l) -> MLcons (typ,c, List.map (f n) l) | MLtuple l -> MLtuple (List.map (f n) l) | MLmagic a -> MLmagic (f n a) - | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ as a -> a + | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom + | MLuint _ | MLfloat _ as a -> a (*s Iter over asts. *) @@ -481,7 +484,8 @@ let ast_iter f = function | MLapp (a,l) -> f a; List.iter f l | MLcons (_,_,l) | MLtuple l -> List.iter f l | MLmagic a -> f a - | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> () + | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom + | MLuint _ | MLfloat _ -> () (*S Operations concerning De Bruijn indices. *) @@ -517,7 +521,7 @@ let nb_occur_match = | MLapp (a,l) -> List.fold_left (fun r a -> r+(nb k a)) (nb k a) l | MLcons (_,_,l) | MLtuple l -> List.fold_left (fun r a -> r+(nb k a)) 0 l | MLmagic a -> nb k a - | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> 0 + | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ -> 0 in nb 1 (* Replace unused variables by _ *) @@ -569,7 +573,7 @@ let dump_unused_vars a = let b' = ren env b in if b' == b then a else MLmagic b' - | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> a + | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ | MLfloat _ -> a and ren_branch env ((ids,p,b) as tr) = let occs = List.map (fun _ -> ref false) ids in @@ -1402,7 +1406,8 @@ let rec ml_size = function | MLfix(_,_,f) -> ml_size_array f | MLletin (_,_,t) -> ml_size t | MLmagic t -> ml_size t - | MLglob _ | MLrel _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> 0 + | MLglob _ | MLrel _ | MLexn _ | MLdummy _ | MLaxiom + | MLuint _ | MLfloat _ -> 0 and ml_size_list l = List.fold_left (fun a t -> a + ml_size t) 0 l diff --git a/plugins/extraction/modutil.ml b/plugins/extraction/modutil.ml index 6b1eef7abb..fe49bfc1ec 100644 --- a/plugins/extraction/modutil.ml +++ b/plugins/extraction/modutil.ml @@ -107,7 +107,7 @@ let ast_iter_references do_term do_cons do_type a = Array.iter (fun (_,p,_) -> patt_iter_references do_cons p) v | MLrel _ | MLlam _ | MLapp _ | MLletin _ | MLtuple _ | MLfix _ | MLexn _ - | MLdummy _ | MLaxiom | MLmagic _ | MLuint _ -> () + | MLdummy _ | MLaxiom | MLmagic _ | MLuint _ | MLfloat _ -> () in iter a let ind_iter_references do_term do_cons do_type kn ind = diff --git a/plugins/extraction/ocaml.ml b/plugins/extraction/ocaml.ml index e7004fe9af..34ddf57b40 100644 --- a/plugins/extraction/ocaml.ml +++ b/plugins/extraction/ocaml.ml @@ -307,6 +307,9 @@ let rec pp_expr par env args = | MLuint i -> assert (args=[]); str "(" ++ str (Uint63.compile i) ++ str ")" + | MLfloat f -> + assert (args=[]); + str "(" ++ str (Float64.compile f) ++ str ")" and pp_record_proj par env typ t pv args = diff --git a/plugins/extraction/scheme.ml b/plugins/extraction/scheme.ml index dd840cd929..c341ec8d57 100644 --- a/plugins/extraction/scheme.ml +++ b/plugins/extraction/scheme.ml @@ -131,6 +131,8 @@ let rec pp_expr env args = | MLaxiom -> paren (str "error \"AXIOM TO BE REALIZED\"") | MLuint _ -> paren (str "Prelude.error \"EXTRACTION OF UINT NOT IMPLEMENTED\"") + | MLfloat _ -> + paren (str "Prelude.error \"EXTRACTION OF FLOAT NOT IMPLEMENTED\"") and pp_cons_args env = function | MLcons (_,r,args) when is_coinductive r -> diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 7be049269c..6db0a1119b 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -692,13 +692,14 @@ let build_proof end | Cast(t,_,_) -> build_proof do_finalize {dyn_infos with info = t} g - | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ -> + | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ | Float _ -> do_finalize dyn_infos g | App(_,_) -> let f,args = decompose_app sigma dyn_infos.info in begin match EConstr.kind sigma f with - | Int _ -> user_err Pp.(str "integer cannot be applied") + | Int _ -> user_err Pp.(str "integer cannot be applied") + | Float _ -> user_err Pp.(str "float cannot be applied") | App _ -> assert false (* we have collected all the app in decompose_app *) | Proj _ -> assert false (*FIXME*) | Var _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _ | Sort _ | Prod _ -> diff --git a/plugins/funind/gen_principle.ml b/plugins/funind/gen_principle.ml index 0452665585..6add56dd5b 100644 --- a/plugins/funind/gen_principle.ml +++ b/plugins/funind/gen_principle.ml @@ -68,7 +68,7 @@ let is_rec names = let check_id id names = Id.Set.mem id names in let rec lookup names gt = match DAst.get gt with | GVar(id) -> check_id id names - | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ -> false + | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ | GFloat _ -> false | GCast(b,_) -> lookup names b | GRec _ -> CErrors.user_err (Pp.str "GRec not handled") | GIf(b,_,lhs,rhs) -> diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 7c17ecdba0..895b6a37ee 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -478,7 +478,7 @@ let rec build_entry_lc env sigma funnames avoid rt : glob_constr build_entry_ret observe (str " Entering : " ++ Printer.pr_glob_constr_env env rt); let open CAst in match DAst.get rt with - | GRef _ | GVar _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ -> + | GRef _ | GVar _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ | GFloat _ -> (* do nothing (except changing type of course) *) mk_result [] rt avoid | GApp(_,_) -> @@ -590,6 +590,7 @@ let rec build_entry_lc env sigma funnames avoid rt : glob_constr build_entry_ret | GRec _ -> user_err Pp.(str "Not handled GRec") | GProd _ -> user_err Pp.(str "Cannot apply a type") | GInt _ -> user_err Pp.(str "Cannot apply an integer") + | GFloat _ -> user_err Pp.(str "Cannot apply a float") end (* end of the application treatement *) | GLambda(n,_,t,b) -> @@ -1231,7 +1232,7 @@ let rebuild_cons env nb_args relname args crossed_types rt = TODO: Find a valid way to deal with implicit arguments here! *) let rec compute_cst_params relnames params gt = DAst.with_val (function - | GRef _ | GVar _ | GEvar _ | GPatVar _ | GInt _ -> params + | GRef _ | GVar _ | GEvar _ | GPatVar _ | GInt _ | GFloat _ -> params | GApp(f,args) -> begin match DAst.get f with | GVar relname' when Id.Set.mem relname' relnames -> diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml index 8abccabae6..5f54bad598 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -115,6 +115,7 @@ let change_vars = | GSort _ as x -> x | GHole _ as x -> x | GInt _ as x -> x + | GFloat _ as x -> x | GCast(b,c) -> GCast(change_vars mapping b, Glob_ops.map_cast_type (change_vars mapping) c) @@ -295,6 +296,7 @@ let rec alpha_rt excluded rt = | GRec _ -> user_err Pp.(str "Not handled GRec") | GSort _ | GInt _ + | GFloat _ | GHole _ as rt -> rt | GCast (b,c) -> GCast(alpha_rt excluded b, @@ -354,7 +356,7 @@ let is_free_in id = | GHole _ -> false | GCast (b,(CastConv t|CastVM t|CastNative t)) -> is_free_in b || is_free_in t | GCast (b,CastCoerce) -> is_free_in b - | GInt _ -> false + | GInt _ | GFloat _ -> false ) x and is_free_in_br {CAst.v=(ids,_,rt)} = (not (Id.List.mem id ids)) && is_free_in rt @@ -447,6 +449,7 @@ let replace_var_by_term x_id term = | GSort _ | GHole _ as rt -> rt | GInt _ as rt -> rt + | GFloat _ as rt -> rt | GCast(b,c) -> GCast(replace_var_by_pattern b, Glob_ops.map_cast_type replace_var_by_pattern c) @@ -529,7 +532,7 @@ let expand_as = | PatCstr(_,patl,_) -> List.fold_left add_as map patl in let rec expand_as map = DAst.map (function - | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ as rt -> rt + | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ | GFloat _ as rt -> rt | GVar id as rt -> begin try diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 29356df81d..66ed1961ba 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -270,7 +270,7 @@ let check_not_nested env sigma forbidden e = let rec check_not_nested e = match EConstr.kind sigma e with | Rel _ -> () - | Int _ -> () + | Int _ | Float _ -> () | Var x -> if Id.List.mem x forbidden then user_err ~hdr:"Recdef.check_not_nested" @@ -452,7 +452,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g = | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_leconstr_env env sigma expr_info.info ++ Pp.str ".") end | Cast(t,_,_) -> travel jinfo continuation_tac {expr_info with info=t} g - | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ -> + | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ | Float _ -> let new_continuation_tac = jinfo.otherS () expr_info continuation_tac in new_continuation_tac expr_info g diff --git a/plugins/ssrmatching/ssrmatching.ml b/plugins/ssrmatching/ssrmatching.ml index 4d7a04f5ee..9682487a22 100644 --- a/plugins/ssrmatching/ssrmatching.ml +++ b/plugins/ssrmatching/ssrmatching.ml @@ -319,7 +319,7 @@ let iter_constr_LR f c = match kind c with for i = 0 to Array.length t - 1 do f t.(i); f b.(i) done | Proj(_,a) -> f a | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ - | Int _) -> () + | Int _ | Float _) -> () (* The comparison used to determine which subterms matches is KEYED *) (* CONVERSION. This looks for convertible terms that either have the same *) diff --git a/plugins/syntax/float_syntax.ml b/plugins/syntax/float_syntax.ml new file mode 100644 index 0000000000..3c2e217d1c --- /dev/null +++ b/plugins/syntax/float_syntax.ml @@ -0,0 +1,50 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +open Names +open Glob_term + +(* Poor's man DECLARE PLUGIN *) +let __coq_plugin_name = "float_syntax_plugin" +let () = Mltop.add_known_module __coq_plugin_name + +(*** Constants for locating float constructors ***) + +let make_dir l = DirPath.make (List.rev_map Id.of_string l) +let make_path dir id = Libnames.make_path (make_dir dir) (Id.of_string id) + +(*** Parsing for float in digital notation ***) + +let interp_float ?loc (sign,n) = + let sign = Constrexpr.(match sign with SPlus -> "" | SMinus -> "-") in + DAst.make ?loc (GFloat (Float64.of_string (sign ^ NumTok.to_string n))) + +(* Pretty printing is already handled in constrextern.ml *) + +let uninterp_float _ = None + +(* Actually declares the interpreter for float *) + +open Notation + +let at_declare_ml_module f x = + Mltop.declare_cache_obj (fun () -> f x) __coq_plugin_name + +let float_module = ["Coq"; "Floats"; "PrimFloat"] +let float_path = make_path float_module "float" +let float_scope = "float_scope" + +let _ = + register_rawnumeral_interpretation float_scope (interp_float,uninterp_float); + at_declare_ml_module enable_prim_token_interpretation + { pt_local = false; + pt_scope = float_scope; + pt_interp_info = Uid float_scope; + pt_required = (float_path,float_module); + pt_refs = []; + pt_in_match = false } diff --git a/plugins/syntax/float_syntax_plugin.mlpack b/plugins/syntax/float_syntax_plugin.mlpack new file mode 100644 index 0000000000..d69f49bcfe --- /dev/null +++ b/plugins/syntax/float_syntax_plugin.mlpack @@ -0,0 +1 @@ +Float_syntax diff --git a/plugins/syntax/plugin_base.dune b/plugins/syntax/plugin_base.dune index 7a23581768..512752135d 100644 --- a/plugins/syntax/plugin_base.dune +++ b/plugins/syntax/plugin_base.dune @@ -25,3 +25,10 @@ (synopsis "Coq syntax plugin: int63") (modules int63_syntax) (libraries coq.vernac)) + +(library + (name float_syntax_plugin) + (public_name coq.plugins.float_syntax) + (synopsis "Coq syntax plugin: float") + (modules float_syntax) + (libraries coq.vernac)) diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml index 43b94aed3d..c78f791a5a 100644 --- a/pretyping/cbv.ml +++ b/pretyping/cbv.ml @@ -220,14 +220,26 @@ module VNativeEntries = | _ -> raise Primred.NativeDestKO) | _ -> raise Primred.NativeDestKO + let get_float () e = + match e with + | VAL(_, cf) -> + (match kind cf with + | Float f -> f + | _ -> raise Primred.NativeDestKO) + | _ -> raise Primred.NativeDestKO + let mkInt env i = VAL(0, mkInt i) + let mkFloat env f = VAL(0, mkFloat f) + let mkBool env b = let (ct,cf) = get_bool_constructors env in CONSTR(Univ.in_punivs (if b then ct else cf), [||]) let int_ty env = VAL(0, mkConst @@ get_int_type env) + let float_ty env = VAL(0, mkConst @@ get_float_type env) + let mkCarry env b e = let (c0,c1) = get_carry_constructors env in CONSTR(Univ.in_punivs (if b then c1 else c0), [|int_ty env;e|]) @@ -237,6 +249,12 @@ module VNativeEntries = let c = get_pair_constructor env in CONSTR(Univ.in_punivs c, [|int_ty;int_ty;e1;e2|]) + let mkFloatIntPair env f i = + let float_ty = float_ty env in + let int_ty = int_ty env in + let c = get_pair_constructor env in + CONSTR(Univ.in_punivs c, [|float_ty;int_ty;f;i|]) + let mkLt env = let (_eq,lt,_gt) = get_cmp_constructors env in CONSTR(Univ.in_punivs lt, [||]) @@ -249,6 +267,66 @@ module VNativeEntries = let (_eq,_lt,gt) = get_cmp_constructors env in CONSTR(Univ.in_punivs gt, [||]) + let mkFLt env = + let (_eq,lt,_gt,_nc) = get_f_cmp_constructors env in + CONSTR(Univ.in_punivs lt, [||]) + + let mkFEq env = + let (eq,_lt,_gt,_nc) = get_f_cmp_constructors env in + CONSTR(Univ.in_punivs eq, [||]) + + let mkFGt env = + let (_eq,_lt,gt,_nc) = get_f_cmp_constructors env in + CONSTR(Univ.in_punivs gt, [||]) + + let mkFNotComparable env = + let (_eq,_lt,_gt,nc) = get_f_cmp_constructors env in + CONSTR(Univ.in_punivs nc, [||]) + + let mkPNormal env = + let (pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + CONSTR(Univ.in_punivs pNormal, [||]) + + let mkNNormal env = + let (_pNormal,nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + CONSTR(Univ.in_punivs nNormal, [||]) + + let mkPSubn env = + let (_pNormal,_nNormal,pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + CONSTR(Univ.in_punivs pSubn, [||]) + + let mkNSubn env = + let (_pNormal,_nNormal,_pSubn,nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + CONSTR(Univ.in_punivs nSubn, [||]) + + let mkPZero env = + let (_pNormal,_nNormal,_pSubn,_nSubn,pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + CONSTR(Univ.in_punivs pZero, [||]) + + let mkNZero env = + let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + CONSTR(Univ.in_punivs nZero, [||]) + + let mkPInf env = + let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,pInf,_nInf,_nan) = + get_f_class_constructors env in + CONSTR(Univ.in_punivs pInf, [||]) + + let mkNInf env = + let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,nInf,_nan) = + get_f_class_constructors env in + CONSTR(Univ.in_punivs nInf, [||]) + + let mkNaN env = + let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,nan) = + get_f_class_constructors env in + CONSTR(Univ.in_punivs nan, [||]) end module VredNative = RedNative(VNativeEntries) @@ -381,7 +459,7 @@ let rec norm_head info env t stack = | Construct c -> (CONSTR(c, [||]), stack) (* neutral cases *) - | (Sort _ | Meta _ | Ind _ | Int _) -> (VAL(0, t), stack) + | (Sort _ | Meta _ | Ind _ | Int _ | Float _) -> (VAL(0, t), stack) | Prod _ -> (CBN(t,env), stack) and norm_head_ref k info env stack normt t = diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index e85c888b2e..d1cc21d82f 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -406,9 +406,10 @@ let matches_core env sigma allow_bound_rels | PEvar (c1,args1), Evar (c2,args2) when Evar.equal c1 c2 -> Array.fold_left2 (sorec ctx env) subst args1 args2 | PInt i1, Int i2 when Uint63.equal i1 i2 -> subst + | PFloat f1, Float f2 when Float64.equal f1 f2 -> subst | (PRef _ | PVar _ | PRel _ | PApp _ | PProj _ | PLambda _ | PProd _ | PLetIn _ | PSort _ | PIf _ | PCase _ - | PFix _ | PCoFix _| PEvar _ | PInt _), _ -> raise PatternMatchingFailure + | PFix _ | PCoFix _| PEvar _ | PInt _ | PFloat _), _ -> raise PatternMatchingFailure in sorec [] env ((Id.Map.empty,Id.Set.empty), Id.Map.empty) pat c @@ -526,7 +527,7 @@ let sub_match ?(closed=true) env sigma pat c = aux env term mk_ctx next with Retyping.RetypeError _ -> next () end - | Construct _| Ind _|Evar _|Const _ | Rel _|Meta _|Var _|Sort _ | Int _ -> + | Construct _|Ind _|Evar _|Const _|Rel _|Meta _|Var _|Sort _|Int _|Float _ -> next () in here next diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index e8c83c7de9..5dd4772bcc 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -834,6 +834,7 @@ and detype_r d flags avoid env sigma t = | Fix (nvn,recdef) -> detype_fix (detype d) flags avoid env sigma nvn recdef | CoFix (n,recdef) -> detype_cofix (detype d) flags avoid env sigma n recdef | Int i -> GInt i + | Float f -> GFloat f and detype_eqns d flags avoid env sigma ci computable constructs consnargsl bl = try @@ -1027,6 +1028,7 @@ let rec subst_glob_constr env subst = DAst.map (function | GVar _ | GEvar _ | GInt _ + | GFloat _ | GPatVar _ as raw -> raw | GApp (r,rl) as raw -> diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 288a349b8b..73d0c6f821 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -138,7 +138,7 @@ let flex_kind_of_term flags env evd c sk = | Evar ev -> if is_frozen flags ev then Rigid else Flexible ev - | Lambda _ | Prod _ | Sort _ | Ind _ | Construct _ | CoFix _ | Int _ -> Rigid + | Lambda _ | Prod _ | Sort _ | Ind _ | Construct _ | CoFix _ | Int _ | Float _ -> Rigid | Meta _ -> Rigid | Fix _ -> Rigid (* happens when the fixpoint is partially applied *) | Cast _ | App _ | Case _ -> assert false @@ -220,7 +220,7 @@ let occur_rigidly flags env evd (evk,_) t = (match aux c with | Rigid b -> Rigid b | _ -> Reducible) - | Meta _ | Fix _ | CoFix _ | Int _ -> Reducible + | Meta _ | Fix _ | CoFix _ | Int _ | Float _ -> Reducible in match aux t with | Rigid b -> b @@ -899,7 +899,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty only if necessary) or the second argument is potentially usable as a canonical projection or canonical value *) let rec is_unnamed (hd, args) = match EConstr.kind i hd with - | (Var _|Construct _|Ind _|Const _|Prod _|Sort _|Int _) -> + | (Var _|Construct _|Ind _|Const _|Prod _|Sort _|Int _ |Float _) -> Stack.not_purely_applicative args | (CoFix _|Meta _|Rel _)-> true | Evar _ -> Stack.not_purely_applicative args @@ -1019,7 +1019,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty | Const _, Const _ | Ind _, Ind _ | Construct _, Construct _ - | Int _, Int _ -> + | Int _, Int _ + | Float _, Float _ -> rigids env evd sk1 term1 sk2 term2 | Evar (sp1,al1), Evar (sp2,al2) -> (* Frozen evars *) @@ -1064,7 +1065,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) flags env evd pbty |Some (sk1',sk2'), Success i' -> evar_conv_x flags env i' CONV (Stack.zip i' (term1,sk1')) (Stack.zip i' (term2,sk2')) end - | (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _ | Int _ | Evar _ | Lambda _), _ -> + | (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _ | Int _ | Float _ | Evar _ | Lambda _), _ -> UnifFailure (evd,NotSameHead) | _, (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _ | Int _ | Evar _ | Lambda _) -> UnifFailure (evd,NotSameHead) diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index 93f5923474..03bb633fa0 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -156,9 +156,10 @@ let mk_glob_constr_eq f c1 c2 = match DAst.get c1, DAst.get c2 with | GCast (c1, t1), GCast (c2, t2) -> f c1 c2 && cast_type_eq f t1 t2 | GInt i1, GInt i2 -> Uint63.equal i1 i2 + | GFloat f1, GFloat f2 -> Float64.equal f1 f2 | (GRef _ | GVar _ | GEvar _ | GPatVar _ | GApp _ | GLambda _ | GProd _ | GLetIn _ | GCases _ | GLetTuple _ | GIf _ | GRec _ | GSort _ | GHole _ | GCast _ | - GInt _), _ -> false + GInt _ | GFloat _), _ -> false let rec glob_constr_eq c = mk_glob_constr_eq glob_constr_eq c @@ -219,7 +220,7 @@ let map_glob_constr_left_to_right f = DAst.map (function let comp1 = f c in let comp2 = map_cast_type f k in GCast (comp1,comp2) - | (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GInt _) as x -> x + | (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GInt _ | GFloat _) as x -> x ) let map_glob_constr = map_glob_constr_left_to_right @@ -251,7 +252,7 @@ let fold_glob_constr f acc = DAst.with_val (function let acc = match k with | CastConv t | CastVM t | CastNative t -> f acc t | CastCoerce -> acc in f acc c - | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GInt _) -> acc + | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GInt _ | GFloat _) -> acc ) let fold_return_type_with_binders f g v acc (na,tyopt) = Option.fold_left (f (Name.fold_right g na v)) acc tyopt @@ -293,7 +294,7 @@ let fold_glob_constr_with_binders g f v acc = DAst.(with_val (function let acc = match k with | CastConv t | CastVM t | CastNative t -> f v acc t | CastCoerce -> acc in f v acc c - | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GInt _) -> acc)) + | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GInt _ | GFloat _) -> acc)) let iter_glob_constr f = fold_glob_constr (fun () -> f) () diff --git a/pretyping/glob_term.ml b/pretyping/glob_term.ml index 10e9d60fd5..44323441b6 100644 --- a/pretyping/glob_term.ml +++ b/pretyping/glob_term.ml @@ -91,6 +91,7 @@ type 'a glob_constr_r = | GHole of Evar_kinds.t * Namegen.intro_pattern_naming_expr * Genarg.glob_generic_argument option | GCast of 'a glob_constr_g * 'a glob_constr_g cast_type | GInt of Uint63.t + | GFloat of Float64.t and 'a glob_constr_g = ('a glob_constr_r, 'a) DAst.t and 'a glob_decl_g = Name.t * binding_kind * 'a glob_constr_g option * 'a glob_constr_g diff --git a/pretyping/heads.ml b/pretyping/heads.ml index 870df62500..7740628c21 100644 --- a/pretyping/heads.ml +++ b/pretyping/heads.ml @@ -79,7 +79,7 @@ and kind_of_head env t = | Proj (p,c) -> RigidHead RigidOther | Case (_,_,c,_) -> aux k [] c true - | Int _ -> ConstructorHead + | Int _ | Float _ -> ConstructorHead | Fix ((i,j),_) -> let n = i.(j) in try aux k [] (List.nth l n) true diff --git a/pretyping/keys.ml b/pretyping/keys.ml index f8eecd80d4..39a4a525ef 100644 --- a/pretyping/keys.ml +++ b/pretyping/keys.ml @@ -26,6 +26,7 @@ type key = | KCoFix | KRel | KInt + | KFloat module KeyOrdered = struct type t = key @@ -42,6 +43,7 @@ module KeyOrdered = struct | KCoFix -> 6 | KRel -> 7 | KInt -> 8 + | KFloat -> 9 let compare gr1 gr2 = match gr1, gr2 with @@ -135,6 +137,7 @@ let constr_key kind c = | Sort _ -> KSort | LetIn _ -> KLet | Int _ -> KInt + | Float _ -> KFloat in Some (aux c) with Not_found -> None @@ -151,6 +154,7 @@ let pr_key pr_global = function | KCoFix -> str"CoFix" | KRel -> str"Rel" | KInt -> str"Int" + | KFloat -> str"Float" let pr_keyset pr_global v = prlist_with_sep spc (pr_key pr_global) (Keyset.elements v) diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index e5aed300a2..0178d5c009 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -208,6 +208,7 @@ let rec nf_val env sigma v typ = mkLambda(name,dom,body) | Vconst n -> construct_of_constr_const env sigma n typ | Vint64 i -> i |> Uint63.of_int64 |> mkInt + | Vfloat64 f -> f |> Float64.of_float |> mkFloat | Vblock b -> let capp,ctyp = construct_of_constr_block env sigma (block_tag b) typ in let args = nf_bargs env sigma b ctyp in diff --git a/pretyping/pattern.ml b/pretyping/pattern.ml index e0beb383b5..2d7a152817 100644 --- a/pretyping/pattern.ml +++ b/pretyping/pattern.ml @@ -40,6 +40,7 @@ type constr_pattern = | PFix of (int array * int) * (Name.t array * constr_pattern array * constr_pattern array) | PCoFix of int * (Name.t array * constr_pattern array * constr_pattern array) | PInt of Uint63.t + | PFloat of Float64.t (** Nota : in a [PCase], the array of branches might be shorter than expected, denoting the use of a final "_ => _" branch *) diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index ccc3b6e83c..0c4312dc77 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -62,9 +62,12 @@ let rec constr_pattern_eq p1 p2 = match p1, p2 with Projection.equal p1 p2 && constr_pattern_eq t1 t2 | PInt i1, PInt i2 -> Uint63.equal i1 i2 +| PFloat f1, PFloat f2 -> + Float64.equal f1 f2 | (PRef _ | PVar _ | PEvar _ | PRel _ | PApp _ | PSoApp _ | PLambda _ | PProd _ | PLetIn _ | PSort _ | PMeta _ - | PIf _ | PCase _ | PFix _ | PCoFix _ | PProj _ | PInt _), _ -> false + | PIf _ | PCase _ | PFix _ | PCoFix _ | PProj _ | PInt _ + | PFloat _), _ -> false (** FIXME: fixpoint and cofixpoint should be relativized to pattern *) and pattern_eq (i1, j1, p1) (i2, j2, p2) = @@ -92,7 +95,7 @@ let rec occur_meta_pattern = function (List.exists (fun (_,_,p) -> occur_meta_pattern p) br) | PMeta _ | PSoApp _ -> true | PEvar _ | PVar _ | PRef _ | PRel _ | PSort _ | PFix _ | PCoFix _ - | PInt _ -> false + | PInt _ | PFloat _ -> false let rec occurn_pattern n = function | PRel p -> Int.equal n p @@ -113,7 +116,7 @@ let rec occurn_pattern n = function (List.exists (fun (_,_,p) -> occurn_pattern n p) br) | PMeta _ | PSoApp _ -> true | PEvar (_,args) -> Array.exists (occurn_pattern n) args - | PVar _ | PRef _ | PSort _ | PInt _ -> false + | PVar _ | PRef _ | PSort _ | PInt _ | PFloat _ -> false | PFix (_,(_,tl,bl)) -> Array.exists (occurn_pattern n) tl || Array.exists (occurn_pattern (n+Array.length tl)) bl | PCoFix (_,(_,tl,bl)) -> @@ -136,7 +139,7 @@ let rec head_pattern_bound t = -> raise BoundPattern (* Perhaps they were arguments, but we don't beta-reduce *) | PLambda _ -> raise BoundPattern - | PCoFix _ | PInt _ -> anomaly ~label:"head_pattern_bound" (Pp.str "not a type.") + | PCoFix _ | PInt _ | PFloat _ -> anomaly ~label:"head_pattern_bound" (Pp.str "not a type.") let head_of_constr_reference sigma c = match EConstr.kind sigma c with | Const (sp,_) -> GlobRef.ConstRef sp @@ -213,7 +216,8 @@ let pattern_of_constr env sigma t = let env' = Array.fold_left2 push env lna tl in PCoFix (ln,(Array.map binder_name lna,Array.map (pattern_of_constr env) tl, Array.map (pattern_of_constr env') bl)) - | Int i -> PInt i in + | Int i -> PInt i + | Float f -> PFloat f in pattern_of_constr env t (* To process patterns, we need a translation without typing at all. *) @@ -235,7 +239,8 @@ let map_pattern_with_binders g f l = function let l' = Array.fold_left (fun l na -> g na l) l lna in PCoFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl)) (* Non recursive *) - | (PVar _ | PEvar _ | PRel _ | PRef _ | PSort _ | PMeta _ | PInt _ as x) -> x + | (PVar _ | PEvar _ | PRel _ | PRef _ | PSort _ | PMeta _ | PInt _ + | PFloat _ as x) -> x let error_instantiate_pattern id l = let is = match l with @@ -290,7 +295,8 @@ let rec subst_pattern env sigma subst pat = | PVar _ | PEvar _ | PRel _ - | PInt _ -> pat + | PInt _ + | PFloat _ -> pat | PProj (p,c) -> let p' = Projection.map (subst_mind subst) p in let c' = subst_pattern env sigma subst c in @@ -495,6 +501,7 @@ let rec pat_of_raw metas vars = DAst.with_loc_val (fun ?loc -> function PCoFix (n, (names, tl, cl)) | GInt i -> PInt i + | GFloat f -> PFloat f | GPatVar _ | GIf _ | GLetTuple _ | GCases _ | GEvar _ -> err ?loc (Pp.str "Non supported pattern.")) diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 4fed526cfc..2e1cb9ff08 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -1026,6 +1026,13 @@ let rec pretype ~program_mode ~poly resolve_tc (tycon : type_constraint) (env : user_err ?loc ~hdr:"pretype" (str "Type of int63 should be registered first.") in inh_conv_coerce_to_tycon ?loc env sigma resj tycon + | GFloat f -> + let resj = + try Typing.judge_of_float !!env f + with Invalid_argument _ -> + user_err ?loc ~hdr:"pretype" (str "Type of float should be registered first.") + in + inh_conv_coerce_to_tycon ?loc env sigma resj tycon and pretype_instance ~program_mode ~poly resolve_tc env sigma loc hyps evk update = let f decl (subst,update,sigma) = diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index df161b747a..2952466fbb 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -848,9 +848,17 @@ struct | Int i -> i | _ -> raise Primred.NativeDestKO + let get_float evd e = + match EConstr.kind evd e with + | Float f -> f + | _ -> raise Primred.NativeDestKO + let mkInt env i = mkInt i + let mkFloat env f = + mkFloat f + let mkBool env b = let (ct,cf) = get_bool_constructors env in mkConstruct (if b then ct else cf) @@ -865,6 +873,12 @@ struct let c = get_pair_constructor env in mkApp(mkConstruct c, [|int_ty;int_ty;e1;e2|]) + let mkFloatIntPair env f i = + let float_ty = mkConst @@ get_float_type env in + let int_ty = mkConst @@ get_int_type env in + let c = get_pair_constructor env in + mkApp(mkConstruct c, [|float_ty;int_ty;f;i|]) + let mkLt env = let (_eq, lt, _gt) = get_cmp_constructors env in mkConstruct lt @@ -877,6 +891,66 @@ struct let (_eq, _lt, gt) = get_cmp_constructors env in mkConstruct gt + let mkFLt env = + let (_eq, lt, _gt, _nc) = get_f_cmp_constructors env in + mkConstruct lt + + let mkFEq env = + let (eq, _lt, _gt, _nc) = get_f_cmp_constructors env in + mkConstruct eq + + let mkFGt env = + let (_eq, _lt, gt, _nc) = get_f_cmp_constructors env in + mkConstruct gt + + let mkFNotComparable env = + let (_eq, _lt, _gt, nc) = get_f_cmp_constructors env in + mkConstruct nc + + let mkPNormal env = + let (pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + mkConstruct pNormal + + let mkNNormal env = + let (_pNormal,nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + mkConstruct nNormal + + let mkPSubn env = + let (_pNormal,_nNormal,pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + mkConstruct pSubn + + let mkNSubn env = + let (_pNormal,_nNormal,_pSubn,nSubn,_pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + mkConstruct nSubn + + let mkPZero env = + let (_pNormal,_nNormal,_pSubn,_nSubn,pZero,_nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + mkConstruct pZero + + let mkNZero env = + let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,nZero,_pInf,_nInf,_nan) = + get_f_class_constructors env in + mkConstruct nZero + + let mkPInf env = + let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,pInf,_nInf,_nan) = + get_f_class_constructors env in + mkConstruct pInf + + let mkNInf env = + let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,nInf,_nan) = + get_f_class_constructors env in + mkConstruct nInf + + let mkNaN env = + let (_pNormal,_nNormal,_pSubn,_nSubn,_pZero,_nZero,_pInf,_nInf,nan) = + get_f_class_constructors env in + mkConstruct nan end module CredNative = RedNative(CNativeEntries) @@ -1135,7 +1209,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = |_ -> fold () else fold () - | Int i -> + | Int _ | Float _ -> begin match Stack.strip_app stack with | (_, Stack.Primitive(p,kn,rargs,kargs,cst_l')::s) -> let more_to_reduce = List.exists (fun k -> CPrimitives.Kwhnf = k) kargs in @@ -1238,7 +1312,7 @@ let local_whd_state_gen flags sigma = else s | Rel _ | Var _ | Sort _ | Prod _ | LetIn _ | Const _ | Ind _ | Proj _ - | Int _ -> s + | Int _ | Float _ -> s in whrec diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index cc341afac3..966c8f6e12 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -146,6 +146,7 @@ let retype ?(polyprop=true) sigma = | Cast (c,_, t) -> t | Sort _ | Prod _ -> mkSort (sort_of env cstr) | Int _ -> EConstr.of_constr (Typeops.type_of_int env) + | Float _ -> EConstr.of_constr (Typeops.type_of_float env) and sort_of env t = match EConstr.kind sigma t with @@ -281,7 +282,7 @@ let relevance_of_term env sigma c = | Fix ((_,i),(lna,_,_)) -> (lna.(i)).binder_relevance | CoFix (i,(lna,_,_)) -> (lna.(i)).binder_relevance | Proj (p, _) -> Retypeops.relevance_of_projection env p - | Int _ -> Sorts.Relevant + | Int _ | Float _ -> Sorts.Relevant | Meta _ | Evar _ -> Sorts.Relevant diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 2db5512ff4..1a145fe1b2 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -319,6 +319,9 @@ let type_of_constructor env sigma ((ind,_ as ctor),u) = let judge_of_int env v = Environ.on_judgment EConstr.of_constr (judge_of_int env v) +let judge_of_float env v = + Environ.on_judgment EConstr.of_constr (judge_of_float env v) + (* cstr must be in n.f. w.r.t. evars and execute returns a judgement where both the term and type are in n.f. *) let rec execute env sigma cstr = @@ -430,6 +433,9 @@ let rec execute env sigma cstr = | Int i -> sigma, judge_of_int env i + | Float f -> + sigma, judge_of_float env f + and execute_recdef env sigma (names,lar,vdef) = let sigma, larj = execute_array env sigma lar in let sigma, lara = Array.fold_left_map (assumption_of_judgment env) sigma larj in diff --git a/pretyping/typing.mli b/pretyping/typing.mli index 63fb0679f1..1b07b2bb78 100644 --- a/pretyping/typing.mli +++ b/pretyping/typing.mli @@ -57,3 +57,4 @@ val judge_of_product : Environ.env -> Name.t -> unsafe_type_judgment -> unsafe_type_judgment -> unsafe_judgment val judge_of_projection : env -> evar_map -> Projection.t -> unsafe_judgment -> unsafe_judgment val judge_of_int : Environ.env -> Uint63.t -> unsafe_judgment +val judge_of_float : Environ.env -> Float64.t -> unsafe_judgment diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 4d34139ec0..7147580b3d 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -566,7 +566,7 @@ let is_rigid_head sigma flags t = match EConstr.kind sigma t with | Const (cst,u) -> not (TransparentState.is_transparent_constant flags.modulo_delta cst) | Ind (i,u) -> true - | Construct _ | Int _ -> true + | Construct _ | Int _ | Float _ -> true | Fix _ | CoFix _ -> true | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Cast (_, _, _) | Prod _ | Lambda _ | LetIn _ | App (_, _) | Case (_, _, _, _) @@ -661,7 +661,7 @@ let rec is_neutral env sigma ts t = | Evar _ | Meta _ -> true | Case (_, p, c, cl) -> is_neutral env sigma ts c | Proj (p, c) -> is_neutral env sigma ts c - | Lambda _ | LetIn _ | Construct _ | CoFix _ | Int _ -> false + | Lambda _ | LetIn _ | Construct _ | CoFix _ | Int _ | Float _ -> false | Sort _ | Cast (_, _, _) | Prod (_, _, _) | Ind _ -> false (* Really? *) | Fix _ -> false (* This is an approximation *) | App _ -> assert false @@ -1821,7 +1821,7 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) = | Cast (_, _, _) (* Is this expected? *) | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _ -> user_err Pp.(str "Match_subterm"))) + | Construct _ | Int _ | Float _ -> user_err Pp.(str "Match_subterm"))) in try matchrec cl with ex when precatchable_exception ex -> @@ -1890,7 +1890,7 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) = | Cast (_, _, _) -> fail "Match_subterm" (* Is this expected? *) | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Const _ | Ind _ - | Construct _ | Int _ -> fail "Match_subterm")) + | Construct _ | Int _ | Float _ -> fail "Match_subterm")) in let res = matchrec cl [] in diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index da0a92f284..d15eb578c3 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -169,6 +169,7 @@ and nf_whd env sigma whd typ = let args = nf_bargs env sigma b ofs ctyp in mkApp(capp,args) | Vint64 i -> i |> Uint63.of_int64 |> mkInt + | Vfloat64 f -> f |> Float64.of_float |> mkFloat | Vatom_stk(Aid idkey, stk) -> constr_type_of_idkey env sigma idkey stk | Vatom_stk(Aind ((mi,i) as ind), stk) -> diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml index ccd7a818b9..58db147b10 100644 --- a/tactics/term_dnet.ml +++ b/tactics/term_dnet.ml @@ -45,6 +45,7 @@ struct | DFix of int array * int * 't array * 't array | DCoFix of int * 't array * 't array | DInt of Uint63.t + | DFloat of Float64.t (* special constructors only inside the left-hand side of DCtx or DApp. Used to encode lists of foralls/letins/apps as contexts *) @@ -63,6 +64,7 @@ struct | DFix _ -> str "fix" | DCoFix _ -> str "cofix" | DInt _ -> str "INT" + | DFloat _ -> str "FLOAT" | DCons ((t,dopt),tl) -> f t ++ (match dopt with Some t' -> str ":=" ++ f t' | None -> str "") ++ spc() ++ str "::" ++ spc() ++ f tl @@ -74,7 +76,7 @@ struct *) let map f = function - | (DRel | DSort | DNil | DRef _ | DInt _) as c -> c + | (DRel | DSort | DNil | DRef _ | DInt _ | DFloat _) as c -> c | DCtx (ctx,c) -> DCtx (f ctx, f c) | DLambda (t,c) -> DLambda (f t, f c) | DApp (t,u) -> DApp (f t,f u) @@ -151,6 +153,10 @@ struct | DInt _, _ -> -1 | _, DInt _ -> 1 + | DFloat f1, DFloat f2 -> Float64.total_compare f1 f2 + + | DFloat _, _ -> -1 | _, DFloat _ -> 1 + | DCons ((t1, ot1), u1), DCons ((t2, ot2), u2) -> let c = cmp t1 t2 in if Int.equal c 0 then @@ -163,7 +169,7 @@ struct | DNil, DNil -> 0 let fold f acc = function - | (DRel | DNil | DSort | DRef _ | DInt _) -> acc + | (DRel | DNil | DSort | DRef _ | DInt _ | DFloat _) -> acc | DCtx (ctx,c) -> f (f acc ctx) c | DLambda (t,c) -> f (f acc t) c | DApp (t,u) -> f (f acc t) u @@ -175,7 +181,7 @@ struct | DCons ((t,topt),u) -> f (Option.fold_left f (f acc t) topt) u let choose f = function - | (DRel | DSort | DNil | DRef _ | DInt _) -> invalid_arg "choose" + | (DRel | DSort | DNil | DRef _ | DInt _ | DFloat _) -> invalid_arg "choose" | DCtx (ctx,c) -> f ctx | DLambda (t,c) -> f t | DApp (t,u) -> f u @@ -192,7 +198,7 @@ struct then invalid_arg "fold2:compare" else match c1,c2 with | (DRel, DRel | DNil, DNil | DSort, DSort | DRef _, DRef _ - | DInt _, DInt _) -> acc + | DInt _, DInt _ | DFloat _, DFloat _) -> acc | (DCtx (c1,t1), DCtx (c2,t2) | DApp (c1,t1), DApp (c2,t2) | DLambda (c1,t1), DLambda (c2,t2)) -> f (f acc c1 c2) t1 t2 @@ -205,7 +211,7 @@ struct | DCons ((t1,topt1),u1), DCons ((t2,topt2),u2) -> f (Option.fold_left2 f (f acc t1 t2) topt1 topt2) u1 u2 | (DRel | DNil | DSort | DRef _ | DCtx _ | DApp _ | DLambda _ | DCase _ - | DFix _ | DCoFix _ | DCons _ | DInt _), _ -> assert false + | DFix _ | DCoFix _ | DCons _ | DInt _ | DFloat _), _ -> assert false let map2 (f:'a -> 'b -> 'c) (c1:'a t) (c2:'b t) : 'c t = let head w = map (fun _ -> ()) w in @@ -213,7 +219,7 @@ struct then invalid_arg "map2_t:compare" else match c1,c2 with | (DRel, DRel | DSort, DSort | DNil, DNil | DRef _, DRef _ - | DInt _, DInt _) as cc -> + | DInt _, DInt _ | DFloat _, DFloat _) as cc -> let (c,_) = cc in c | DCtx (c1,t1), DCtx (c2,t2) -> DCtx (f c1 c2, f t1 t2) | DLambda (t1,c1), DLambda (t2,c2) -> DLambda (f t1 t2, f c1 c2) @@ -227,10 +233,10 @@ struct | DCons ((t1,topt1),u1), DCons ((t2,topt2),u2) -> DCons ((f t1 t2,Option.lift2 f topt1 topt2), f u1 u2) | (DRel | DNil | DSort | DRef _ | DCtx _ | DApp _ | DLambda _ | DCase _ - | DFix _ | DCoFix _ | DCons _ | DInt _), _ -> assert false + | DFix _ | DCoFix _ | DCons _ | DInt _ | DFloat _), _ -> assert false let terminal = function - | (DRel | DSort | DNil | DRef _ | DInt _) -> true + | (DRel | DSort | DNil | DRef _ | DInt _ | DFloat _) -> true | DLambda _ | DApp _ | DCase _ | DFix _ | DCoFix _ | DCtx _ | DCons _ -> false @@ -325,6 +331,7 @@ struct | Proj (p,c) -> Term (DApp (Term (DRef (ConstRef (Projection.constant p))), pat_of_constr c)) | Int i -> Term (DInt i) + | Float f -> Term (DFloat f) and ctx_of_constr ctx c = match Constr.kind c with | Prod (_,t,c) -> ctx_of_constr (Term(DCons((pat_of_constr t,None),ctx))) c diff --git a/test-suite/Makefile b/test-suite/Makefile index d9c89d7a8a..1744138d29 100644 --- a/test-suite/Makefile +++ b/test-suite/Makefile @@ -102,7 +102,7 @@ INTERACTIVE := interactive UNIT_TESTS := unit-tests VSUBSYSTEMS := prerequisite success failure $(BUGS) output output-coqtop \ output-modulo-time $(INTERACTIVE) micromega $(COMPLEXITY) modules stm \ - coqdoc ssr arithmetic ltac2 + coqdoc ssr primitive/uint63 primitive/float ltac2 # All subsystems SUBSYSTEMS := $(VSUBSYSTEMS) misc bugs ide vio coqchk coqwc coq-makefile tools $(UNIT_TESTS) @@ -175,6 +175,7 @@ summary: $(call summary_dir, "Miscellaneous tests", misc); \ $(call summary_dir, "Complexity tests", complexity); \ $(call summary_dir, "Module tests", modules); \ + $(call summary_dir, "Primitive tests", primitive); \ $(call summary_dir, "STM tests", stm); \ $(call summary_dir, "SSR tests", ssr); \ $(call summary_dir, "IDE tests", ide); \ @@ -330,7 +331,7 @@ $(addsuffix .log,$(wildcard prerequisite/*.v)): %.v.log: %.v } > "$@" ssr: $(wildcard ssr/*.v:%.v=%.v.log) -$(addsuffix .log,$(wildcard ssr/*.v success/*.v micromega/*.v modules/*.v arithmetic/*.v ltac2/*.v)): %.v.log: %.v $(PREREQUISITELOG) +$(addsuffix .log,$(wildcard ssr/*.v success/*.v micromega/*.v modules/*.v primitive/*/*.v ltac2/*.v)): %.v.log: %.v $(PREREQUISITELOG) @echo "TEST $< $(call get_coq_prog_args_in_parens,"$<")" $(HIDE){ \ opts="$(if $(findstring modules/,$<),-R modules Mods)"; \ diff --git a/test-suite/output/FloatExtraction.out b/test-suite/output/FloatExtraction.out new file mode 100644 index 0000000000..cfd6633752 --- /dev/null +++ b/test-suite/output/FloatExtraction.out @@ -0,0 +1,67 @@ + +(** val infinity : Float64.t **) + +let infinity = + (Float64.of_float (infinity)) + +(** val neg_infinity : Float64.t **) + +let neg_infinity = + (Float64.of_float (neg_infinity)) + +(** val nan : Float64.t **) + +let nan = + (Float64.of_float (nan)) + +(** val one : Float64.t **) + +let one = + (Float64.of_float (0x1p+0)) + +(** val zero : Float64.t **) + +let zero = + (Float64.of_float (0x0p+0)) + +(** val two : Float64.t **) + +let two = + (Float64.of_float (0x1p+1)) + +(** val list_floats : Float64.t list **) + +let list_floats = + nan :: (infinity :: (neg_infinity :: (zero :: (one :: (two :: ((Float64.of_float (0x1p-1)) :: ((Float64.of_float (0x1.47ae147ae147bp-7)) :: ((Float64.of_float (-0x1p-1)) :: ((Float64.of_float (-0x1.47ae147ae147bp-7)) :: ((Float64.of_float (0x1.e42d130773b76p+1023)) :: ((Float64.of_float (-0x0.c396c98f8d899p-1022)) :: []))))))))))) + + +(** val sqrt : Float64.t -> Float64.t **) + +let sqrt = Float64.sqrt + +(** val opp : Float64.t -> Float64.t **) + +let opp = Float64.opp + +(** val mul : Float64.t -> Float64.t -> Float64.t **) + +let mul = Float64.mul + +(** val sub : Float64.t -> Float64.t -> Float64.t **) + +let sub = Float64.sub + +(** val div : Float64.t -> Float64.t -> Float64.t **) + +let div = Float64.div + +(** val discr : Float64.t -> Float64.t -> Float64.t -> Float64.t **) + +let discr a b c = + sub (mul b b) (mul (mul (Float64.of_float (0x1p+2)) a) c) + +(** val x1 : Float64.t -> Float64.t -> Float64.t -> Float64.t **) + +let x1 a b c = + div (sub (opp b) (sqrt (discr a b c))) (mul (Float64.of_float (0x1p+1)) a) + diff --git a/test-suite/output/FloatExtraction.v b/test-suite/output/FloatExtraction.v new file mode 100644 index 0000000000..f296e8e871 --- /dev/null +++ b/test-suite/output/FloatExtraction.v @@ -0,0 +1,33 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(* * (see LICENSE file for the text of the license) *) +(************************************************************************) + +Require Import Floats ExtrOCamlFloats. + +Require Import List. Import ListNotations. + +(* from Require Import ExtrOcamlBasic. *) +Extract Inductive list => list [ "[]" "( :: )" ]. + +Local Open Scope float_scope. + +(* Avoid exponents with less than three digits as they are usually + displayed with two digits (1e7 is displayed 1e+07) except on + Windows where three digits are used (1e+007). *) +Definition list_floats := + [nan; infinity; neg_infinity; zero; one; two; + 0.5; 0.01; -0.5; -0.01; 1.7e+308; -1.7e-308]. + +Recursive Extraction list_floats. + +Definition discr a b c := b * b - 4.0 * a * c. + +Definition x1 a b c := (- b - sqrt (discr a b c)) / (2.0 * a). + +Recursive Extraction x1. diff --git a/test-suite/output/FloatSyntax.out b/test-suite/output/FloatSyntax.out new file mode 100644 index 0000000000..668a55977d --- /dev/null +++ b/test-suite/output/FloatSyntax.out @@ -0,0 +1,40 @@ +2%float + : float +2.5%float + : float +(-2.5)%float + : float +2.4999999999999999e+123%float + : float +(-2.5000000000000001e-123)%float + : float +(2 + 2)%float + : float +(2.5 + 2.5)%float + : float +2 + : float +2.5 + : float +-2.5 + : float +2.4999999999999999e+123 + : float +-2.5000000000000001e-123 + : float +2 + 2 + : float +2.5 + 2.5 + : float +2 + : nat +2%float + : float +t = 2%flt + : float +t = 2%flt + : float +2 + : nat +2 + : float diff --git a/test-suite/output/FloatSyntax.v b/test-suite/output/FloatSyntax.v new file mode 100644 index 0000000000..85f611352c --- /dev/null +++ b/test-suite/output/FloatSyntax.v @@ -0,0 +1,37 @@ +Require Import Floats. + +Check 2%float. +Check 2.5%float. +Check (-2.5)%float. +(* Avoid exponents with less than three digits as they are usually + displayed with two digits (1e7 is displayed 1e+07) except on + Windows where three digits are used (1e+007). *) +Check 2.5e123%float. +Check (-2.5e-123)%float. +Check (2 + 2)%float. +Check (2.5 + 2.5)%float. + +Open Scope float_scope. + +Check 2. +Check 2.5. +Check (-2.5). +Check 2.5e123. +Check (-2.5e-123). +Check (2 + 2). +Check (2.5 + 2.5). + +Open Scope nat_scope. + +Check 2. +Check 2%float. + +Delimit Scope float_scope with flt. +Definition t := 2%float. +Print t. +Delimit Scope nat_scope with float. +Print t. +Check 2. +Close Scope nat_scope. +Check 2. +Close Scope float_scope. diff --git a/test-suite/primitive/float/add.v b/test-suite/primitive/float/add.v new file mode 100644 index 0000000000..f8c5939d0a --- /dev/null +++ b/test-suite/primitive/float/add.v @@ -0,0 +1,63 @@ +Require Import ZArith Int63 Floats. + +Open Scope float_scope. + +Definition two := Eval compute in of_int63 2%int63. +Definition three := Eval compute in of_int63 3%int63. +Definition five := Eval compute in of_int63 5%int63. + +Check (eq_refl : two + three = five). +Check (eq_refl five <: two + three = five). +Check (eq_refl five <<: two + three = five). +Definition compute1 := Eval compute in two + three. +Check (eq_refl compute1 : five = five). + +Definition huge := Eval compute in ldexp one 1023%Z. +Definition tiny := Eval compute in ldexp one (-1023)%Z. + +Check (eq_refl : huge + tiny = huge). +Check (eq_refl huge <: huge + tiny = huge). +Check (eq_refl huge <<: huge + tiny = huge). +Definition compute2 := Eval compute in huge + tiny. +Check (eq_refl compute2 : huge = huge). + +Check (eq_refl : huge + huge = infinity). +Check (eq_refl infinity <: huge + huge = infinity). +Check (eq_refl infinity <<: huge + huge = infinity). +Definition compute3 := Eval compute in huge + huge. +Check (eq_refl compute3 : infinity = infinity). + +Check (eq_refl : one + nan = nan). +Check (eq_refl nan <: one + nan = nan). +Check (eq_refl nan <<: one + nan = nan). +Definition compute4 := Eval compute in one + nan. +Check (eq_refl compute4 : nan = nan). + +Check (eq_refl : infinity + infinity = infinity). +Check (eq_refl infinity <: infinity + infinity = infinity). +Check (eq_refl infinity <<: infinity + infinity = infinity). +Definition compute5 := Eval compute in infinity + infinity. +Check (eq_refl compute5 : infinity = infinity). + +Check (eq_refl : infinity + neg_infinity = nan). +Check (eq_refl nan <: infinity + neg_infinity = nan). +Check (eq_refl nan <<: infinity + neg_infinity = nan). +Definition compute6 := Eval compute in infinity + neg_infinity. +Check (eq_refl compute6 : nan = nan). + +Check (eq_refl : zero + zero = zero). +Check (eq_refl zero <: zero + zero = zero). +Check (eq_refl zero <<: zero + zero = zero). +Check (eq_refl : neg_zero + zero = zero). +Check (eq_refl zero <: neg_zero + zero = zero). +Check (eq_refl zero <<: neg_zero + zero = zero). +Check (eq_refl : neg_zero + neg_zero = neg_zero). +Check (eq_refl neg_zero <: neg_zero + neg_zero = neg_zero). +Check (eq_refl neg_zero <<: neg_zero + neg_zero = neg_zero). +Check (eq_refl : zero + neg_zero = zero). +Check (eq_refl zero <: zero + neg_zero = zero). +Check (eq_refl zero <<: zero + neg_zero = zero). + +Check (eq_refl : huge + neg_infinity = neg_infinity). +Check (eq_refl neg_infinity <: huge + neg_infinity = neg_infinity). +Check (eq_refl neg_infinity <<: huge + neg_infinity = neg_infinity). diff --git a/test-suite/primitive/float/classify.v b/test-suite/primitive/float/classify.v new file mode 100644 index 0000000000..22e3fca844 --- /dev/null +++ b/test-suite/primitive/float/classify.v @@ -0,0 +1,33 @@ +Require Import ZArith Floats. + +Definition epsilon := Eval compute in ldexp one (-1024)%Z. + +Check (eq_refl : classify one = PNormal). +Check (eq_refl : classify (- one)%float = NNormal). +Check (eq_refl : classify epsilon = PSubn). +Check (eq_refl : classify (- epsilon)%float = NSubn). +Check (eq_refl : classify zero = PZero). +Check (eq_refl : classify neg_zero = NZero). +Check (eq_refl : classify infinity = PInf). +Check (eq_refl : classify neg_infinity = NInf). +Check (eq_refl : classify nan = NaN). + +Check (eq_refl PNormal <: classify one = PNormal). +Check (eq_refl NNormal <: classify (- one)%float = NNormal). +Check (eq_refl PSubn <: classify epsilon = PSubn). +Check (eq_refl NSubn <: classify (- epsilon)%float = NSubn). +Check (eq_refl PZero <: classify zero = PZero). +Check (eq_refl NZero <: classify neg_zero = NZero). +Check (eq_refl PInf <: classify infinity = PInf). +Check (eq_refl NInf <: classify neg_infinity = NInf). +Check (eq_refl NaN <: classify nan = NaN). + +Check (eq_refl PNormal <<: classify one = PNormal). +Check (eq_refl NNormal <<: classify (- one)%float = NNormal). +Check (eq_refl PSubn <<: classify epsilon = PSubn). +Check (eq_refl NSubn <<: classify (- epsilon)%float = NSubn). +Check (eq_refl PZero <<: classify zero = PZero). +Check (eq_refl NZero <<: classify neg_zero = NZero). +Check (eq_refl PInf <<: classify infinity = PInf). +Check (eq_refl NInf <<: classify neg_infinity = NInf). +Check (eq_refl NaN <<: classify nan = NaN). diff --git a/test-suite/primitive/float/compare.v b/test-suite/primitive/float/compare.v new file mode 100644 index 0000000000..23d1e5bbae --- /dev/null +++ b/test-suite/primitive/float/compare.v @@ -0,0 +1,385 @@ +(* DO NOT EDIT THIS FILE: automatically generated by ./gen_compare.sh *) +Require Import ZArith Floats. +Local Open Scope float_scope. + +Definition min_denorm := Eval compute in ldexp one (-1074)%Z. + +Definition min_norm := Eval compute in ldexp one (-1024)%Z. + +Check (eq_refl false : nan == nan = false). +Check (eq_refl false : nan == nan = false). +Check (eq_refl false : nan < nan = false). +Check (eq_refl false : nan < nan = false). +Check (eq_refl false : nan <= nan = false). +Check (eq_refl false : nan <= nan = false). +Check (eq_refl FNotComparable : nan ?= nan = FNotComparable). +Check (eq_refl FNotComparable : nan ?= nan = FNotComparable). + +Check (eq_refl false <: nan == nan = false). +Check (eq_refl false <: nan == nan = false). +Check (eq_refl false <: nan < nan = false). +Check (eq_refl false <: nan < nan = false). +Check (eq_refl false <: nan <= nan = false). +Check (eq_refl false <: nan <= nan = false). +Check (eq_refl FNotComparable <: nan ?= nan = FNotComparable). +Check (eq_refl FNotComparable <: nan ?= nan = FNotComparable). + +Check (eq_refl false <<: nan == nan = false). +Check (eq_refl false <<: nan == nan = false). +Check (eq_refl false <<: nan < nan = false). +Check (eq_refl false <<: nan < nan = false). +Check (eq_refl false <<: nan <= nan = false). +Check (eq_refl false <<: nan <= nan = false). +Check (eq_refl FNotComparable <<: nan ?= nan = FNotComparable). +Check (eq_refl FNotComparable <<: nan ?= nan = FNotComparable). + +Check (eq_refl false : nan == - nan = false). +Check (eq_refl false : - nan == nan = false). +Check (eq_refl false : nan < - nan = false). +Check (eq_refl false : - nan < nan = false). +Check (eq_refl false : nan <= - nan = false). +Check (eq_refl false : - nan <= nan = false). +Check (eq_refl FNotComparable : nan ?= - nan = FNotComparable). +Check (eq_refl FNotComparable : - nan ?= nan = FNotComparable). + +Check (eq_refl false <: nan == - nan = false). +Check (eq_refl false <: - nan == nan = false). +Check (eq_refl false <: nan < - nan = false). +Check (eq_refl false <: - nan < nan = false). +Check (eq_refl false <: nan <= - nan = false). +Check (eq_refl false <: - nan <= nan = false). +Check (eq_refl FNotComparable <: nan ?= - nan = FNotComparable). +Check (eq_refl FNotComparable <: - nan ?= nan = FNotComparable). + +Check (eq_refl false <<: nan == - nan = false). +Check (eq_refl false <<: - nan == nan = false). +Check (eq_refl false <<: nan < - nan = false). +Check (eq_refl false <<: - nan < nan = false). +Check (eq_refl false <<: nan <= - nan = false). +Check (eq_refl false <<: - nan <= nan = false). +Check (eq_refl FNotComparable <<: nan ?= - nan = FNotComparable). +Check (eq_refl FNotComparable <<: - nan ?= nan = FNotComparable). + +Check (eq_refl true : one == one = true). +Check (eq_refl true : one == one = true). +Check (eq_refl false : one < one = false). +Check (eq_refl false : one < one = false). +Check (eq_refl true : one <= one = true). +Check (eq_refl true : one <= one = true). +Check (eq_refl FEq : one ?= one = FEq). +Check (eq_refl FEq : one ?= one = FEq). + +Check (eq_refl true <: one == one = true). +Check (eq_refl true <: one == one = true). +Check (eq_refl false <: one < one = false). +Check (eq_refl false <: one < one = false). +Check (eq_refl true <: one <= one = true). +Check (eq_refl true <: one <= one = true). +Check (eq_refl FEq <: one ?= one = FEq). +Check (eq_refl FEq <: one ?= one = FEq). + +Check (eq_refl true <<: one == one = true). +Check (eq_refl true <<: one == one = true). +Check (eq_refl false <<: one < one = false). +Check (eq_refl false <<: one < one = false). +Check (eq_refl true <<: one <= one = true). +Check (eq_refl true <<: one <= one = true). +Check (eq_refl FEq <<: one ?= one = FEq). +Check (eq_refl FEq <<: one ?= one = FEq). + +Check (eq_refl true : zero == zero = true). +Check (eq_refl true : zero == zero = true). +Check (eq_refl false : zero < zero = false). +Check (eq_refl false : zero < zero = false). +Check (eq_refl true : zero <= zero = true). +Check (eq_refl true : zero <= zero = true). +Check (eq_refl FEq : zero ?= zero = FEq). +Check (eq_refl FEq : zero ?= zero = FEq). + +Check (eq_refl true <: zero == zero = true). +Check (eq_refl true <: zero == zero = true). +Check (eq_refl false <: zero < zero = false). +Check (eq_refl false <: zero < zero = false). +Check (eq_refl true <: zero <= zero = true). +Check (eq_refl true <: zero <= zero = true). +Check (eq_refl FEq <: zero ?= zero = FEq). +Check (eq_refl FEq <: zero ?= zero = FEq). + +Check (eq_refl true <<: zero == zero = true). +Check (eq_refl true <<: zero == zero = true). +Check (eq_refl false <<: zero < zero = false). +Check (eq_refl false <<: zero < zero = false). +Check (eq_refl true <<: zero <= zero = true). +Check (eq_refl true <<: zero <= zero = true). +Check (eq_refl FEq <<: zero ?= zero = FEq). +Check (eq_refl FEq <<: zero ?= zero = FEq). + +Check (eq_refl true : zero == - zero = true). +Check (eq_refl true : - zero == zero = true). +Check (eq_refl false : zero < - zero = false). +Check (eq_refl false : - zero < zero = false). +Check (eq_refl true : zero <= - zero = true). +Check (eq_refl true : - zero <= zero = true). +Check (eq_refl FEq : zero ?= - zero = FEq). +Check (eq_refl FEq : - zero ?= zero = FEq). + +Check (eq_refl true <: zero == - zero = true). +Check (eq_refl true <: - zero == zero = true). +Check (eq_refl false <: zero < - zero = false). +Check (eq_refl false <: - zero < zero = false). +Check (eq_refl true <: zero <= - zero = true). +Check (eq_refl true <: - zero <= zero = true). +Check (eq_refl FEq <: zero ?= - zero = FEq). +Check (eq_refl FEq <: - zero ?= zero = FEq). + +Check (eq_refl true <<: zero == - zero = true). +Check (eq_refl true <<: - zero == zero = true). +Check (eq_refl false <<: zero < - zero = false). +Check (eq_refl false <<: - zero < zero = false). +Check (eq_refl true <<: zero <= - zero = true). +Check (eq_refl true <<: - zero <= zero = true). +Check (eq_refl FEq <<: zero ?= - zero = FEq). +Check (eq_refl FEq <<: - zero ?= zero = FEq). + +Check (eq_refl true : - zero == - zero = true). +Check (eq_refl true : - zero == - zero = true). +Check (eq_refl false : - zero < - zero = false). +Check (eq_refl false : - zero < - zero = false). +Check (eq_refl true : - zero <= - zero = true). +Check (eq_refl true : - zero <= - zero = true). +Check (eq_refl FEq : - zero ?= - zero = FEq). +Check (eq_refl FEq : - zero ?= - zero = FEq). + +Check (eq_refl true <: - zero == - zero = true). +Check (eq_refl true <: - zero == - zero = true). +Check (eq_refl false <: - zero < - zero = false). +Check (eq_refl false <: - zero < - zero = false). +Check (eq_refl true <: - zero <= - zero = true). +Check (eq_refl true <: - zero <= - zero = true). +Check (eq_refl FEq <: - zero ?= - zero = FEq). +Check (eq_refl FEq <: - zero ?= - zero = FEq). + +Check (eq_refl true <<: - zero == - zero = true). +Check (eq_refl true <<: - zero == - zero = true). +Check (eq_refl false <<: - zero < - zero = false). +Check (eq_refl false <<: - zero < - zero = false). +Check (eq_refl true <<: - zero <= - zero = true). +Check (eq_refl true <<: - zero <= - zero = true). +Check (eq_refl FEq <<: - zero ?= - zero = FEq). +Check (eq_refl FEq <<: - zero ?= - zero = FEq). + +Check (eq_refl true : infinity == infinity = true). +Check (eq_refl true : infinity == infinity = true). +Check (eq_refl false : infinity < infinity = false). +Check (eq_refl false : infinity < infinity = false). +Check (eq_refl true : infinity <= infinity = true). +Check (eq_refl true : infinity <= infinity = true). +Check (eq_refl FEq : infinity ?= infinity = FEq). +Check (eq_refl FEq : infinity ?= infinity = FEq). + +Check (eq_refl true <: infinity == infinity = true). +Check (eq_refl true <: infinity == infinity = true). +Check (eq_refl false <: infinity < infinity = false). +Check (eq_refl false <: infinity < infinity = false). +Check (eq_refl true <: infinity <= infinity = true). +Check (eq_refl true <: infinity <= infinity = true). +Check (eq_refl FEq <: infinity ?= infinity = FEq). +Check (eq_refl FEq <: infinity ?= infinity = FEq). + +Check (eq_refl true <<: infinity == infinity = true). +Check (eq_refl true <<: infinity == infinity = true). +Check (eq_refl false <<: infinity < infinity = false). +Check (eq_refl false <<: infinity < infinity = false). +Check (eq_refl true <<: infinity <= infinity = true). +Check (eq_refl true <<: infinity <= infinity = true). +Check (eq_refl FEq <<: infinity ?= infinity = FEq). +Check (eq_refl FEq <<: infinity ?= infinity = FEq). + +Check (eq_refl true : - infinity == - infinity = true). +Check (eq_refl true : - infinity == - infinity = true). +Check (eq_refl false : - infinity < - infinity = false). +Check (eq_refl false : - infinity < - infinity = false). +Check (eq_refl true : - infinity <= - infinity = true). +Check (eq_refl true : - infinity <= - infinity = true). +Check (eq_refl FEq : - infinity ?= - infinity = FEq). +Check (eq_refl FEq : - infinity ?= - infinity = FEq). + +Check (eq_refl true <: - infinity == - infinity = true). +Check (eq_refl true <: - infinity == - infinity = true). +Check (eq_refl false <: - infinity < - infinity = false). +Check (eq_refl false <: - infinity < - infinity = false). +Check (eq_refl true <: - infinity <= - infinity = true). +Check (eq_refl true <: - infinity <= - infinity = true). +Check (eq_refl FEq <: - infinity ?= - infinity = FEq). +Check (eq_refl FEq <: - infinity ?= - infinity = FEq). + +Check (eq_refl true <<: - infinity == - infinity = true). +Check (eq_refl true <<: - infinity == - infinity = true). +Check (eq_refl false <<: - infinity < - infinity = false). +Check (eq_refl false <<: - infinity < - infinity = false). +Check (eq_refl true <<: - infinity <= - infinity = true). +Check (eq_refl true <<: - infinity <= - infinity = true). +Check (eq_refl FEq <<: - infinity ?= - infinity = FEq). +Check (eq_refl FEq <<: - infinity ?= - infinity = FEq). + +Check (eq_refl false : min_denorm == min_norm = false). +Check (eq_refl false : min_norm == min_denorm = false). +Check (eq_refl true : min_denorm < min_norm = true). +Check (eq_refl false : min_norm < min_denorm = false). +Check (eq_refl true : min_denorm <= min_norm = true). +Check (eq_refl false : min_norm <= min_denorm = false). +Check (eq_refl FLt : min_denorm ?= min_norm = FLt). +Check (eq_refl FGt : min_norm ?= min_denorm = FGt). + +Check (eq_refl false <: min_denorm == min_norm = false). +Check (eq_refl false <: min_norm == min_denorm = false). +Check (eq_refl true <: min_denorm < min_norm = true). +Check (eq_refl false <: min_norm < min_denorm = false). +Check (eq_refl true <: min_denorm <= min_norm = true). +Check (eq_refl false <: min_norm <= min_denorm = false). +Check (eq_refl FLt <: min_denorm ?= min_norm = FLt). +Check (eq_refl FGt <: min_norm ?= min_denorm = FGt). + +Check (eq_refl false <<: min_denorm == min_norm = false). +Check (eq_refl false <<: min_norm == min_denorm = false). +Check (eq_refl true <<: min_denorm < min_norm = true). +Check (eq_refl false <<: min_norm < min_denorm = false). +Check (eq_refl true <<: min_denorm <= min_norm = true). +Check (eq_refl false <<: min_norm <= min_denorm = false). +Check (eq_refl FLt <<: min_denorm ?= min_norm = FLt). +Check (eq_refl FGt <<: min_norm ?= min_denorm = FGt). + +Check (eq_refl false : min_denorm == one = false). +Check (eq_refl false : one == min_denorm = false). +Check (eq_refl true : min_denorm < one = true). +Check (eq_refl false : one < min_denorm = false). +Check (eq_refl true : min_denorm <= one = true). +Check (eq_refl false : one <= min_denorm = false). +Check (eq_refl FLt : min_denorm ?= one = FLt). +Check (eq_refl FGt : one ?= min_denorm = FGt). + +Check (eq_refl false <: min_denorm == one = false). +Check (eq_refl false <: one == min_denorm = false). +Check (eq_refl true <: min_denorm < one = true). +Check (eq_refl false <: one < min_denorm = false). +Check (eq_refl true <: min_denorm <= one = true). +Check (eq_refl false <: one <= min_denorm = false). +Check (eq_refl FLt <: min_denorm ?= one = FLt). +Check (eq_refl FGt <: one ?= min_denorm = FGt). + +Check (eq_refl false <<: min_denorm == one = false). +Check (eq_refl false <<: one == min_denorm = false). +Check (eq_refl true <<: min_denorm < one = true). +Check (eq_refl false <<: one < min_denorm = false). +Check (eq_refl true <<: min_denorm <= one = true). +Check (eq_refl false <<: one <= min_denorm = false). +Check (eq_refl FLt <<: min_denorm ?= one = FLt). +Check (eq_refl FGt <<: one ?= min_denorm = FGt). + +Check (eq_refl false : min_norm == one = false). +Check (eq_refl false : one == min_norm = false). +Check (eq_refl true : min_norm < one = true). +Check (eq_refl false : one < min_norm = false). +Check (eq_refl true : min_norm <= one = true). +Check (eq_refl false : one <= min_norm = false). +Check (eq_refl FLt : min_norm ?= one = FLt). +Check (eq_refl FGt : one ?= min_norm = FGt). + +Check (eq_refl false <: min_norm == one = false). +Check (eq_refl false <: one == min_norm = false). +Check (eq_refl true <: min_norm < one = true). +Check (eq_refl false <: one < min_norm = false). +Check (eq_refl true <: min_norm <= one = true). +Check (eq_refl false <: one <= min_norm = false). +Check (eq_refl FLt <: min_norm ?= one = FLt). +Check (eq_refl FGt <: one ?= min_norm = FGt). + +Check (eq_refl false <<: min_norm == one = false). +Check (eq_refl false <<: one == min_norm = false). +Check (eq_refl true <<: min_norm < one = true). +Check (eq_refl false <<: one < min_norm = false). +Check (eq_refl true <<: min_norm <= one = true). +Check (eq_refl false <<: one <= min_norm = false). +Check (eq_refl FLt <<: min_norm ?= one = FLt). +Check (eq_refl FGt <<: one ?= min_norm = FGt). + +Check (eq_refl false : one == infinity = false). +Check (eq_refl false : infinity == one = false). +Check (eq_refl true : one < infinity = true). +Check (eq_refl false : infinity < one = false). +Check (eq_refl true : one <= infinity = true). +Check (eq_refl false : infinity <= one = false). +Check (eq_refl FLt : one ?= infinity = FLt). +Check (eq_refl FGt : infinity ?= one = FGt). + +Check (eq_refl false <: one == infinity = false). +Check (eq_refl false <: infinity == one = false). +Check (eq_refl true <: one < infinity = true). +Check (eq_refl false <: infinity < one = false). +Check (eq_refl true <: one <= infinity = true). +Check (eq_refl false <: infinity <= one = false). +Check (eq_refl FLt <: one ?= infinity = FLt). +Check (eq_refl FGt <: infinity ?= one = FGt). + +Check (eq_refl false <<: one == infinity = false). +Check (eq_refl false <<: infinity == one = false). +Check (eq_refl true <<: one < infinity = true). +Check (eq_refl false <<: infinity < one = false). +Check (eq_refl true <<: one <= infinity = true). +Check (eq_refl false <<: infinity <= one = false). +Check (eq_refl FLt <<: one ?= infinity = FLt). +Check (eq_refl FGt <<: infinity ?= one = FGt). + +Check (eq_refl false : - infinity == infinity = false). +Check (eq_refl false : infinity == - infinity = false). +Check (eq_refl true : - infinity < infinity = true). +Check (eq_refl false : infinity < - infinity = false). +Check (eq_refl true : - infinity <= infinity = true). +Check (eq_refl false : infinity <= - infinity = false). +Check (eq_refl FLt : - infinity ?= infinity = FLt). +Check (eq_refl FGt : infinity ?= - infinity = FGt). + +Check (eq_refl false <: - infinity == infinity = false). +Check (eq_refl false <: infinity == - infinity = false). +Check (eq_refl true <: - infinity < infinity = true). +Check (eq_refl false <: infinity < - infinity = false). +Check (eq_refl true <: - infinity <= infinity = true). +Check (eq_refl false <: infinity <= - infinity = false). +Check (eq_refl FLt <: - infinity ?= infinity = FLt). +Check (eq_refl FGt <: infinity ?= - infinity = FGt). + +Check (eq_refl false <<: - infinity == infinity = false). +Check (eq_refl false <<: infinity == - infinity = false). +Check (eq_refl true <<: - infinity < infinity = true). +Check (eq_refl false <<: infinity < - infinity = false). +Check (eq_refl true <<: - infinity <= infinity = true). +Check (eq_refl false <<: infinity <= - infinity = false). +Check (eq_refl FLt <<: - infinity ?= infinity = FLt). +Check (eq_refl FGt <<: infinity ?= - infinity = FGt). + +Check (eq_refl false : - infinity == one = false). +Check (eq_refl false : one == - infinity = false). +Check (eq_refl true : - infinity < one = true). +Check (eq_refl false : one < - infinity = false). +Check (eq_refl true : - infinity <= one = true). +Check (eq_refl false : one <= - infinity = false). +Check (eq_refl FLt : - infinity ?= one = FLt). +Check (eq_refl FGt : one ?= - infinity = FGt). + +Check (eq_refl false <: - infinity == one = false). +Check (eq_refl false <: one == - infinity = false). +Check (eq_refl true <: - infinity < one = true). +Check (eq_refl false <: one < - infinity = false). +Check (eq_refl true <: - infinity <= one = true). +Check (eq_refl false <: one <= - infinity = false). +Check (eq_refl FLt <: - infinity ?= one = FLt). +Check (eq_refl FGt <: one ?= - infinity = FGt). + +Check (eq_refl false <<: - infinity == one = false). +Check (eq_refl false <<: one == - infinity = false). +Check (eq_refl true <<: - infinity < one = true). +Check (eq_refl false <<: one < - infinity = false). +Check (eq_refl true <<: - infinity <= one = true). +Check (eq_refl false <<: one <= - infinity = false). +Check (eq_refl FLt <<: - infinity ?= one = FLt). +Check (eq_refl FGt <<: one ?= - infinity = FGt). diff --git a/test-suite/primitive/float/coq_env_double_array.v b/test-suite/primitive/float/coq_env_double_array.v new file mode 100644 index 0000000000..754ccb69aa --- /dev/null +++ b/test-suite/primitive/float/coq_env_double_array.v @@ -0,0 +1,13 @@ +Require Import Floats. + +Goal True. +pose (f := one). +pose (f' := (-f)%float). + +(* this used to segfault when the coq_env array given to the + coq_interprete C function was an unboxed OCaml Double_array + (created by Array.map in csymtable.ml just before calling + eval_tcode) *) +vm_compute in f'. + +Abort. diff --git a/test-suite/primitive/float/div.v b/test-suite/primitive/float/div.v new file mode 100644 index 0000000000..8e971f575b --- /dev/null +++ b/test-suite/primitive/float/div.v @@ -0,0 +1,87 @@ +Require Import ZArith Int63 Floats. + +Open Scope float_scope. + +Definition two := Eval compute in of_int63 2%int63. +Definition three := Eval compute in of_int63 3%int63. +Definition six := Eval compute in of_int63 6%int63. + +Check (eq_refl : six / three = two). +Check (eq_refl two <: six / three = two). +Check (eq_refl two <<: six / three = two). +Definition compute1 := Eval compute in six / three. +Check (eq_refl compute1 : two = two). + +Definition huge := Eval compute in ldexp one 1023%Z. +Definition tiny := Eval compute in ldexp one (-1023)%Z. + +Check (eq_refl : huge / tiny = infinity). +Check (eq_refl infinity <: huge / tiny = infinity). +Check (eq_refl infinity <<: huge / tiny = infinity). +Definition compute2 := Eval compute in huge / tiny. +Check (eq_refl compute2 : infinity = infinity). + +Check (eq_refl : huge / huge = one). +Check (eq_refl one <: huge / huge = one). +Check (eq_refl one <<: huge / huge = one). +Definition compute3 := Eval compute in huge / huge. +Check (eq_refl compute3 : one = one). + +Check (eq_refl : one / nan = nan). +Check (eq_refl nan <: one / nan = nan). +Check (eq_refl nan <<: one / nan = nan). +Definition compute4 := Eval compute in one / nan. +Check (eq_refl compute4 : nan = nan). + +Check (eq_refl : infinity / infinity = nan). +Check (eq_refl nan <: infinity / infinity = nan). +Check (eq_refl nan <<: infinity / infinity = nan). +Definition compute5 := Eval compute in infinity / infinity. +Check (eq_refl compute5 : nan = nan). + +Check (eq_refl : infinity / neg_infinity = nan). +Check (eq_refl nan <: infinity / neg_infinity = nan). +Check (eq_refl nan <<: infinity / neg_infinity = nan). +Definition compute6 := Eval compute in infinity / neg_infinity. +Check (eq_refl compute6 : nan = nan). + +Check (eq_refl : zero / zero = nan). +Check (eq_refl nan <: zero / zero = nan). +Check (eq_refl nan <<: zero / zero = nan). +Check (eq_refl : neg_zero / zero = nan). +Check (eq_refl nan <: neg_zero / zero = nan). +Check (eq_refl nan <<: neg_zero / zero = nan). + +Check (eq_refl : huge / neg_infinity = neg_zero). +Check (eq_refl neg_zero <: huge / neg_infinity = neg_zero). +Check (eq_refl neg_zero <<: huge / neg_infinity = neg_zero). + +Check (eq_refl : one / tiny = huge). +Check (eq_refl huge <: one / tiny = huge). +Check (eq_refl huge <<: one / tiny = huge). +Check (eq_refl : one / huge = tiny). +Check (eq_refl tiny <: one / huge = tiny). +Check (eq_refl tiny <<: one / huge = tiny). +Check (eq_refl : zero / huge = zero). +Check (eq_refl zero <: zero / huge = zero). +Check (eq_refl zero <<: zero / huge = zero). +Check (eq_refl : zero / (-huge) = neg_zero). +Check (eq_refl neg_zero <: zero / (-huge) = neg_zero). +Check (eq_refl neg_zero <<: zero / (-huge) = neg_zero). + +Check (eq_refl : tiny / one = tiny). +Check (eq_refl tiny <: tiny / one = tiny). +Check (eq_refl tiny <<: tiny / one = tiny). +Check (eq_refl : huge / one = huge). +Check (eq_refl huge <: huge / one = huge). +Check (eq_refl huge <<: huge / one = huge). +Check (eq_refl : infinity / one = infinity). +Check (eq_refl infinity <: infinity / one = infinity). +Check (eq_refl infinity <<: infinity / one = infinity). + +Check (eq_refl : zero / infinity = zero). +Check (eq_refl zero <: zero / infinity = zero). +Check (eq_refl zero <<: zero / infinity = zero). +Check (eq_refl : infinity / zero = infinity). +Check (eq_refl infinity <: infinity / zero = infinity). +Check (eq_refl infinity <<: infinity / zero = infinity). diff --git a/test-suite/primitive/float/double_rounding.v b/test-suite/primitive/float/double_rounding.v new file mode 100644 index 0000000000..e2356cdd7b --- /dev/null +++ b/test-suite/primitive/float/double_rounding.v @@ -0,0 +1,38 @@ +Require Import Floats ZArith. + +(* This test check that there is no double rounding with 80 bits registers inside float computations *) + +Definition big_cbn := Eval cbn in ldexp one (53)%Z. +Definition small_cbn := Eval cbn in (one + ldexp one (-52)%Z)%float. +Definition result_cbn := Eval cbn in (big_cbn + small_cbn)%float. +Definition check_cbn := Eval cbn in (big_cbn + one)%float. + +Check (eq_refl : (result_cbn ?= big_cbn)%float = FGt). +Check (eq_refl : (check_cbn ?= big_cbn)%float = FEq). + + +Definition big_cbv := Eval cbv in ldexp one (53)%Z. +Definition small_cbv := Eval cbv in (one + ldexp one (-52)%Z)%float. +Definition result_cbv := Eval cbv in (big_cbv + small_cbv)%float. +Definition check_cbv := Eval cbv in (big_cbv + one)%float. + +Check (eq_refl : (result_cbv ?= big_cbv)%float = FGt). +Check (eq_refl : (check_cbv ?= big_cbv)%float = FEq). + + +Definition big_vm := Eval vm_compute in ldexp one (53)%Z. +Definition small_vm := Eval vm_compute in (one + ldexp one (-52)%Z)%float. +Definition result_vm := Eval vm_compute in (big_vm + small_vm)%float. +Definition check_vm := Eval vm_compute in (big_vm + one)%float. + +Check (eq_refl : (result_vm ?= big_vm)%float = FGt). +Check (eq_refl : (check_vm ?= big_vm)%float = FEq). + + +Definition big_native := Eval native_compute in ldexp one (53)%Z. +Definition small_native := Eval native_compute in (one + ldexp one (-52)%Z)%float. +Definition result_native := Eval native_compute in (big_native + small_native)%float. +Definition check_native := Eval native_compute in (big_native + one)%float. + +Check (eq_refl : (result_native ?= big_native)%float = FGt). +Check (eq_refl : (check_native ?= big_native)%float = FEq). diff --git a/test-suite/primitive/float/frexp.v b/test-suite/primitive/float/frexp.v new file mode 100644 index 0000000000..2a600429b1 --- /dev/null +++ b/test-suite/primitive/float/frexp.v @@ -0,0 +1,28 @@ +Require Import ZArith Floats. + +Definition denorm := Eval compute in ldexp one (-1074)%Z. +Definition neg_one := Eval compute in (-one)%float. + +Check (eq_refl : let (m,e) := frexp infinity in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF infinity)). +Check (eq_refl (SFfrexp prec emax (Prim2SF infinity)) <: let (m,e) := frexp infinity in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF infinity)). +Check (eq_refl (SFfrexp prec emax (Prim2SF infinity)) <<: let (m,e) := frexp infinity in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF infinity)). + +Check (eq_refl : let (m,e) := frexp nan in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF nan)). +Check (eq_refl (SFfrexp prec emax (Prim2SF nan)) <: let (m,e) := frexp nan in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF nan)). +Check (eq_refl (SFfrexp prec emax (Prim2SF nan)) <<: let (m,e) := frexp nan in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF nan)). + +Check (eq_refl : let (m,e) := frexp zero in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF zero)). +Check (eq_refl (SFfrexp prec emax (Prim2SF zero)) <: let (m,e) := frexp zero in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF zero)). +Check (eq_refl (SFfrexp prec emax (Prim2SF zero)) <<: let (m,e) := frexp zero in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF zero)). + +Check (eq_refl : let (m,e) := frexp one in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF one)). +Check (eq_refl (SFfrexp prec emax (Prim2SF one)) <: let (m,e) := frexp one in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF one)). +Check (eq_refl (SFfrexp prec emax (Prim2SF one)) <<: let (m,e) := frexp one in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF one)). + +Check (eq_refl : let (m,e) := frexp neg_one in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF neg_one)). +Check (eq_refl (SFfrexp prec emax (Prim2SF neg_one)) <: let (m,e) := frexp neg_one in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF neg_one)). +Check (eq_refl (SFfrexp prec emax (Prim2SF neg_one)) <<: let (m,e) := frexp neg_one in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF neg_one)). + +Check (eq_refl : let (m,e) := frexp denorm in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF denorm)). +Check (eq_refl (SFfrexp prec emax (Prim2SF denorm)) <: let (m,e) := frexp denorm in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF denorm)). +Check (eq_refl (SFfrexp prec emax (Prim2SF denorm)) <<: let (m,e) := frexp denorm in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF denorm)). diff --git a/test-suite/primitive/float/gen_compare.sh b/test-suite/primitive/float/gen_compare.sh new file mode 100755 index 0000000000..cd87eb4e5b --- /dev/null +++ b/test-suite/primitive/float/gen_compare.sh @@ -0,0 +1,65 @@ +#!/bin/bash +# -*- compile-command: "mv -f compare.v{,~} && ./gen_compare.sh" -*- +set -e + +exec > compare.v + +cat <<EOF +(* DO NOT EDIT THIS FILE: automatically generated by ./gen_compare.sh *) +Require Import ZArith Floats. +Local Open Scope float_scope. + +Definition min_denorm := Eval compute in ldexp one (-1074)%Z. + +Definition min_norm := Eval compute in ldexp one (-1024)%Z. + +EOF + +genTest() { + if [ $# -ne 10 ]; then + echo >&2 "genTest expects 10 arguments" + fi + TACTICS=(":" "<:" "<<:") + OPS=("==" "<" "<=" "?=") + x="$1" + y="$2" + OPS1=("$3" "$4" "$5" "$6") # for x y + OPS2=("$7" "$8" "$9" "${10}") # for y x + for tac in "${TACTICS[@]}"; do + for i in {0..3}; do + op="${OPS[$i]}" + op1="${OPS1[$i]}" + op2="${OPS2[$i]}" + echo "Check (eq_refl $op1 $tac $x $op $y = $op1)." + echo "Check (eq_refl $op2 $tac $y $op $x = $op2)." + done + echo + done +} + +genTest nan nan \ + false false false FNotComparable \ + false false false FNotComparable +genTest nan "- nan" \ + false false false FNotComparable \ + false false false FNotComparable + +EQ=(true false true FEq \ + true false true FEq) + +genTest one one "${EQ[@]}" +genTest zero zero "${EQ[@]}" +genTest zero "- zero" "${EQ[@]}" +genTest "- zero" "- zero" "${EQ[@]}" +genTest infinity infinity "${EQ[@]}" +genTest "- infinity" "- infinity" "${EQ[@]}" + +LT=(false true true FLt \ + false false false FGt) + +genTest min_denorm min_norm "${LT[@]}" +genTest min_denorm one "${LT[@]}" +genTest min_norm one "${LT[@]}" +genTest one infinity "${LT[@]}" +genTest "- infinity" infinity "${LT[@]}" +genTest "- infinity" one "${LT[@]}" diff --git a/test-suite/primitive/float/ldexp.v b/test-suite/primitive/float/ldexp.v new file mode 100644 index 0000000000..a725deeeca --- /dev/null +++ b/test-suite/primitive/float/ldexp.v @@ -0,0 +1,21 @@ +Require Import ZArith Int63 Floats. + +Check (eq_refl : ldexp one 9223372036854773807%Z = infinity). +Check (eq_refl infinity <: ldexp one 9223372036854773807%Z = infinity). +Check (eq_refl infinity <<: ldexp one 9223372036854773807%Z = infinity). + +Check (eq_refl : ldshiftexp one 9223372036854775807 = infinity). +Check (eq_refl infinity <: ldshiftexp one 9223372036854775807 = infinity). +Check (eq_refl infinity <<: ldshiftexp one 9223372036854775807 = infinity). + +Check (eq_refl : ldexp one (-2102) = 0%float). +Check (eq_refl 0%float <: ldexp one (-2102) = 0%float). +Check (eq_refl 0%float <<: ldexp one (-2102) = 0%float). + +Check (eq_refl : ldexp one (-3) = 0.125%float). +Check (eq_refl 0.125%float <: ldexp one (-3) = 0.125%float). +Check (eq_refl 0.125%float <<: ldexp one (-3) = 0.125%float). + +Check (eq_refl : ldexp one 3 = 8%float). +Check (eq_refl 8%float <: ldexp one 3 = 8%float). +Check (eq_refl 8%float <<: ldexp one 3 = 8%float). diff --git a/test-suite/primitive/float/mul.v b/test-suite/primitive/float/mul.v new file mode 100644 index 0000000000..91fe7e9791 --- /dev/null +++ b/test-suite/primitive/float/mul.v @@ -0,0 +1,83 @@ +Require Import ZArith Int63 Floats. + +Open Scope float_scope. + +Definition two := Eval compute in of_int63 2%int63. +Definition three := Eval compute in of_int63 3%int63. +Definition six := Eval compute in of_int63 6%int63. + +Check (eq_refl : three * two = six). +Check (eq_refl six <: three * two = six). +Check (eq_refl six <<: three * two = six). +Definition compute1 := Eval compute in three * two. +Check (eq_refl compute1 : six = six). + +Definition huge := Eval compute in ldexp one 1023%Z. +Definition tiny := Eval compute in ldexp one (-1023)%Z. + +Check (eq_refl : huge * tiny = one). +Check (eq_refl one <: huge * tiny = one). +Check (eq_refl one <<: huge * tiny = one). +Definition compute2 := Eval compute in huge * tiny. +Check (eq_refl compute2 : one = one). + +Check (eq_refl : huge * huge = infinity). +Check (eq_refl infinity <: huge * huge = infinity). +Check (eq_refl infinity <<: huge * huge = infinity). +Definition compute3 := Eval compute in huge * huge. +Check (eq_refl compute3 : infinity = infinity). + +Check (eq_refl : one * nan = nan). +Check (eq_refl nan <: one * nan = nan). +Check (eq_refl nan <<: one * nan = nan). +Definition compute4 := Eval compute in one * nan. +Check (eq_refl compute4 : nan = nan). + +Check (eq_refl : infinity * infinity = infinity). +Check (eq_refl infinity <: infinity * infinity = infinity). +Check (eq_refl infinity <<: infinity * infinity = infinity). +Definition compute5 := Eval compute in infinity * infinity. +Check (eq_refl compute5 : infinity = infinity). + +Check (eq_refl : infinity * neg_infinity = neg_infinity). +Check (eq_refl neg_infinity <: infinity * neg_infinity = neg_infinity). +Check (eq_refl neg_infinity <<: infinity * neg_infinity = neg_infinity). +Definition compute6 := Eval compute in infinity * neg_infinity. +Check (eq_refl compute6 : neg_infinity = neg_infinity). + +Check (eq_refl : zero * zero = zero). +Check (eq_refl zero <: zero * zero = zero). +Check (eq_refl zero <<: zero * zero = zero). +Check (eq_refl : neg_zero * zero = neg_zero). +Check (eq_refl neg_zero <: neg_zero * zero = neg_zero). +Check (eq_refl neg_zero <<: neg_zero * zero = neg_zero). +Check (eq_refl : neg_zero * neg_zero = zero). +Check (eq_refl zero <: neg_zero * neg_zero = zero). +Check (eq_refl zero <<: neg_zero * neg_zero = zero). +Check (eq_refl : zero * neg_zero = neg_zero). +Check (eq_refl neg_zero <: zero * neg_zero = neg_zero). +Check (eq_refl neg_zero <<: zero * neg_zero = neg_zero). + +Check (eq_refl : huge * neg_infinity = neg_infinity). +Check (eq_refl neg_infinity <: huge * neg_infinity = neg_infinity). +Check (eq_refl neg_infinity <<: huge * neg_infinity = neg_infinity). + +Check (eq_refl : one * tiny = tiny). +Check (eq_refl tiny <: one * tiny = tiny). +Check (eq_refl tiny <<: one * tiny = tiny). +Check (eq_refl : one * huge = huge). +Check (eq_refl huge <: one * huge = huge). +Check (eq_refl huge <<: one * huge = huge). +Check (eq_refl : zero * huge = zero). +Check (eq_refl zero <: zero * huge = zero). +Check (eq_refl zero <<: zero * huge = zero). +Check (eq_refl : zero * (-huge) = neg_zero). +Check (eq_refl neg_zero <: zero * (-huge) = neg_zero). +Check (eq_refl neg_zero <<: zero * (-huge) = neg_zero). + +Check (eq_refl : zero * infinity = nan). +Check (eq_refl nan <: zero * infinity = nan). +Check (eq_refl nan <<: zero * infinity = nan). +Check (eq_refl : neg_infinity * zero = nan). +Check (eq_refl nan <: neg_infinity * zero = nan). +Check (eq_refl nan <<: neg_infinity * zero = nan). diff --git a/test-suite/primitive/float/next_up_down.v b/test-suite/primitive/float/next_up_down.v new file mode 100644 index 0000000000..4f8427dc5b --- /dev/null +++ b/test-suite/primitive/float/next_up_down.v @@ -0,0 +1,122 @@ +Require Import ZArith Int63 Floats. + +Open Scope float_scope. + +Definition f0 := zero. +Definition f1 := neg_zero. +Definition f2 := Eval compute in ldexp one 0. +Definition f3 := Eval compute in -f1. +(* smallest positive float *) +Definition f4 := Eval compute in ldexp one (-1074). +Definition f5 := Eval compute in -f3. +Definition f6 := infinity. +Definition f7 := neg_infinity. +Definition f8 := Eval compute in ldexp one (-1). +Definition f9 := Eval compute in -f8. +Definition f10 := Eval compute in of_int63 42. +Definition f11 := Eval compute in -f10. +(* max float *) +Definition f12 := Eval compute in ldexp (of_int63 9007199254740991) 1024. +Definition f13 := Eval compute in -f12. +(* smallest positive normalized float *) +Definition f14 := Eval compute in ldexp one (-1022). +Definition f15 := Eval compute in -f14. + +Check (eq_refl : Prim2SF (next_up f0) = SF64succ (Prim2SF f0)). +Check (eq_refl : Prim2SF (next_down f0) = SF64pred (Prim2SF f0)). +Check (eq_refl : Prim2SF (next_up f1) = SF64succ (Prim2SF f1)). +Check (eq_refl : Prim2SF (next_down f1) = SF64pred (Prim2SF f1)). +Check (eq_refl : Prim2SF (next_up f2) = SF64succ (Prim2SF f2)). +Check (eq_refl : Prim2SF (next_down f2) = SF64pred (Prim2SF f2)). +Check (eq_refl : Prim2SF (next_up f3) = SF64succ (Prim2SF f3)). +Check (eq_refl : Prim2SF (next_down f3) = SF64pred (Prim2SF f3)). +Check (eq_refl : Prim2SF (next_up f4) = SF64succ (Prim2SF f4)). +Check (eq_refl : Prim2SF (next_down f4) = SF64pred (Prim2SF f4)). +Check (eq_refl : Prim2SF (next_up f5) = SF64succ (Prim2SF f5)). +Check (eq_refl : Prim2SF (next_down f5) = SF64pred (Prim2SF f5)). +Check (eq_refl : Prim2SF (next_up f6) = SF64succ (Prim2SF f6)). +Check (eq_refl : Prim2SF (next_down f6) = SF64pred (Prim2SF f6)). +Check (eq_refl : Prim2SF (next_up f7) = SF64succ (Prim2SF f7)). +Check (eq_refl : Prim2SF (next_down f7) = SF64pred (Prim2SF f7)). +Check (eq_refl : Prim2SF (next_up f8) = SF64succ (Prim2SF f8)). +Check (eq_refl : Prim2SF (next_down f8) = SF64pred (Prim2SF f8)). +Check (eq_refl : Prim2SF (next_up f9) = SF64succ (Prim2SF f9)). +Check (eq_refl : Prim2SF (next_down f9) = SF64pred (Prim2SF f9)). +Check (eq_refl : Prim2SF (next_up f10) = SF64succ (Prim2SF f10)). +Check (eq_refl : Prim2SF (next_down f10) = SF64pred (Prim2SF f10)). +Check (eq_refl : Prim2SF (next_up f11) = SF64succ (Prim2SF f11)). +Check (eq_refl : Prim2SF (next_down f11) = SF64pred (Prim2SF f11)). +Check (eq_refl : Prim2SF (next_up f12) = SF64succ (Prim2SF f12)). +Check (eq_refl : Prim2SF (next_down f12) = SF64pred (Prim2SF f12)). +Check (eq_refl : Prim2SF (next_up f13) = SF64succ (Prim2SF f13)). +Check (eq_refl : Prim2SF (next_down f13) = SF64pred (Prim2SF f13)). +Check (eq_refl : Prim2SF (next_up f14) = SF64succ (Prim2SF f14)). +Check (eq_refl : Prim2SF (next_down f14) = SF64pred (Prim2SF f14)). +Check (eq_refl : Prim2SF (next_up f15) = SF64succ (Prim2SF f15)). +Check (eq_refl : Prim2SF (next_down f15) = SF64pred (Prim2SF f15)). + +Check (eq_refl (SF64succ (Prim2SF f0)) <: Prim2SF (next_up f0) = SF64succ (Prim2SF f0)). +Check (eq_refl (SF64pred (Prim2SF f0)) <: Prim2SF (next_down f0) = SF64pred (Prim2SF f0)). +Check (eq_refl (SF64succ (Prim2SF f1)) <: Prim2SF (next_up f1) = SF64succ (Prim2SF f1)). +Check (eq_refl (SF64pred (Prim2SF f1)) <: Prim2SF (next_down f1) = SF64pred (Prim2SF f1)). +Check (eq_refl (SF64succ (Prim2SF f2)) <: Prim2SF (next_up f2) = SF64succ (Prim2SF f2)). +Check (eq_refl (SF64pred (Prim2SF f2)) <: Prim2SF (next_down f2) = SF64pred (Prim2SF f2)). +Check (eq_refl (SF64succ (Prim2SF f3)) <: Prim2SF (next_up f3) = SF64succ (Prim2SF f3)). +Check (eq_refl (SF64pred (Prim2SF f3)) <: Prim2SF (next_down f3) = SF64pred (Prim2SF f3)). +Check (eq_refl (SF64succ (Prim2SF f4)) <: Prim2SF (next_up f4) = SF64succ (Prim2SF f4)). +Check (eq_refl (SF64pred (Prim2SF f4)) <: Prim2SF (next_down f4) = SF64pred (Prim2SF f4)). +Check (eq_refl (SF64succ (Prim2SF f5)) <: Prim2SF (next_up f5) = SF64succ (Prim2SF f5)). +Check (eq_refl (SF64pred (Prim2SF f5)) <: Prim2SF (next_down f5) = SF64pred (Prim2SF f5)). +Check (eq_refl (SF64succ (Prim2SF f6)) <: Prim2SF (next_up f6) = SF64succ (Prim2SF f6)). +Check (eq_refl (SF64pred (Prim2SF f6)) <: Prim2SF (next_down f6) = SF64pred (Prim2SF f6)). +Check (eq_refl (SF64succ (Prim2SF f7)) <: Prim2SF (next_up f7) = SF64succ (Prim2SF f7)). +Check (eq_refl (SF64pred (Prim2SF f7)) <: Prim2SF (next_down f7) = SF64pred (Prim2SF f7)). +Check (eq_refl (SF64succ (Prim2SF f8)) <: Prim2SF (next_up f8) = SF64succ (Prim2SF f8)). +Check (eq_refl (SF64pred (Prim2SF f8)) <: Prim2SF (next_down f8) = SF64pred (Prim2SF f8)). +Check (eq_refl (SF64succ (Prim2SF f9)) <: Prim2SF (next_up f9) = SF64succ (Prim2SF f9)). +Check (eq_refl (SF64pred (Prim2SF f9)) <: Prim2SF (next_down f9) = SF64pred (Prim2SF f9)). +Check (eq_refl (SF64succ (Prim2SF f10)) <: Prim2SF (next_up f10) = SF64succ (Prim2SF f10)). +Check (eq_refl (SF64pred (Prim2SF f10)) <: Prim2SF (next_down f10) = SF64pred (Prim2SF f10)). +Check (eq_refl (SF64succ (Prim2SF f11)) <: Prim2SF (next_up f11) = SF64succ (Prim2SF f11)). +Check (eq_refl (SF64pred (Prim2SF f11)) <: Prim2SF (next_down f11) = SF64pred (Prim2SF f11)). +Check (eq_refl (SF64succ (Prim2SF f12)) <: Prim2SF (next_up f12) = SF64succ (Prim2SF f12)). +Check (eq_refl (SF64pred (Prim2SF f12)) <: Prim2SF (next_down f12) = SF64pred (Prim2SF f12)). +Check (eq_refl (SF64succ (Prim2SF f13)) <: Prim2SF (next_up f13) = SF64succ (Prim2SF f13)). +Check (eq_refl (SF64pred (Prim2SF f13)) <: Prim2SF (next_down f13) = SF64pred (Prim2SF f13)). +Check (eq_refl (SF64succ (Prim2SF f14)) <: Prim2SF (next_up f14) = SF64succ (Prim2SF f14)). +Check (eq_refl (SF64pred (Prim2SF f14)) <: Prim2SF (next_down f14) = SF64pred (Prim2SF f14)). +Check (eq_refl (SF64succ (Prim2SF f15)) <: Prim2SF (next_up f15) = SF64succ (Prim2SF f15)). +Check (eq_refl (SF64pred (Prim2SF f15)) <: Prim2SF (next_down f15) = SF64pred (Prim2SF f15)). + +Check (eq_refl (SF64succ (Prim2SF f0)) <<: Prim2SF (next_up f0) = SF64succ (Prim2SF f0)). +Check (eq_refl (SF64pred (Prim2SF f0)) <<: Prim2SF (next_down f0) = SF64pred (Prim2SF f0)). +Check (eq_refl (SF64succ (Prim2SF f1)) <<: Prim2SF (next_up f1) = SF64succ (Prim2SF f1)). +Check (eq_refl (SF64pred (Prim2SF f1)) <<: Prim2SF (next_down f1) = SF64pred (Prim2SF f1)). +Check (eq_refl (SF64succ (Prim2SF f2)) <<: Prim2SF (next_up f2) = SF64succ (Prim2SF f2)). +Check (eq_refl (SF64pred (Prim2SF f2)) <<: Prim2SF (next_down f2) = SF64pred (Prim2SF f2)). +Check (eq_refl (SF64succ (Prim2SF f3)) <<: Prim2SF (next_up f3) = SF64succ (Prim2SF f3)). +Check (eq_refl (SF64pred (Prim2SF f3)) <<: Prim2SF (next_down f3) = SF64pred (Prim2SF f3)). +Check (eq_refl (SF64succ (Prim2SF f4)) <<: Prim2SF (next_up f4) = SF64succ (Prim2SF f4)). +Check (eq_refl (SF64pred (Prim2SF f4)) <<: Prim2SF (next_down f4) = SF64pred (Prim2SF f4)). +Check (eq_refl (SF64succ (Prim2SF f5)) <<: Prim2SF (next_up f5) = SF64succ (Prim2SF f5)). +Check (eq_refl (SF64pred (Prim2SF f5)) <<: Prim2SF (next_down f5) = SF64pred (Prim2SF f5)). +Check (eq_refl (SF64succ (Prim2SF f6)) <<: Prim2SF (next_up f6) = SF64succ (Prim2SF f6)). +Check (eq_refl (SF64pred (Prim2SF f6)) <<: Prim2SF (next_down f6) = SF64pred (Prim2SF f6)). +Check (eq_refl (SF64succ (Prim2SF f7)) <<: Prim2SF (next_up f7) = SF64succ (Prim2SF f7)). +Check (eq_refl (SF64pred (Prim2SF f7)) <<: Prim2SF (next_down f7) = SF64pred (Prim2SF f7)). +Check (eq_refl (SF64succ (Prim2SF f8)) <<: Prim2SF (next_up f8) = SF64succ (Prim2SF f8)). +Check (eq_refl (SF64pred (Prim2SF f8)) <<: Prim2SF (next_down f8) = SF64pred (Prim2SF f8)). +Check (eq_refl (SF64succ (Prim2SF f9)) <<: Prim2SF (next_up f9) = SF64succ (Prim2SF f9)). +Check (eq_refl (SF64pred (Prim2SF f9)) <<: Prim2SF (next_down f9) = SF64pred (Prim2SF f9)). +Check (eq_refl (SF64succ (Prim2SF f10)) <<: Prim2SF (next_up f10) = SF64succ (Prim2SF f10)). +Check (eq_refl (SF64pred (Prim2SF f10)) <<: Prim2SF (next_down f10) = SF64pred (Prim2SF f10)). +Check (eq_refl (SF64succ (Prim2SF f11)) <<: Prim2SF (next_up f11) = SF64succ (Prim2SF f11)). +Check (eq_refl (SF64pred (Prim2SF f11)) <<: Prim2SF (next_down f11) = SF64pred (Prim2SF f11)). +Check (eq_refl (SF64succ (Prim2SF f12)) <<: Prim2SF (next_up f12) = SF64succ (Prim2SF f12)). +Check (eq_refl (SF64pred (Prim2SF f12)) <<: Prim2SF (next_down f12) = SF64pred (Prim2SF f12)). +Check (eq_refl (SF64succ (Prim2SF f13)) <<: Prim2SF (next_up f13) = SF64succ (Prim2SF f13)). +Check (eq_refl (SF64pred (Prim2SF f13)) <<: Prim2SF (next_down f13) = SF64pred (Prim2SF f13)). +Check (eq_refl (SF64succ (Prim2SF f14)) <<: Prim2SF (next_up f14) = SF64succ (Prim2SF f14)). +Check (eq_refl (SF64pred (Prim2SF f14)) <<: Prim2SF (next_down f14) = SF64pred (Prim2SF f14)). +Check (eq_refl (SF64succ (Prim2SF f15)) <<: Prim2SF (next_up f15) = SF64succ (Prim2SF f15)). +Check (eq_refl (SF64pred (Prim2SF f15)) <<: Prim2SF (next_down f15) = SF64pred (Prim2SF f15)). diff --git a/test-suite/primitive/float/normfr_mantissa.v b/test-suite/primitive/float/normfr_mantissa.v new file mode 100644 index 0000000000..28bd1c03f5 --- /dev/null +++ b/test-suite/primitive/float/normfr_mantissa.v @@ -0,0 +1,28 @@ +Require Import Int63 ZArith Floats. + +Definition half := ldexp one (-1)%Z. +Definition three_quarters := (half + (ldexp one (-2)%Z))%float. + +Check (eq_refl : normfr_mantissa one = 0%int63). +Check (eq_refl : normfr_mantissa half = (1 << 52)%int63). +Check (eq_refl : normfr_mantissa (-half) = (1 << 52)%int63). +Check (eq_refl : normfr_mantissa (-one) = 0%int63). +Check (eq_refl : normfr_mantissa zero = 0%int63). +Check (eq_refl : normfr_mantissa nan = 0%int63). +Check (eq_refl : normfr_mantissa three_quarters = (3 << 51)%int63). + +Check (eq_refl 0%int63 <: normfr_mantissa one = 0%int63). +Check (eq_refl (1 << 52)%int63 <: normfr_mantissa half = (1 << 52)%int63). +Check (eq_refl (1 << 52)%int63 <: normfr_mantissa (-half) = (1 << 52)%int63). +Check (eq_refl 0%int63 <: normfr_mantissa (-one) = 0%int63). +Check (eq_refl 0%int63 <: normfr_mantissa zero = 0%int63). +Check (eq_refl 0%int63 <: normfr_mantissa nan = 0%int63). +Check (eq_refl (3 << 51)%int63 <: normfr_mantissa three_quarters = (3 << 51)%int63). + +Check (eq_refl 0%int63 <<: normfr_mantissa one = 0%int63). +Check (eq_refl (1 << 52)%int63 <<: normfr_mantissa half = (1 << 52)%int63). +Check (eq_refl (1 << 52)%int63 <<: normfr_mantissa (-half) = (1 << 52)%int63). +Check (eq_refl 0%int63 <<: normfr_mantissa (-one) = 0%int63). +Check (eq_refl 0%int63 <<: normfr_mantissa zero = 0%int63). +Check (eq_refl 0%int63 <<: normfr_mantissa nan = 0%int63). +Check (eq_refl (3 << 51)%int63 <<: normfr_mantissa three_quarters = (3 << 51)%int63). diff --git a/test-suite/primitive/float/spec_conv.v b/test-suite/primitive/float/spec_conv.v new file mode 100644 index 0000000000..a610d39671 --- /dev/null +++ b/test-suite/primitive/float/spec_conv.v @@ -0,0 +1,46 @@ +Require Import ZArith Floats. + +Definition two := Eval compute in (one + one)%float. +Definition half := Eval compute in (one / two)%float. +Definition huge := Eval compute in ldexp one (1023)%Z. +Definition tiny := Eval compute in ldexp one (-1023)%Z. +Definition denorm := Eval compute in ldexp one (-1074)%Z. + +Check (eq_refl : SF2Prim (Prim2SF zero) = zero). +Check (eq_refl : SF2Prim (Prim2SF neg_zero) = neg_zero). +Check (eq_refl : SF2Prim (Prim2SF one) = one). +Check (eq_refl : SF2Prim (Prim2SF (-one)) = (-one)%float). +Check (eq_refl : SF2Prim (Prim2SF infinity) = infinity). +Check (eq_refl : SF2Prim (Prim2SF neg_infinity) = neg_infinity). +Check (eq_refl : SF2Prim (Prim2SF huge) = huge). +Check (eq_refl : SF2Prim (Prim2SF tiny) = tiny). +Check (eq_refl : SF2Prim (Prim2SF denorm) = denorm). +Check (eq_refl : SF2Prim (Prim2SF nan) = nan). +Check (eq_refl : SF2Prim (Prim2SF two) = two). +Check (eq_refl : SF2Prim (Prim2SF half) = half). + +Check (eq_refl zero <: SF2Prim (Prim2SF zero) = zero). +Check (eq_refl neg_zero <: SF2Prim (Prim2SF neg_zero) = neg_zero). +Check (eq_refl one <: SF2Prim (Prim2SF one) = one). +Check (eq_refl (-one)%float <: SF2Prim (Prim2SF (-one)) = (-one)%float). +Check (eq_refl infinity <: SF2Prim (Prim2SF infinity) = infinity). +Check (eq_refl neg_infinity <: SF2Prim (Prim2SF neg_infinity) = neg_infinity). +Check (eq_refl huge <: SF2Prim (Prim2SF huge) = huge). +Check (eq_refl tiny <: SF2Prim (Prim2SF tiny) = tiny). +Check (eq_refl denorm <: SF2Prim (Prim2SF denorm) = denorm). +Check (eq_refl nan <: SF2Prim (Prim2SF nan) = nan). +Check (eq_refl two <: SF2Prim (Prim2SF two) = two). +Check (eq_refl half <: SF2Prim (Prim2SF half) = half). + +Check (eq_refl zero <<: SF2Prim (Prim2SF zero) = zero). +Check (eq_refl neg_zero <<: SF2Prim (Prim2SF neg_zero) = neg_zero). +Check (eq_refl one <<: SF2Prim (Prim2SF one) = one). +Check (eq_refl (-one)%float <<: SF2Prim (Prim2SF (-one)) = (-one)%float). +Check (eq_refl infinity <<: SF2Prim (Prim2SF infinity) = infinity). +Check (eq_refl neg_infinity <<: SF2Prim (Prim2SF neg_infinity) = neg_infinity). +Check (eq_refl huge <<: SF2Prim (Prim2SF huge) = huge). +Check (eq_refl tiny <<: SF2Prim (Prim2SF tiny) = tiny). +Check (eq_refl denorm <<: SF2Prim (Prim2SF denorm) = denorm). +Check (eq_refl nan <<: SF2Prim (Prim2SF nan) = nan). +Check (eq_refl two <<: SF2Prim (Prim2SF two) = two). +Check (eq_refl half <<: SF2Prim (Prim2SF half) = half). diff --git a/test-suite/primitive/float/sqrt.v b/test-suite/primitive/float/sqrt.v new file mode 100644 index 0000000000..04c8ab035d --- /dev/null +++ b/test-suite/primitive/float/sqrt.v @@ -0,0 +1,49 @@ +Require Import ZArith Int63 Floats. + +Open Scope float_scope. + +Definition three := Eval compute in of_int63 3%int63. +Definition nine := Eval compute in of_int63 9%int63. + +Check (eq_refl : sqrt nine = three). +Check (eq_refl three <: sqrt nine = three). +Definition compute1 := Eval compute in sqrt nine. +Check (eq_refl : three = three). + +Definition huge := Eval compute in ldexp one (1023)%Z. +Definition tiny := Eval compute in ldexp one (-1023)%Z. +Definition denorm := Eval compute in ldexp one (-1074)%Z. + +Goal (Prim2SF (sqrt huge) = SF64sqrt (Prim2SF huge)). + now compute. Undo. now vm_compute. +Qed. + +Goal (Prim2SF (sqrt tiny) = SF64sqrt (Prim2SF tiny)). + now compute. Undo. now vm_compute. +Qed. + +Goal (Prim2SF (sqrt denorm) = SF64sqrt (Prim2SF denorm)). + now compute. Undo. now vm_compute. +Qed. + +Check (eq_refl : sqrt neg_zero = neg_zero). +Check (eq_refl neg_zero <: sqrt neg_zero = neg_zero). +Check (eq_refl neg_zero <<: sqrt neg_zero = neg_zero). +Check (eq_refl : sqrt zero = zero). +Check (eq_refl zero <: sqrt zero = zero). +Check (eq_refl zero <<: sqrt zero = zero). +Check (eq_refl : sqrt one = one). +Check (eq_refl one <: sqrt one = one). +Check (eq_refl one <<: sqrt one = one). +Check (eq_refl : sqrt (-one) = nan). +Check (eq_refl nan <: sqrt (-one) = nan). +Check (eq_refl nan <<: sqrt (-one) = nan). +Check (eq_refl : sqrt infinity = infinity). +Check (eq_refl infinity <: sqrt infinity = infinity). +Check (eq_refl infinity <<: sqrt infinity = infinity). +Check (eq_refl : sqrt neg_infinity = nan). +Check (eq_refl nan <: sqrt neg_infinity = nan). +Check (eq_refl nan <<: sqrt neg_infinity = nan). +Check (eq_refl : sqrt infinity = infinity). +Check (eq_refl infinity <: sqrt infinity = infinity). +Check (eq_refl infinity <<: sqrt infinity = infinity). diff --git a/test-suite/primitive/float/sub.v b/test-suite/primitive/float/sub.v new file mode 100644 index 0000000000..fc068cb585 --- /dev/null +++ b/test-suite/primitive/float/sub.v @@ -0,0 +1,62 @@ +Require Import ZArith Int63 Floats. + +Open Scope float_scope. + +Definition two := Eval compute in of_int63 2%int63. +Definition three := Eval compute in of_int63 3%int63. + +Check (eq_refl : three - two = one). +Check (eq_refl one <: three - two = one). +Check (eq_refl one <<: three - two = one). +Definition compute1 := Eval compute in three - two. +Check (eq_refl compute1 : one = one). + +Definition huge := Eval compute in ldexp one 1023%Z. +Definition tiny := Eval compute in ldexp one (-1023)%Z. + +Check (eq_refl : huge - tiny = huge). +Check (eq_refl huge <: huge - tiny = huge). +Check (eq_refl huge <<: huge - tiny = huge). +Definition compute2 := Eval compute in huge - tiny. +Check (eq_refl compute2 : huge = huge). + +Check (eq_refl : huge - huge = zero). +Check (eq_refl zero <: huge - huge = zero). +Check (eq_refl zero <<: huge - huge = zero). +Definition compute3 := Eval compute in huge - huge. +Check (eq_refl compute3 : zero = zero). + +Check (eq_refl : one - nan = nan). +Check (eq_refl nan <: one - nan = nan). +Check (eq_refl nan <<: one - nan = nan). +Definition compute4 := Eval compute in one - nan. +Check (eq_refl compute4 : nan = nan). + +Check (eq_refl : infinity - infinity = nan). +Check (eq_refl nan <: infinity - infinity = nan). +Check (eq_refl nan <<: infinity - infinity = nan). +Definition compute5 := Eval compute in infinity - infinity. +Check (eq_refl compute5 : nan = nan). + +Check (eq_refl : infinity - neg_infinity = infinity). +Check (eq_refl infinity <: infinity - neg_infinity = infinity). +Check (eq_refl infinity <<: infinity - neg_infinity = infinity). +Definition compute6 := Eval compute in infinity - neg_infinity. +Check (eq_refl compute6 : infinity = infinity). + +Check (eq_refl : zero - zero = zero). +Check (eq_refl zero <: zero - zero = zero). +Check (eq_refl zero <<: zero - zero = zero). +Check (eq_refl : neg_zero - zero = neg_zero). +Check (eq_refl neg_zero <: neg_zero - zero = neg_zero). +Check (eq_refl neg_zero <<: neg_zero - zero = neg_zero). +Check (eq_refl : neg_zero - neg_zero = zero). +Check (eq_refl zero <: neg_zero - neg_zero = zero). +Check (eq_refl zero <<: neg_zero - neg_zero = zero). +Check (eq_refl : zero - neg_zero = zero). +Check (eq_refl zero <: zero - neg_zero = zero). +Check (eq_refl zero <<: zero - neg_zero = zero). + +Check (eq_refl : huge - neg_infinity = infinity). +Check (eq_refl infinity <: huge - neg_infinity = infinity). +Check (eq_refl infinity <<: huge - neg_infinity = infinity). diff --git a/test-suite/primitive/float/syntax.v b/test-suite/primitive/float/syntax.v new file mode 100644 index 0000000000..cc0bbcf628 --- /dev/null +++ b/test-suite/primitive/float/syntax.v @@ -0,0 +1,13 @@ +Require Import Floats. + +Open Scope float_scope. + +Definition two := Eval compute in one + one. +Definition half := Eval compute in one / two. + +Check (eq_refl : 1.5 = one + half). +Check (eq_refl : 15e-1 = one + half). +Check (eq_refl : 150e-2 = one + half). +Check (eq_refl : 0.15e+1 = one + half). +Check (eq_refl : 0.15e1 = one + half). +Check (eq_refl : 0.0015e3 = one + half). diff --git a/test-suite/primitive/float/valid_binary_conv.v b/test-suite/primitive/float/valid_binary_conv.v new file mode 100644 index 0000000000..82e00b8532 --- /dev/null +++ b/test-suite/primitive/float/valid_binary_conv.v @@ -0,0 +1,46 @@ +Require Import ZArith Floats. + +Definition two := Eval compute in (one + one)%float. +Definition half := Eval compute in (one / two)%float. +Definition huge := Eval compute in ldexp one (1023)%Z. +Definition tiny := Eval compute in ldexp one (-1022)%Z. +Definition denorm := Eval compute in ldexp one (-1074)%Z. + +Check (eq_refl : valid_binary (Prim2SF zero) = true). +Check (eq_refl : valid_binary (Prim2SF neg_zero) = true). +Check (eq_refl : valid_binary (Prim2SF one) = true). +Check (eq_refl : valid_binary (Prim2SF (-one)) = true). +Check (eq_refl : valid_binary (Prim2SF infinity) = true). +Check (eq_refl : valid_binary (Prim2SF neg_infinity) = true). +Check (eq_refl : valid_binary (Prim2SF huge) = true). +Check (eq_refl : valid_binary (Prim2SF tiny) = true). +Check (eq_refl : valid_binary (Prim2SF denorm) = true). +Check (eq_refl : valid_binary (Prim2SF nan) = true). +Check (eq_refl : valid_binary (Prim2SF two) = true). +Check (eq_refl : valid_binary (Prim2SF half) = true). + +Check (eq_refl true <: valid_binary (Prim2SF zero) = true). +Check (eq_refl true <: valid_binary (Prim2SF neg_zero) = true). +Check (eq_refl true <: valid_binary (Prim2SF one) = true). +Check (eq_refl true <: valid_binary (Prim2SF (-one)) = true). +Check (eq_refl true <: valid_binary (Prim2SF infinity) = true). +Check (eq_refl true <: valid_binary (Prim2SF neg_infinity) = true). +Check (eq_refl true <: valid_binary (Prim2SF huge) = true). +Check (eq_refl true <: valid_binary (Prim2SF tiny) = true). +Check (eq_refl true <: valid_binary (Prim2SF denorm) = true). +Check (eq_refl true <: valid_binary (Prim2SF nan) = true). +Check (eq_refl true <: valid_binary (Prim2SF two) = true). +Check (eq_refl true <: valid_binary (Prim2SF half) = true). + +Check (eq_refl true <<: valid_binary (Prim2SF zero) = true). +Check (eq_refl true <<: valid_binary (Prim2SF neg_zero) = true). +Check (eq_refl true <<: valid_binary (Prim2SF one) = true). +Check (eq_refl true <<: valid_binary (Prim2SF (-one)) = true). +Check (eq_refl true <<: valid_binary (Prim2SF infinity) = true). +Check (eq_refl true <<: valid_binary (Prim2SF neg_infinity) = true). +Check (eq_refl true <<: valid_binary (Prim2SF huge) = true). +Check (eq_refl true <<: valid_binary (Prim2SF tiny) = true). +Check (eq_refl true <<: valid_binary (Prim2SF denorm) = true). +Check (eq_refl true <<: valid_binary (Prim2SF nan) = true). +Check (eq_refl true <<: valid_binary (Prim2SF two) = true). +Check (eq_refl true <<: valid_binary (Prim2SF half) = true). diff --git a/test-suite/primitive/float/zero.v b/test-suite/primitive/float/zero.v new file mode 100644 index 0000000000..75348d4657 --- /dev/null +++ b/test-suite/primitive/float/zero.v @@ -0,0 +1,7 @@ +Require Import ZArith Int63 Floats. + +Open Scope float_scope. + +Fail Check (eq_refl : zero = neg_zero). +Fail Check (eq_refl <: zero = neg_zero). +Fail Check (eq_refl <<: zero = neg_zero). diff --git a/test-suite/arithmetic/add.v b/test-suite/primitive/uint63/add.v index fb7eb1d53c..fb7eb1d53c 100644 --- a/test-suite/arithmetic/add.v +++ b/test-suite/primitive/uint63/add.v diff --git a/test-suite/arithmetic/addc.v b/test-suite/primitive/uint63/addc.v index 432aec0291..432aec0291 100644 --- a/test-suite/arithmetic/addc.v +++ b/test-suite/primitive/uint63/addc.v diff --git a/test-suite/arithmetic/addcarryc.v b/test-suite/primitive/uint63/addcarryc.v index a4430769ca..a4430769ca 100644 --- a/test-suite/arithmetic/addcarryc.v +++ b/test-suite/primitive/uint63/addcarryc.v diff --git a/test-suite/arithmetic/addmuldiv.v b/test-suite/primitive/uint63/addmuldiv.v index 72b0164b49..72b0164b49 100644 --- a/test-suite/arithmetic/addmuldiv.v +++ b/test-suite/primitive/uint63/addmuldiv.v diff --git a/test-suite/arithmetic/compare.v b/test-suite/primitive/uint63/compare.v index a8d1ea1226..a8d1ea1226 100644 --- a/test-suite/arithmetic/compare.v +++ b/test-suite/primitive/uint63/compare.v diff --git a/test-suite/arithmetic/div.v b/test-suite/primitive/uint63/div.v index 0ee0b91580..0ee0b91580 100644 --- a/test-suite/arithmetic/div.v +++ b/test-suite/primitive/uint63/div.v diff --git a/test-suite/arithmetic/diveucl.v b/test-suite/primitive/uint63/diveucl.v index 8f88a0f356..8f88a0f356 100644 --- a/test-suite/arithmetic/diveucl.v +++ b/test-suite/primitive/uint63/diveucl.v diff --git a/test-suite/arithmetic/diveucl_21.v b/test-suite/primitive/uint63/diveucl_21.v index b12dba429c..b12dba429c 100644 --- a/test-suite/arithmetic/diveucl_21.v +++ b/test-suite/primitive/uint63/diveucl_21.v diff --git a/test-suite/arithmetic/eqb.v b/test-suite/primitive/uint63/eqb.v index dcc0b71f6d..dcc0b71f6d 100644 --- a/test-suite/arithmetic/eqb.v +++ b/test-suite/primitive/uint63/eqb.v diff --git a/test-suite/arithmetic/head0.v b/test-suite/primitive/uint63/head0.v index f4234d2605..f4234d2605 100644 --- a/test-suite/arithmetic/head0.v +++ b/test-suite/primitive/uint63/head0.v diff --git a/test-suite/arithmetic/isint.v b/test-suite/primitive/uint63/isint.v index c215caa878..c215caa878 100644 --- a/test-suite/arithmetic/isint.v +++ b/test-suite/primitive/uint63/isint.v diff --git a/test-suite/arithmetic/land.v b/test-suite/primitive/uint63/land.v index 0ea6fd90b6..0ea6fd90b6 100644 --- a/test-suite/arithmetic/land.v +++ b/test-suite/primitive/uint63/land.v diff --git a/test-suite/arithmetic/leb.v b/test-suite/primitive/uint63/leb.v index 5354919978..5354919978 100644 --- a/test-suite/arithmetic/leb.v +++ b/test-suite/primitive/uint63/leb.v diff --git a/test-suite/arithmetic/lor.v b/test-suite/primitive/uint63/lor.v index 9c3b85c054..9c3b85c054 100644 --- a/test-suite/arithmetic/lor.v +++ b/test-suite/primitive/uint63/lor.v diff --git a/test-suite/arithmetic/lsl.v b/test-suite/primitive/uint63/lsl.v index 70f3b90140..70f3b90140 100644 --- a/test-suite/arithmetic/lsl.v +++ b/test-suite/primitive/uint63/lsl.v diff --git a/test-suite/arithmetic/lsr.v b/test-suite/primitive/uint63/lsr.v index c36c24e237..c36c24e237 100644 --- a/test-suite/arithmetic/lsr.v +++ b/test-suite/primitive/uint63/lsr.v diff --git a/test-suite/arithmetic/ltb.v b/test-suite/primitive/uint63/ltb.v index 7ae5ac6493..7ae5ac6493 100644 --- a/test-suite/arithmetic/ltb.v +++ b/test-suite/primitive/uint63/ltb.v diff --git a/test-suite/arithmetic/lxor.v b/test-suite/primitive/uint63/lxor.v index b453fc7697..b453fc7697 100644 --- a/test-suite/arithmetic/lxor.v +++ b/test-suite/primitive/uint63/lxor.v diff --git a/test-suite/arithmetic/mod.v b/test-suite/primitive/uint63/mod.v index 5307eed493..5307eed493 100644 --- a/test-suite/arithmetic/mod.v +++ b/test-suite/primitive/uint63/mod.v diff --git a/test-suite/arithmetic/mul.v b/test-suite/primitive/uint63/mul.v index 9480e8cd46..9480e8cd46 100644 --- a/test-suite/arithmetic/mul.v +++ b/test-suite/primitive/uint63/mul.v diff --git a/test-suite/arithmetic/mulc.v b/test-suite/primitive/uint63/mulc.v index e10855bafa..e10855bafa 100644 --- a/test-suite/arithmetic/mulc.v +++ b/test-suite/primitive/uint63/mulc.v diff --git a/test-suite/arithmetic/reduction.v b/test-suite/primitive/uint63/reduction.v index 00e067ac5a..00e067ac5a 100644 --- a/test-suite/arithmetic/reduction.v +++ b/test-suite/primitive/uint63/reduction.v diff --git a/test-suite/arithmetic/sub.v b/test-suite/primitive/uint63/sub.v index 1606fd2aa1..1606fd2aa1 100644 --- a/test-suite/arithmetic/sub.v +++ b/test-suite/primitive/uint63/sub.v diff --git a/test-suite/arithmetic/subc.v b/test-suite/primitive/uint63/subc.v index fc4067e48b..fc4067e48b 100644 --- a/test-suite/arithmetic/subc.v +++ b/test-suite/primitive/uint63/subc.v diff --git a/test-suite/arithmetic/subcarryc.v b/test-suite/primitive/uint63/subcarryc.v index e81b6536b2..e81b6536b2 100644 --- a/test-suite/arithmetic/subcarryc.v +++ b/test-suite/primitive/uint63/subcarryc.v diff --git a/test-suite/arithmetic/tail0.v b/test-suite/primitive/uint63/tail0.v index c9d426087a..c9d426087a 100644 --- a/test-suite/arithmetic/tail0.v +++ b/test-suite/primitive/uint63/tail0.v diff --git a/test-suite/arithmetic/unsigned.v b/test-suite/primitive/uint63/unsigned.v index 82920bd201..82920bd201 100644 --- a/test-suite/arithmetic/unsigned.v +++ b/test-suite/primitive/uint63/unsigned.v diff --git a/theories/Floats/FloatAxioms.v b/theories/Floats/FloatAxioms.v new file mode 100644 index 0000000000..8ca64aac42 --- /dev/null +++ b/theories/Floats/FloatAxioms.v @@ -0,0 +1,58 @@ +Require Import ZArith Int63 SpecFloat PrimFloat FloatOps. + +(** * Properties of the primitive operators for the Binary64 format *) + +Notation valid_binary := (valid_binary prec emax). + +Definition SF64classify := SFclassify prec. +Definition SF64mul := SFmul prec emax. +Definition SF64add := SFadd prec emax. +Definition SF64sub := SFsub prec emax. +Definition SF64div := SFdiv prec emax. +Definition SF64sqrt := SFsqrt prec emax. +Definition SF64succ := SFsucc prec emax. +Definition SF64pred := SFpred prec emax. + +Axiom Prim2SF_valid : forall x, valid_binary (Prim2SF x) = true. +Axiom SF2Prim_Prim2SF : forall x, SF2Prim (Prim2SF x) = x. +Axiom Prim2SF_SF2Prim : forall x, valid_binary x = true -> Prim2SF (SF2Prim x) = x. + +Theorem Prim2SF_inj : forall x y, Prim2SF x = Prim2SF y -> x = y. + intros. rewrite <- SF2Prim_Prim2SF. symmetry. rewrite <- SF2Prim_Prim2SF. now rewrite H. +Qed. + +Theorem SF2Prim_inj : forall x y, SF2Prim x = SF2Prim y -> valid_binary x = true -> valid_binary y = true -> x = y. + intros. rewrite <- Prim2SF_SF2Prim by assumption. symmetry. rewrite <- Prim2SF_SF2Prim by assumption. rewrite H. reflexivity. +Qed. + +Axiom opp_spec : forall x, Prim2SF (-x)%float = SFopp (Prim2SF x). +Axiom abs_spec : forall x, Prim2SF (abs x) = SFabs (Prim2SF x). + +Axiom eqb_spec : forall x y, (x == y)%float = SFeqb (Prim2SF x) (Prim2SF y). +Axiom ltb_spec : forall x y, (x < y)%float = SFltb (Prim2SF x) (Prim2SF y). +Axiom leb_spec : forall x y, (x <= y)%float = SFleb (Prim2SF x) (Prim2SF y). + +Definition flatten_cmp_opt c := + match c with + | None => FNotComparable + | Some Eq => FEq + | Some Lt => FLt + | Some Gt => FGt + end. +Axiom compare_spec : forall x y, (x ?= y)%float = flatten_cmp_opt (SFcompare (Prim2SF x) (Prim2SF y)). + +Axiom classify_spec : forall x, classify x = SF64classify (Prim2SF x). +Axiom mul_spec : forall x y, Prim2SF (x * y)%float = SF64mul (Prim2SF x) (Prim2SF y). +Axiom add_spec : forall x y, Prim2SF (x + y)%float = SF64add (Prim2SF x) (Prim2SF y). +Axiom sub_spec : forall x y, Prim2SF (x - y)%float = SF64sub (Prim2SF x) (Prim2SF y). +Axiom div_spec : forall x y, Prim2SF (x / y)%float = SF64div (Prim2SF x) (Prim2SF y). +Axiom sqrt_spec : forall x, Prim2SF (sqrt x) = SF64sqrt (Prim2SF x). + +Axiom of_int63_spec : forall n, Prim2SF (of_int63 n) = binary_normalize prec emax (to_Z n) 0%Z false. +Axiom normfr_mantissa_spec : forall f, to_Z (normfr_mantissa f) = Z.of_N (SFnormfr_mantissa prec (Prim2SF f)). + +Axiom frshiftexp_spec : forall f, let (m,e) := frshiftexp f in (Prim2SF m, ((to_Z e) - shift)%Z) = SFfrexp prec emax (Prim2SF f). +Axiom ldshiftexp_spec : forall f e, Prim2SF (ldshiftexp f e) = SFldexp prec emax (Prim2SF f) ((to_Z e) - shift). + +Axiom next_up_spec : forall x, Prim2SF (next_up x) = SF64succ (Prim2SF x). +Axiom next_down_spec : forall x, Prim2SF (next_down x) = SF64pred (Prim2SF x). diff --git a/theories/Floats/FloatClass.v b/theories/Floats/FloatClass.v new file mode 100644 index 0000000000..627cb648f9 --- /dev/null +++ b/theories/Floats/FloatClass.v @@ -0,0 +1,2 @@ +Variant float_class : Set := + | PNormal | NNormal | PSubn | NSubn | PZero | NZero | PInf | NInf | NaN. diff --git a/theories/Floats/FloatLemmas.v b/theories/Floats/FloatLemmas.v new file mode 100644 index 0000000000..81cb7120e0 --- /dev/null +++ b/theories/Floats/FloatLemmas.v @@ -0,0 +1,319 @@ +Require Import ZArith Int63 SpecFloat PrimFloat FloatOps FloatAxioms. +Require Import Psatz. + +(** * Support results involving frexp and ldexp *) + +Lemma shift_value : shift = (2*emax + prec)%Z. + reflexivity. +Qed. + +Theorem frexp_spec : forall f, let (m,e) := frexp f in (Prim2SF m, e) = SFfrexp prec emax (Prim2SF f). + intro. + unfold frexp. + case_eq (frshiftexp f). + intros. + assert (H' := frshiftexp_spec f). + now rewrite H in H'. +Qed. + +Theorem ldexp_spec : forall f e, Prim2SF (ldexp f e) = SFldexp prec emax (Prim2SF f) e. + intros. + unfold ldexp. + rewrite (ldshiftexp_spec f _). + assert (Hv := Prim2SF_valid f). + destruct (Prim2SF f); auto. + unfold SFldexp. + unfold binary_round. + assert (Hmod_elim : forall e, ([| of_Z (Z.max (Z.min e (emax - emin)) (emin - emax - 1) + shift)|]%int63 - shift = Z.max (Z.min e (emax - emin)) (emin - emax - 1))%Z). + { + intro e1. + rewrite of_Z_spec, shift_value. + unfold wB, size; simpl. + unfold Z.pow_pos; simpl. + set (n := Z.max (Z.min _ _) _). + set (wB := 9223372036854775808%Z). (* Z.pow_pos 2 63 *) + assert (-2099 <= n <= 2098)%Z by (unfold n; lia). + rewrite Z.mod_small by (unfold wB; lia). + now rewrite Z.add_simpl_r. + } + rewrite Hmod_elim. + clear Hmod_elim. + revert Hv. + unfold valid_binary, bounded, canonical_mantissa. + unfold fexp. + rewrite Bool.andb_true_iff. + intro H'. + destruct H' as (H1,H2). + apply Zeq_bool_eq in H1. + apply Z.max_case_strong. + apply Z.min_case_strong. + - reflexivity. + - intros He _. + destruct (Z.max_spec (Z.pos (digits2_pos m) + e0 - prec) emin) as [ (H, Hm) | (H, Hm) ]. + + rewrite Hm in H1. + rewrite <- H1. + rewrite !Z.max_l by (revert He; unfold emax, emin, prec; lia). + replace (emin + _)%Z with emax by ring. + unfold shl_align. + rewrite <- H1 in H. + replace (Z.pos _ + _ - _ - _)%Z with (Z.pos (digits2_pos m) - prec)%Z by ring. + remember (Zpos _ - _)%Z as z'. + destruct z' ; [ lia | lia | ]. + unfold binary_round_aux. + unfold shr_fexp. + unfold fexp. + unfold Zdigits2. + unfold shr_record_of_loc, shr. + rewrite !Z.max_l by (revert H He; unfold emax, emin, prec; lia). + replace (_ - _)%Z with (Z.pos (digits2_pos (shift_pos p m)) - prec)%Z by ring. + assert (Hs : (Z.pos (digits2_pos (shift_pos p m)) <= prec)%Z). + { + assert (H' : forall p p', digits2_pos (shift_pos p p') = (digits2_pos p' + p)%positive). + { + induction p0. + intro p'. + simpl. + rewrite IHp0. + rewrite IHp0. + lia. + intro p'. + simpl. + rewrite IHp0. + rewrite IHp0. + lia. + intro p'. + simpl. + lia. + } + rewrite H'. + lia. + } + replace (Z.pos (digits2_pos m) + (emin + e) - prec - (emin + e))%Z with (Z.neg p) by lia. + unfold shr_m, loc_of_shr_record. + unfold round_nearest_even. + remember (Z.pos (digits2_pos (shift_pos p m)) - prec)%Z as ds. + destruct ds. + * rewrite Z.max_l by (revert He; unfold emax, emin, prec; lia). + replace (_ - _)%Z with Z0 by lia. + replace (_ <=? _)%Z with false by (symmetry; rewrite Z.leb_gt; lia). + rewrite Z.max_l by (revert He; unfold emax, emin, prec; lia). + replace (_ - _)%Z with Z0 by lia. + rewrite Z.max_l by (revert He; unfold emax, emin, prec; lia). + replace (_ - _)%Z with Z0 by lia. + replace (_ <=? _)%Z with false by (symmetry; rewrite Z.leb_gt; lia). + reflexivity. + * exfalso; lia. + * rewrite Z.max_l by (revert He; unfold emax, emin, prec; lia). + replace (_ - _)%Z with (Zneg p0) by lia. + replace (_ <=? _)%Z with false by (symmetry; rewrite Z.leb_gt; lia). + rewrite Z.max_l by (revert He; unfold emax, emin, prec; lia). + replace (_ - _)%Z with (Zneg p0) by lia. + rewrite Z.max_l by (revert He; unfold emax, emin, prec; lia). + replace (_ - _)%Z with (Zneg p0) by lia. + replace (_ <=? _)%Z with false by (symmetry; rewrite Z.leb_gt; lia). + reflexivity. + + rewrite !Z.max_l by (revert H He; unfold emax, emin, prec; lia). + rewrite Hm in H1. + clear Hm. + replace (Zpos _ + _ - _)%Z with (e0 + (emax - emin))%Z by (rewrite <- H1 at 1; ring). + replace (Zpos _ + _ - _)%Z with (e0 + e)%Z by (rewrite <- H1 at 1; ring). + unfold shl_align. + replace (_ - _)%Z with Z0 by ring. + replace (e0 + e - _)%Z with Z0 by ring. + unfold binary_round_aux. + unfold shr_fexp. + unfold fexp. + unfold Zdigits2. + rewrite !Z.max_l by (revert H He; unfold emax, emin, prec; lia). + unfold shr_record_of_loc. + unfold shr. + unfold Zdigits2. + replace (Zpos _ + _ - _ - _)%Z with Z0 by lia. + unfold shr_m. + unfold loc_of_shr_record. + unfold round_nearest_even. + rewrite Z.max_l by (revert H He; unfold emax, emin, prec; lia). + replace (Zpos _ + _ - _ - _)%Z with Z0 by lia. + replace (_ <=? _)%Z with false by (symmetry; rewrite Z.leb_gt; lia). + replace (Zpos _ + _ - _ - _)%Z with Z0 by lia. + rewrite Z.max_l by (revert H He; unfold emax, emin, prec; lia). + replace (Zpos _ + _ - _ - _)%Z with Z0 by lia. + replace (_ <=? _)%Z with false by (symmetry; rewrite Z.leb_gt; lia). + reflexivity. + - rewrite Z.min_le_iff. + intro H. + destruct H as [ He | Habs ]; [ | revert Habs; now unfold emin, emax ]. + unfold shl_align. + assert (Hprec : (Z.pos (digits2_pos m) <= prec)%Z). + { + destruct (Z.max_spec (Z.pos (digits2_pos m) + e0 - prec) emin) as [ (Hpi, Hpe) | (Hpi, Hpe) ]; rewrite Hpe in H1; lia. + } + + assert (Hshr : forall p s, Zdigits2 (shr_m (iter_pos shr_1 p s)) = Z.max Z0 (Zdigits2 (shr_m s) - Z.pos p)%Z). + { + assert (Hshr1 : forall s, Zdigits2 (shr_m (shr_1 s)) = Z.max 0 (Zdigits2 (shr_m s) - 1)%Z). + { + intro s0. + destruct s0. + unfold shr_1. + destruct shr_m; try (simpl; lia). + - destruct p; unfold Zdigits2, shr_m, digits2_pos; lia. + - destruct p; unfold Zdigits2, shr_m, digits2_pos; lia. + } + induction p. + simpl. + intro s0. + do 2 rewrite IHp. + rewrite Hshr1. + lia. + intros. + simpl. + do 2 rewrite IHp. + lia. + apply Hshr1. + } + + assert (Hd0 : forall z, Zdigits2 z = 0%Z -> z = 0%Z). + { + intro z. + unfold Zdigits2. + now destruct z. + } + + assert (Hshr_p0 : forall p0, (prec < Z.pos p0)%Z -> shr_m (iter_pos shr_1 p0 {| shr_m := Z.pos m; shr_r := false; shr_s := false |}) = Z0). + { + intros p0 Hp0. + apply Hd0. + rewrite Hshr. + rewrite Z.max_l; [ reflexivity | ]. + unfold shr_m. + unfold Zdigits2. + lia. + } + + assert (Hshr_p0_r : forall p0, (prec < Z.pos p0)%Z -> shr_r (iter_pos shr_1 p0 {| shr_m := Z.pos m; shr_r := false; shr_s := false |}) = false). + { + intros p0 Hp0. + + assert (Hshr_p0m1 : shr_m (iter_pos shr_1 (p0-1) {| shr_m := Z.pos m; shr_r := false; shr_s := false |}) = Z0). + { + apply Hd0. + rewrite Hshr. + rewrite Z.max_l; [ reflexivity | ]. + unfold shr_m. + unfold Zdigits2. + lia. + } + + assert (Hiter_pos : forall A (f : A -> A) p e, iter_pos f (p + 1) e = f (iter_pos f p e)). + { + assert (Hiter_pos' : forall A (f : A -> A) p e, iter_pos f p (f e) = f (iter_pos f p e)). + { + intros A f'. + induction p. + intro e'. + simpl. + now do 2 rewrite IHp. + intro e'. + simpl. + now do 2 rewrite IHp. + intro e'. + now simpl. + } + intros A f'. + induction p. + intros. + simpl. + rewrite <- Pos.add_1_r. + do 2 rewrite IHp. + now do 3 rewrite Hiter_pos'. + intros. + simpl. + now do 2 rewrite Hiter_pos'. + intros. + now simpl. + } + replace p0 with (p0 - 1 + 1)%positive. + rewrite Hiter_pos. + unfold shr_1 at 1. + remember (iter_pos _ _ _) as shr_p0m1. + destruct shr_p0m1. + unfold SpecFloat.shr_m in Hshr_p0m1. + now rewrite Hshr_p0m1. + rewrite Pos.add_1_r. + rewrite Pos.sub_1_r. + apply Pos.succ_pred. + lia. + } + + rewrite Z.leb_le in H2. + + destruct (Z.max_spec (Z.pos (digits2_pos m) + (e0 + (emin - emax - 1)) - prec) emin) as [ (H, Hm) | (H, Hm) ]. + + rewrite Hm. + replace (_ - _)%Z with (emax - e0 + 1)%Z by ring. + remember (emax - e0 + 1)%Z as z'. + destruct z'; [ exfalso; lia | | exfalso; lia ]. + unfold binary_round_aux. + unfold shr_fexp, fexp. + unfold shr, shr_record_of_loc. + unfold Zdigits2. + rewrite Hm. + replace (_ - _)%Z with (Z.pos p) by (rewrite Heqz'; ring). + set (rne := round_nearest_even _ _). + assert (rne = 0%Z). + { + unfold rne. + unfold round_nearest_even. + + assert (Hp0 : (prec < Z.pos p)%Z) by lia. + + unfold loc_of_shr_record. + specialize (Hshr_p0_r _ Hp0). + specialize (Hshr_p0 _ Hp0). + revert Hshr_p0_r Hshr_p0. + set (shr_p0 := iter_pos shr_1 _ _). + destruct shr_p0. + unfold SpecFloat.shr_r, SpecFloat.shr_m. + intros Hshr_r Hshr_m. + rewrite Hshr_r, Hshr_m. + now destruct shr_s. + } + + rewrite H0. + rewrite Z.max_r by (rewrite Heqz'; unfold prec; lia). + replace (_ - _)%Z with 0%Z by lia. + unfold shr_m. + + rewrite Z.max_r by lia. + remember (emin - (e0 + e))%Z as eminmze. + destruct eminmze; [ exfalso; lia | | exfalso; lia ]. + + rewrite Z.max_r by lia. + rewrite <- Heqeminmze. + + set (rne' := round_nearest_even _ _). + assert (Hrne'0 : rne' = 0%Z). + { + unfold rne'. + unfold round_nearest_even. + + assert (Hp1 : (prec < Z.pos p0)%Z) by lia. + + unfold loc_of_shr_record. + specialize (Hshr_p0_r _ Hp1). + specialize (Hshr_p0 _ Hp1). + revert Hshr_p0_r Hshr_p0. + set (shr_p1 := iter_pos shr_1 _ _). + destruct shr_p1. + unfold SpecFloat.shr_r, SpecFloat.shr_m. + intros Hshr_r Hshr_m. + rewrite Hshr_r, Hshr_m. + now destruct shr_s. + } + + rewrite Hrne'0. + rewrite Z.max_r by (rewrite Heqeminmze; unfold prec; lia). + replace (_ - _)%Z with 0%Z by lia. + reflexivity. + + exfalso; lia. +Qed. diff --git a/theories/Floats/FloatOps.v b/theories/Floats/FloatOps.v new file mode 100644 index 0000000000..f0d3bcced9 --- /dev/null +++ b/theories/Floats/FloatOps.v @@ -0,0 +1,48 @@ +Require Import ZArith Int63 SpecFloat PrimFloat. + +(** * Derived operations and mapping between primitive [float]s and [spec_float]s *) + +Definition prec := 53%Z. +Definition emax := 1024%Z. +Notation emin := (emin prec emax). + +Definition shift := 2101%Z. (** [= 2*emax + prec] *) + +Definition frexp f := + let (m, se) := frshiftexp f in + (m, ([| se |] - shift)%Z%int63). + +Definition ldexp f e := + let e' := Z.max (Z.min e (emax - emin)) (emin - emax - 1) in + ldshiftexp f (of_Z (e' + shift)). + +Definition ulp f := ldexp one (fexp prec emax (snd (frexp f))). + +(** [Prim2SF] is an injective function that will be useful to express +the properties of the implemented Binary64 format (see [FloatAxioms]). +*) +Definition Prim2SF f := + if is_nan f then S754_nan + else if is_zero f then S754_zero (get_sign f) + else if is_infinity f then S754_infinity (get_sign f) + else + let (r, exp) := frexp f in + let e := (exp - prec)%Z in + let (shr, e') := shr_fexp prec emax [| normfr_mantissa r |]%int63 e loc_Exact in + match shr_m shr with + | Zpos p => S754_finite (get_sign f) p e' + | Zneg _ | Z0 => S754_zero false (* must never occur *) + end. + +Definition SF2Prim ef := + match ef with + | S754_nan => nan + | S754_zero false => zero + | S754_zero true => neg_zero + | S754_infinity false => infinity + | S754_infinity true => neg_infinity + | S754_finite s m e => + let pm := of_int63 (of_Z (Zpos m)) in + let f := ldexp pm e in + if s then (-f)%float else f + end. diff --git a/theories/Floats/Floats.v b/theories/Floats/Floats.v new file mode 100644 index 0000000000..700c69b99d --- /dev/null +++ b/theories/Floats/Floats.v @@ -0,0 +1,17 @@ +(** The Floats library is split in 6 theories: +- FloatClass: define the [float_class] inductive +- PrimFloat: define the floating-point values and operators as kernel primitives +- SpecFloat: specify the floating-point operators with binary integers +- FloatOps: define conversion functions between [spec_float] and [float] +- FloatAxioms: state properties of the primitive operators w.r.t. [spec_float] +- FloatLemmas: prove a few results involving frexp and ldexp + +For a brief overview of the Floats library, +see {{https://coq.inria.fr/distrib/current/refman/language/coq-library.html#floats-library}} *) + +Require Export FloatClass. +Require Export PrimFloat. +Require Export SpecFloat. +Require Export FloatOps. +Require Export FloatAxioms. +Require Export FloatLemmas. diff --git a/theories/Floats/PrimFloat.v b/theories/Floats/PrimFloat.v new file mode 100644 index 0000000000..bc1727469d --- /dev/null +++ b/theories/Floats/PrimFloat.v @@ -0,0 +1,118 @@ +Require Import Int63 FloatClass. + +(** * Definition of the interface for primitive floating-point arithmetic + +This interface provides processor operators for the Binary64 format of the +IEEE 754-2008 standard. *) + +(** ** Type definition for the co-domain of [compare] *) +Variant float_comparison : Set := FEq | FLt | FGt | FNotComparable. + +Register float_comparison as kernel.ind_f_cmp. + +Register float_class as kernel.ind_f_class. + +(** ** The main type *) +(** [float]: primitive type for Binary64 floating-point numbers. *) +Primitive float := #float64_type. + +(** ** Syntax support *) +Declare Scope float_scope. +Delimit Scope float_scope with float. +Bind Scope float_scope with float. + +Declare ML Module "float_syntax_plugin". + +(** ** Floating-point operators *) +Primitive classify := #float64_classify. + +Primitive abs := #float64_abs. + +Primitive sqrt := #float64_sqrt. + +Primitive opp := #float64_opp. +Notation "- x" := (opp x) : float_scope. + +Primitive eqb := #float64_eq. +Notation "x == y" := (eqb x y) (at level 70, no associativity) : float_scope. + +Primitive ltb := #float64_lt. +Notation "x < y" := (ltb x y) (at level 70, no associativity) : float_scope. + +Primitive leb := #float64_le. +Notation "x <= y" := (leb x y) (at level 70, no associativity) : float_scope. + +Primitive compare := #float64_compare. +Notation "x ?= y" := (compare x y) (at level 70, no associativity) : float_scope. + +Primitive mul := #float64_mul. +Notation "x * y" := (mul x y) : float_scope. + +Primitive add := #float64_add. +Notation "x + y" := (add x y) : float_scope. + +Primitive sub := #float64_sub. +Notation "x - y" := (sub x y) : float_scope. + +Primitive div := #float64_div. +Notation "x / y" := (div x y) : float_scope. + +(** ** Conversions *) + +(** [of_int63]: convert a primitive integer into a float value. + The value is rounded if need be. *) +Primitive of_int63 := #float64_of_int63. + +(** Specification of [normfr_mantissa]: +- If the input is a float value with an absolute value inside $[0.5, 1.)$#[0.5, 1.)#; +- Then return its mantissa as a primitive integer. + The mantissa will be a 53-bit integer with its most significant bit set to 1; +- Else return zero. + +The sign bit is always ignored. *) +Primitive normfr_mantissa := #float64_normfr_mantissa. + +(** ** Exponent manipulation functions *) +(** [frshiftexp]: convert a float to fractional part in $[0.5, 1.)$#[0.5, 1.)# +and integer part. *) +Primitive frshiftexp := #float64_frshiftexp. + +(** [ldshiftexp]: multiply a float by an integral power of 2. *) +Primitive ldshiftexp := #float64_ldshiftexp. + +(** ** Predecesor/Successor functions *) + +(** [next_up]: return the next float towards positive infinity. *) +Primitive next_up := #float64_next_up. + +(** [next_down]: return the next float towards negative infinity. *) +Primitive next_down := #float64_next_down. + +(** ** Special values (needed for pretty-printing) *) +Definition infinity := Eval compute in div (of_int63 1) (of_int63 0). +Definition neg_infinity := Eval compute in opp infinity. +Definition nan := Eval compute in div (of_int63 0) (of_int63 0). + +Register infinity as num.float.infinity. +Register neg_infinity as num.float.neg_infinity. +Register nan as num.float.nan. + +(** ** Other special values *) +Definition one := Eval compute in (of_int63 1). +Definition zero := Eval compute in (of_int63 0). +Definition neg_zero := Eval compute in (-zero)%float. +Definition two := Eval compute in (of_int63 2). + +(** ** Predicates and helper functions *) +Definition is_nan f := negb (f == f)%float. + +Definition is_zero f := (f == zero)%float. (* note: 0 == -0 with floats *) + +Definition is_infinity f := (abs f == infinity)%float. + +Definition is_finite (x : float) := negb (is_nan x || is_infinity x). + +(** [get_sign]: return [true] for [-] sign, [false] for [+] sign. *) +Definition get_sign f := + let f := if is_zero f then (one / f)%float else f in + (f < zero)%float. diff --git a/theories/Floats/SpecFloat.v b/theories/Floats/SpecFloat.v new file mode 100644 index 0000000000..fd0aa5e075 --- /dev/null +++ b/theories/Floats/SpecFloat.v @@ -0,0 +1,416 @@ +Require Import ZArith FloatClass. + +(** * Specification of floating-point arithmetic + +This specification is mostly borrowed from the [IEEE754.Binary] module +of the Flocq library (see {{http://flocq.gforge.inria.fr/}}) *) + +(** ** Inductive specification of floating-point numbers + +Similar to [Flocq.IEEE754.Binary.full_float], but with no NaN payload. *) +Variant spec_float := + | S754_zero (s : bool) + | S754_infinity (s : bool) + | S754_nan + | S754_finite (s : bool) (m : positive) (e : Z). + +(** ** Parameterized definitions + +[prec] is the number of bits of the mantissa including the implicit one; +[emax] is the exponent of the infinities. + +For instance, Binary64 is defined by [prec = 53] and [emax = 1024]. *) +Section FloatOps. + Variable prec emax : Z. + + Definition emin := (3-emax-prec)%Z. + Definition fexp e := Z.max (e - prec) emin. + + Section Zdigits2. + Fixpoint digits2_pos (n : positive) : positive := + match n with + | xH => xH + | xO p => Pos.succ (digits2_pos p) + | xI p => Pos.succ (digits2_pos p) + end. + + Definition Zdigits2 n := + match n with + | Z0 => n + | Zpos p => Zpos (digits2_pos p) + | Zneg p => Zpos (digits2_pos p) + end. + End Zdigits2. + + Section ValidBinary. + Definition canonical_mantissa m e := + Zeq_bool (fexp (Zpos (digits2_pos m) + e)) e. + + Definition bounded m e := + andb (canonical_mantissa m e) (Zle_bool e (emax - prec)). + + Definition valid_binary x := + match x with + | S754_finite _ m e => bounded m e + | _ => true + end. + End ValidBinary. + + Section Iter. + Context {A : Type}. + Variable (f : A -> A). + + Fixpoint iter_pos (n : positive) (x : A) {struct n} : A := + match n with + | xI n' => iter_pos n' (iter_pos n' (f x)) + | xO n' => iter_pos n' (iter_pos n' x) + | xH => f x + end. + End Iter. + + Section Rounding. + Inductive location := loc_Exact | loc_Inexact : comparison -> location. + + Record shr_record := { shr_m : Z ; shr_r : bool ; shr_s : bool }. + + Definition shr_1 mrs := + let '(Build_shr_record m r s) := mrs in + let s := orb r s in + match m with + | Z0 => Build_shr_record Z0 false s + | Zpos xH => Build_shr_record Z0 true s + | Zpos (xO p) => Build_shr_record (Zpos p) false s + | Zpos (xI p) => Build_shr_record (Zpos p) true s + | Zneg xH => Build_shr_record Z0 true s + | Zneg (xO p) => Build_shr_record (Zneg p) false s + | Zneg (xI p) => Build_shr_record (Zneg p) true s + end. + + Definition loc_of_shr_record mrs := + match mrs with + | Build_shr_record _ false false => loc_Exact + | Build_shr_record _ false true => loc_Inexact Lt + | Build_shr_record _ true false => loc_Inexact Eq + | Build_shr_record _ true true => loc_Inexact Gt + end. + + Definition shr_record_of_loc m l := + match l with + | loc_Exact => Build_shr_record m false false + | loc_Inexact Lt => Build_shr_record m false true + | loc_Inexact Eq => Build_shr_record m true false + | loc_Inexact Gt => Build_shr_record m true true + end. + + Definition shr mrs e n := + match n with + | Zpos p => (iter_pos shr_1 p mrs, (e + n)%Z) + | _ => (mrs, e) + end. + + Definition shr_fexp m e l := + shr (shr_record_of_loc m l) e (fexp (Zdigits2 m + e) - e). + + Definition round_nearest_even mx lx := + match lx with + | loc_Exact => mx + | loc_Inexact Lt => mx + | loc_Inexact Eq => if Z.even mx then mx else (mx + 1)%Z + | loc_Inexact Gt => (mx + 1)%Z + end. + + Definition binary_round_aux sx mx ex lx := + let '(mrs', e') := shr_fexp mx ex lx in + let '(mrs'', e'') := shr_fexp (round_nearest_even (shr_m mrs') (loc_of_shr_record mrs')) e' loc_Exact in + match shr_m mrs'' with + | Z0 => S754_zero sx + | Zpos m => if Zle_bool e'' (emax - prec) then S754_finite sx m e'' else S754_infinity sx + | _ => S754_nan + end. + + Definition shl_align mx ex ex' := + match (ex' - ex)%Z with + | Zneg d => (shift_pos d mx, ex') + | _ => (mx, ex) + end. + + Definition binary_round sx mx ex := + let '(mz, ez) := shl_align mx ex (fexp (Zpos (digits2_pos mx) + ex))in + binary_round_aux sx (Zpos mz) ez loc_Exact. + + Definition binary_normalize m e szero := + match m with + | Z0 => S754_zero szero + | Zpos m => binary_round false m e + | Zneg m => binary_round true m e + end. + End Rounding. + + (** ** Define operations *) + + Definition SFopp x := + match x with + | S754_nan => S754_nan + | S754_infinity sx => S754_infinity (negb sx) + | S754_finite sx mx ex => S754_finite (negb sx) mx ex + | S754_zero sx => S754_zero (negb sx) + end. + + Definition SFabs x := + match x with + | S754_nan => S754_nan + | S754_infinity sx => S754_infinity false + | S754_finite sx mx ex => S754_finite false mx ex + | S754_zero sx => S754_zero false + end. + + Definition SFcompare f1 f2 := + match f1, f2 with + | S754_nan , _ | _, S754_nan => None + | S754_infinity s1, S754_infinity s2 => + Some match s1, s2 with + | true, true => Eq + | false, false => Eq + | true, false => Lt + | false, true => Gt + end + | S754_infinity s, _ => Some (if s then Lt else Gt) + | _, S754_infinity s => Some (if s then Gt else Lt) + | S754_finite s _ _, S754_zero _ => Some (if s then Lt else Gt) + | S754_zero _, S754_finite s _ _ => Some (if s then Gt else Lt) + | S754_zero _, S754_zero _ => Some Eq + | S754_finite s1 m1 e1, S754_finite s2 m2 e2 => + Some match s1, s2 with + | true, false => Lt + | false, true => Gt + | false, false => + match Z.compare e1 e2 with + | Lt => Lt + | Gt => Gt + | Eq => Pcompare m1 m2 Eq + end + | true, true => + match Z.compare e1 e2 with + | Lt => Gt + | Gt => Lt + | Eq => CompOpp (Pcompare m1 m2 Eq) + end + end + end. + + Definition SFeqb f1 f2 := + match SFcompare f1 f2 with + | Some Eq => true + | _ => false + end. + + Definition SFltb f1 f2 := + match SFcompare f1 f2 with + | Some Lt => true + | _ => false + end. + + Definition SFleb f1 f2 := + match SFcompare f1 f2 with + | Some Le => true + | _ => false + end. + + Definition SFclassify f := + match f with + | S754_nan => NaN + | S754_infinity false => PInf + | S754_infinity true => NInf + | S754_zero false => NZero + | S754_zero true => PZero + | S754_finite false m _ => + if (digits2_pos m =? Z.to_pos prec)%positive then PNormal + else PSubn + | S754_finite true m _ => + if (digits2_pos m =? Z.to_pos prec)%positive then NNormal + else NSubn + end. + + Definition SFmul x y := + match x, y with + | S754_nan, _ | _, S754_nan => S754_nan + | S754_infinity sx, S754_infinity sy => S754_infinity (xorb sx sy) + | S754_infinity sx, S754_finite sy _ _ => S754_infinity (xorb sx sy) + | S754_finite sx _ _, S754_infinity sy => S754_infinity (xorb sx sy) + | S754_infinity _, S754_zero _ => S754_nan + | S754_zero _, S754_infinity _ => S754_nan + | S754_finite sx _ _, S754_zero sy => S754_zero (xorb sx sy) + | S754_zero sx, S754_finite sy _ _ => S754_zero (xorb sx sy) + | S754_zero sx, S754_zero sy => S754_zero (xorb sx sy) + | S754_finite sx mx ex, S754_finite sy my ey => + binary_round_aux (xorb sx sy) (Zpos (mx * my)) (ex + ey) loc_Exact + end. + + Definition cond_Zopp (b : bool) m := if b then Z.opp m else m. + + Definition SFadd x y := + match x, y with + | S754_nan, _ | _, S754_nan => S754_nan + | S754_infinity sx, S754_infinity sy => + if Bool.eqb sx sy then x else S754_nan + | S754_infinity _, _ => x + | _, S754_infinity _ => y + | S754_zero sx, S754_zero sy => + if Bool.eqb sx sy then x else + S754_zero false + | S754_zero _, _ => y + | _, S754_zero _ => x + | S754_finite sx mx ex, S754_finite sy my ey => + let ez := Z.min ex ey in + binary_normalize (Zplus (cond_Zopp sx (Zpos (fst (shl_align mx ex ez)))) (cond_Zopp sy (Zpos (fst (shl_align my ey ez))))) + ez false + end. + + Definition SFsub x y := + match x, y with + | S754_nan, _ | _, S754_nan => S754_nan + | S754_infinity sx, S754_infinity sy => + if Bool.eqb sx (negb sy) then x else S754_nan + | S754_infinity _, _ => x + | _, S754_infinity sy => S754_infinity (negb sy) + | S754_zero sx, S754_zero sy => + if Bool.eqb sx (negb sy) then x else + S754_zero false + | S754_zero _, S754_finite sy my ey => S754_finite (negb sy) my ey + | _, S754_zero _ => x + | S754_finite sx mx ex, S754_finite sy my ey => + let ez := Z.min ex ey in + binary_normalize (Zminus (cond_Zopp sx (Zpos (fst (shl_align mx ex ez)))) (cond_Zopp sy (Zpos (fst (shl_align my ey ez))))) + ez false + end. + + Definition new_location_even nb_steps k := + if Zeq_bool k 0 then loc_Exact + else loc_Inexact (Z.compare (2 * k) nb_steps). + + Definition new_location_odd nb_steps k := + if Zeq_bool k 0 then loc_Exact + else + loc_Inexact + match Z.compare (2 * k + 1) nb_steps with + | Lt => Lt + | Eq => Lt + | Gt => Gt + end. + + Definition new_location nb_steps := + if Z.even nb_steps then new_location_even nb_steps else new_location_odd nb_steps. + + Definition SFdiv_core_binary m1 e1 m2 e2 := + let d1 := Zdigits2 m1 in + let d2 := Zdigits2 m2 in + let e' := Z.min (fexp (d1 + e1 - (d2 + e2))) (e1 - e2) in + let s := (e1 - e2 - e')%Z in + let m' := + match s with + | Zpos _ => Z.shiftl m1 s + | Z0 => m1 + | Zneg _ => Z0 + end in + let '(q, r) := Z.div_eucl m' m2 in + (q, e', new_location m2 r). + + Definition SFdiv x y := + match x, y with + | S754_nan, _ | _, S754_nan => S754_nan + | S754_infinity sx, S754_infinity sy => S754_nan + | S754_infinity sx, S754_finite sy _ _ => S754_infinity (xorb sx sy) + | S754_finite sx _ _, S754_infinity sy => S754_zero (xorb sx sy) + | S754_infinity sx, S754_zero sy => S754_infinity (xorb sx sy) + | S754_zero sx, S754_infinity sy => S754_zero (xorb sx sy) + | S754_finite sx _ _, S754_zero sy => S754_infinity (xorb sx sy) + | S754_zero sx, S754_finite sy _ _ => S754_zero (xorb sx sy) + | S754_zero sx, S754_zero sy => S754_nan + | S754_finite sx mx ex, S754_finite sy my ey => + let '(mz, ez, lz) := SFdiv_core_binary (Zpos mx) ex (Zpos my) ey in + binary_round_aux (xorb sx sy) mz ez lz + end. + + Definition SFsqrt_core_binary m e := + let d := Zdigits2 m in + let e' := Z.min (fexp (Z.div2 (d + e + 1))) (Z.div2 e) in + let s := (e - 2 * e')%Z in + let m' := + match s with + | Zpos p => Z.shiftl m s + | Z0 => m + | Zneg _ => Z0 + end in + let (q, r) := Z.sqrtrem m' in + let l := + if Zeq_bool r 0 then loc_Exact + else loc_Inexact (if Zle_bool r q then Lt else Gt) in + (q, e', l). + + Definition SFsqrt x := + match x with + | S754_nan => S754_nan + | S754_infinity false => x + | S754_infinity true => S754_nan + | S754_finite true _ _ => S754_nan + | S754_zero _ => x + | S754_finite sx mx ex => + let '(mz, ez, lz) := SFsqrt_core_binary (Zpos mx) ex in + binary_round_aux false mz ez lz + end. + + Definition SFnormfr_mantissa f := + match f with + | S754_finite _ mx ex => + if Z.eqb ex (-prec) then Npos mx else 0%N + | _ => 0%N + end. + + Definition SFldexp f e := + match f with + | S754_finite sx mx ex => binary_round sx mx (ex+e) + | _ => f + end. + + Definition SFfrexp f := + match f with + | S754_finite sx mx ex => + if (Z.to_pos prec <=? digits2_pos mx)%positive then + (S754_finite sx mx (-prec), (ex+prec)%Z) + else + let d := (prec - Z.pos (digits2_pos mx))%Z in + (S754_finite sx (shift_pos (Z.to_pos d) mx) (-prec), (ex+prec-d)%Z) + | _ => (f, (-2*emax-prec)%Z) + end. + + Definition SFone := binary_round false 1 0. + + Definition SFulp x := SFldexp SFone (fexp (snd (SFfrexp x))). + + Definition SFpred_pos x := + match x with + | S754_finite _ mx _ => + let d := + if (mx~0 =? shift_pos (Z.to_pos prec) 1)%positive then + SFldexp SFone (fexp (snd (SFfrexp x) - 1)) + else + SFulp x in + SFsub x d + | _ => x + end. + + Definition SFmax_float := + S754_finite false (shift_pos (Z.to_pos prec) 1 - 1) (emax - prec). + + Definition SFsucc x := + match x with + | S754_zero _ => SFldexp SFone emin + | S754_infinity false => x + | S754_infinity true => SFopp SFmax_float + | S754_nan => x + | S754_finite false _ _ => SFadd x (SFulp x) + | S754_finite true _ _ => SFopp (SFpred_pos (SFopp x)) + end. + + Definition SFpred f := SFopp (SFsucc (SFopp f)). +End FloatOps. diff --git a/user-contrib/Ltac2/Constr.v b/user-contrib/Ltac2/Constr.v index 942cbe8916..a0b25afc37 100644 --- a/user-contrib/Ltac2/Constr.v +++ b/user-contrib/Ltac2/Constr.v @@ -45,6 +45,7 @@ Ltac2 Type kind := [ | CoFix (int, ident option binder_annot array, constr array, constr array) | Proj (projection, constr) | Uint63 (uint63) +| Float (float) ]. Ltac2 @ external kind : constr -> kind := "ltac2" "constr_kind". diff --git a/user-contrib/Ltac2/Init.v b/user-contrib/Ltac2/Init.v index 6eed261554..65f0a362b1 100644 --- a/user-contrib/Ltac2/Init.v +++ b/user-contrib/Ltac2/Init.v @@ -17,6 +17,7 @@ Ltac2 Type string. Ltac2 Type char. Ltac2 Type ident. Ltac2 Type uint63. +Ltac2 Type float. (** Constr-specific built-in types *) Ltac2 Type meta. diff --git a/user-contrib/Ltac2/tac2core.ml b/user-contrib/Ltac2/tac2core.ml index 0268e8f9ef..55cd7f7692 100644 --- a/user-contrib/Ltac2/tac2core.ml +++ b/user-contrib/Ltac2/tac2core.ml @@ -450,6 +450,8 @@ let () = define1 "constr_kind" constr begin fun c -> |] | Int n -> v_blk 17 [|Value.of_uint63 n|] + | Float f -> + v_blk 18 [|Value.of_float f|] end end @@ -530,6 +532,9 @@ let () = define1 "constr_make" valexpr begin fun knd -> | (17, [|n|]) -> let n = Value.to_uint63 n in EConstr.mkInt n + | (18, [|f|]) -> + let f = Value.to_float f in + EConstr.mkFloat f | _ -> assert false in return (Value.of_constr c) diff --git a/user-contrib/Ltac2/tac2ffi.ml b/user-contrib/Ltac2/tac2ffi.ml index 7c9613f31b..9ae17bf9bc 100644 --- a/user-contrib/Ltac2/tac2ffi.ml +++ b/user-contrib/Ltac2/tac2ffi.ml @@ -33,6 +33,8 @@ type valexpr = (** Arbitrary data *) | ValUint63 of Uint63.t (** Primitive integers *) +| ValFloat of Float64.t + (** Primitive floats *) and closure = MLTactic : (valexpr, 'v) arity0 * 'v -> closure @@ -50,21 +52,21 @@ type t = valexpr let is_int = function | ValInt _ -> true -| ValBlk _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ | ValUint63 _ -> false +| ValBlk _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ | ValUint63 _ | ValFloat _ -> false let tag v = match v with | ValBlk (n, _) -> n -| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ | ValUint63 _ -> +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ | ValUint63 _ | ValFloat _ -> CErrors.anomaly (Pp.str "Unexpected value shape") let field v n = match v with | ValBlk (_, v) -> v.(n) -| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ | ValUint63 _ -> +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ | ValUint63 _ | ValFloat _ -> CErrors.anomaly (Pp.str "Unexpected value shape") let set_field v n w = match v with | ValBlk (_, v) -> v.(n) <- w -| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ | ValUint63 _ -> +| ValInt _ | ValStr _ | ValCls _ | ValOpn _ | ValExt _ | ValUint63 _ | ValFloat _ -> CErrors.anomaly (Pp.str "Unexpected value shape") let make_block tag v = ValBlk (tag, v) @@ -196,7 +198,7 @@ let of_closure cls = ValCls cls let to_closure = function | ValCls cls -> cls -| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ | ValUint63 _ -> assert false +| ValExt _ | ValInt _ | ValBlk _ | ValStr _ | ValOpn _ | ValUint63 _ | ValFloat _ -> assert false let closure = { r_of = of_closure; @@ -333,6 +335,17 @@ let uint63 = { r_id = false; } +let of_float f = ValFloat f +let to_float = function +| ValFloat f -> f +| _ -> assert false + +let float = { + r_of = of_float; + r_to = to_float; + r_id = false; +} + let of_constant c = of_ext val_constant c let to_constant c = to_ext val_constant c let constant = repr_ext val_constant diff --git a/user-contrib/Ltac2/tac2ffi.mli b/user-contrib/Ltac2/tac2ffi.mli index d3c9596e8f..ee13f00568 100644 --- a/user-contrib/Ltac2/tac2ffi.mli +++ b/user-contrib/Ltac2/tac2ffi.mli @@ -32,6 +32,8 @@ type valexpr = (** Arbitrary data *) | ValUint63 of Uint63.t (** Primitive integers *) +| ValFloat of Float64.t + (** Primitive floats *) type 'a arity @@ -151,6 +153,10 @@ val of_uint63 : Uint63.t -> valexpr val to_uint63 : valexpr -> Uint63.t val uint63 : Uint63.t repr +val of_float : Float64.t -> valexpr +val to_float : valexpr -> Float64.t +val float : Float64.t repr + type ('a, 'b) fun1 val app_fun1 : ('a, 'b) fun1 -> 'a repr -> 'b repr -> 'a -> 'b Proofview.tactic diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index 98fe436a22..5822a1a586 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -248,6 +248,7 @@ let build_beq_scheme mode kn = | Meta _ -> raise (EqUnknown "meta-variable") | Evar _ -> raise (EqUnknown "existential variable") | Int _ -> raise (EqUnknown "int") + | Float _ -> raise (EqUnknown "float") in aux t in diff --git a/vernac/g_vernac.mlg b/vernac/g_vernac.mlg index 1387ca4675..b4c0a33585 100644 --- a/vernac/g_vernac.mlg +++ b/vernac/g_vernac.mlg @@ -244,7 +244,8 @@ GRAMMAR EXTEND Gram ; register_type_token: - [ [ "#int63_type" -> { CPrimitives.PT_int63 } ] ] + [ [ "#int63_type" -> { CPrimitives.PT_int63 } + | "#float64_type" -> { CPrimitives.PT_float64 } ] ] ; register_prim_token: @@ -272,6 +273,24 @@ GRAMMAR EXTEND Gram | "#int63_lt" -> { CPrimitives.Int63lt } | "#int63_le" -> { CPrimitives.Int63le } | "#int63_compare" -> { CPrimitives.Int63compare } + | "#float64_opp" -> { CPrimitives.Float64opp } + | "#float64_abs" -> { CPrimitives.Float64abs } + | "#float64_eq" -> { CPrimitives.Float64eq } + | "#float64_lt" -> { CPrimitives.Float64lt } + | "#float64_le" -> { CPrimitives.Float64le } + | "#float64_compare" -> { CPrimitives.Float64compare } + | "#float64_classify" -> { CPrimitives.Float64classify } + | "#float64_add" -> { CPrimitives.Float64add } + | "#float64_sub" -> { CPrimitives.Float64sub } + | "#float64_mul" -> { CPrimitives.Float64mul } + | "#float64_div" -> { CPrimitives.Float64div } + | "#float64_sqrt" -> { CPrimitives.Float64sqrt } + | "#float64_of_int63" -> { CPrimitives.Float64ofInt63 } + | "#float64_normfr_mantissa" -> { CPrimitives.Float64normfr_mantissa } + | "#float64_frshiftexp" -> { CPrimitives.Float64frshiftexp } + | "#float64_ldshiftexp" -> { CPrimitives.Float64ldshiftexp } + | "#float64_next_up" -> { CPrimitives.Float64next_up } + | "#float64_next_down" -> { CPrimitives.Float64next_down } ] ] ; diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index ade291918e..6dfba02ae9 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -1844,11 +1844,13 @@ let vernac_register qid r = if DirPath.equal (dirpath_of_string "kernel") ns then begin if Global.sections_are_opened () then user_err Pp.(str "Registering a kernel type is not allowed in sections"); - let pind = match Id.to_string id with - | "ind_bool" -> CPrimitives.PIT_bool - | "ind_carry" -> CPrimitives.PIT_carry - | "ind_pair" -> CPrimitives.PIT_pair - | "ind_cmp" -> CPrimitives.PIT_cmp + let CPrimitives.PIE pind = match Id.to_string id with + | "ind_bool" -> CPrimitives.(PIE PIT_bool) + | "ind_carry" -> CPrimitives.(PIE PIT_carry) + | "ind_pair" -> CPrimitives.(PIE PIT_pair) + | "ind_cmp" -> CPrimitives.(PIE PIT_cmp) + | "ind_f_cmp" -> CPrimitives.(PIE PIT_f_cmp) + | "ind_f_class" -> CPrimitives.(PIE PIT_f_class) | k -> CErrors.user_err Pp.(str "Register: unknown identifier “" ++ str k ++ str "” in the “kernel” namespace") in match gr with |
