diff options
| author | Pierre-Marie Pédrot | 2019-02-04 17:22:36 +0100 |
|---|---|---|
| committer | Pierre-Marie Pédrot | 2019-02-04 17:22:36 +0100 |
| commit | c70412ec8b0bb34b7a5607c07d34607a147d834c (patch) | |
| tree | 0cc6cd76a8f70dfd2f5b55b0db96db4de2ff07a2 | |
| parent | 720ee2730684cc289cef588482323d177e0bea59 (diff) | |
| parent | 191e253d1d1ebd6c76c63b3d83f4228e46196a6e (diff) | |
Merge PR #6914: Primitive integers
Ack-by: JasonGross
Ack-by: SkySkimmer
Ack-by: ejgallego
Ack-by: gares
Ack-by: maximedenes
Ack-by: ppedrot
204 files changed, 6905 insertions, 2772 deletions
diff --git a/.gitignore b/.gitignore index 2e5529ccfb..c30fd850a1 100644 --- a/.gitignore +++ b/.gitignore @@ -147,6 +147,7 @@ plugins/ssr/ssrvernac.ml kernel/byterun/coq_jumptbl.h kernel/copcodes.ml +kernel/uint63.ml ide/index_urls.txt .lia.cache .nia.cache diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index a6858c6802..7ebc2d8a4d 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -9,7 +9,7 @@ stages: variables: # Format: $IMAGE-V$DATE [Cache is not used as of today but kept here # for reference] - CACHEKEY: "bionic_coq-V2018-12-14-V1" + CACHEKEY: "bionic_coq-V2019-01-28-V1" IMAGE: "$CI_REGISTRY_IMAGE:$CACHEKEY" # By default, jobs run in the base switch; override to select another switch OPAM_SWITCH: "base" diff --git a/CHANGES.md b/CHANGES.md index 9d912a63b1..1a0b53f84a 100644 --- a/CHANGES.md +++ b/CHANGES.md @@ -44,6 +44,10 @@ Specification language, type inference solved by writing an explicit `return` clause, sometimes even simply an explicit `return _` clause. +Kernel + +- Added primitive integers + Notations - New command `Declare Scope` to explicitly declare a scope name @@ -155,6 +159,8 @@ Standard Library - Added `Coq.Structures.OrderedTypeEx.String_as_OT` to make strings an ordered type (using lexical order). +- The `Coq.Numbers.Cyclic.Int31` library is deprecated. + Universes - Added `Print Universes Subgraph` variant of `Print Universes`. diff --git a/META.coq.in b/META.coq.in index 159984d87a..ab142ccc43 100644 --- a/META.coq.in +++ b/META.coq.in @@ -411,30 +411,6 @@ package "plugins" ( archive(native) = "nsatz_plugin.cmx" ) - package "natsyntax" ( - - description = "Coq natsyntax plugin" - version = "8.10" - - requires = "" - directory = "syntax" - - archive(byte) = "nat_syntax_plugin.cmo" - archive(native) = "nat_syntax_plugin.cmx" - ) - - package "zsyntax" ( - - description = "Coq zsyntax plugin" - version = "8.10" - - requires = "" - directory = "syntax" - - archive(byte) = "z_syntax_plugin.cmo" - archive(native) = "z_syntax_plugin.cmx" - ) - package "rsyntax" ( description = "Coq rsyntax plugin" @@ -447,16 +423,16 @@ package "plugins" ( archive(native) = "r_syntax_plugin.cmx" ) - package "int31syntax" ( + package "int63syntax" ( - description = "Coq int31syntax plugin" + description = "Coq int63syntax plugin" version = "8.10" requires = "" directory = "syntax" - archive(byte) = "int31_syntax_plugin.cmo" - archive(native) = "int31_syntax_plugin.cmx" + archive(byte) = "int63_syntax_plugin.cmo" + archive(native) = "int63_syntax_plugin.cmx" ) package "string_notation" ( @@ -100,7 +100,7 @@ GENMLGFILES:= $(MLGFILES:.mlg=.ml) export GRAMFILES=$(addprefix gramlib/.pack/gramlib__,Ploc Plexing Gramext Grammar) export GRAMMLFILES := $(addsuffix .ml, $(GRAMFILES)) $(addsuffix .mli, $(GRAMFILES)) export GENGRAMFILES := $(GRAMMLFILES) gramlib/.pack/gramlib.ml -export GENMLFILES:=$(LEXFILES:.mll=.ml) $(YACCFILES:.mly=.ml) $(GENMLGFILES) ide/coqide_os_specific.ml kernel/copcodes.ml +export GENMLFILES:=$(LEXFILES:.mll=.ml) $(YACCFILES:.mly=.ml) $(GENMLGFILES) ide/coqide_os_specific.ml kernel/copcodes.ml kernel/uint63.ml export GENHFILES:=kernel/byterun/coq_jumptbl.h export GENFILES:=$(GENMLFILES) $(GENMLIFILES) $(GENHFILES) diff --git a/Makefile.build b/Makefile.build index 5775569712..26e2819990 100644 --- a/Makefile.build +++ b/Makefile.build @@ -338,6 +338,13 @@ $(COQPP): $(COQPPCMO) coqpp/coqpp_main.ml $(HIDE)$(OCAMLC) -I coqpp $^ -linkall -o $@ ########################################################################### +# Specific rules for Uint63 +########################################################################### +kernel/uint63.ml: kernel/write_uint63.ml kernel/uint63_x86.ml kernel/uint63_amd64.ml + $(SHOW)'WRITE $@' + $(HIDE)(cd kernel && ocaml unix.cma $(shell basename $<)) + +########################################################################### # Main targets (coqtop.opt, coqtop.byte) ########################################################################### @@ -504,7 +511,7 @@ $(FAKEIDEBYTE): $(FAKEIDECMO) | $(IDETOPBYTE) # votour: a small vo explorer (based on the checker) -VOTOURCMO:=clib/cObj.cmo checker/analyze.cmo checker/values.cmo checker/votour.cmo +VOTOURCMO:=clib/cObj.cmo kernel/uint63.cmo checker/analyze.cmo checker/values.cmo checker/votour.cmo bin/votour: $(call bestobj, $(VOTOURCMO)) $(SHOW)'OCAMLBEST -o $@' diff --git a/Makefile.common b/Makefile.common index f998ea867b..8292158ef8 100644 --- a/Makefile.common +++ b/Makefile.common @@ -140,7 +140,7 @@ BTAUTOCMO:=plugins/btauto/btauto_plugin.cmo RTAUTOCMO:=plugins/rtauto/rtauto_plugin.cmo SYNTAXCMO:=$(addprefix plugins/syntax/, \ r_syntax_plugin.cmo \ - int31_syntax_plugin.cmo \ + int63_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 63324bff20..77e70318dd 100644 --- a/checker/analyze.ml +++ b/checker/analyze.ml @@ -106,6 +106,7 @@ end type repr = | RInt of int +| RInt63 of Uint63.t | RBlock of (int * int) (* tag × len *) | RString of string | RPointer of int @@ -119,6 +120,7 @@ type data = type obj = | Struct of int * data array (* tag × data *) +| Int63 of Uint63.t (* Primitive integer *) | String of string module type Input = @@ -255,6 +257,28 @@ let input_header64 chan = in (tag, len) +let input_cstring chan : string = + let buff = Buffer.create 17 in + let rec loop () = + match input_char chan with + | '\o000' -> Buffer.contents buff + | c -> Buffer.add_char buff c |> loop + in loop () + +let input_intL chan : int64 = + 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 + (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 parse_object chan = let data = input_byte chan in if prefix_small_block <= data then @@ -297,6 +321,11 @@ let parse_object chan = let addr = input_int32u chan in for _i = 0 to 15 do ignore (input_byte chan); done; RCode addr + | CODE_CUSTOM -> + begin match input_cstring chan with + | "_j" -> RInt63 (Uint63.of_int64 (input_intL chan)) + | s -> Printf.eprintf "Unhandled custom code: %s" s; assert false + end | CODE_DOUBLE_ARRAY32_LITTLE | CODE_DOUBLE_BIG | CODE_DOUBLE_LITTLE @@ -304,8 +333,7 @@ let parse_object chan = | CODE_DOUBLE_ARRAY8_LITTLE | CODE_DOUBLE_ARRAY32_BIG | CODE_INFIXPOINTER - | CODE_CUSTOM -> - Printf.eprintf "Unhandled code %04x\n%!" data; assert false + -> Printf.eprintf "Unhandled code %04x\n%!" data; assert false let parse chan = let (magic, len, _, _, size) = parse_header chan in @@ -337,6 +365,11 @@ let parse chan = | RCode addr -> let data = Fun addr in data, None + | RInt63 i -> + let data = Ptr !current_object in + let () = LargeArray.set memory !current_object (Int63 i) in + let () = incr current_object in + data, None in let rec fill block off accu = @@ -400,6 +433,7 @@ let instantiate (p, mem) = for i = 0 to len - 1 do let obj = match LargeArray.get mem i with | Struct (tag, blk) -> Obj.new_block tag (Array.length blk) + | Int63 i -> Obj.repr i | String str -> Obj.repr str in LargeArray.set ans i obj @@ -418,6 +452,7 @@ let instantiate (p, mem) = for k = 0 to Array.length blk - 1 do Obj.set_field obj k (get_data blk.(k)) done + | Int63 _ | String _ -> () done; get_data p diff --git a/checker/analyze.mli b/checker/analyze.mli index d7770539df..029f595959 100644 --- a/checker/analyze.mli +++ b/checker/analyze.mli @@ -1,3 +1,4 @@ +(** Representation of data allocated on the OCaml heap. *) type data = | Int of int | Ptr of int @@ -6,6 +7,7 @@ type data = type obj = | Struct of int * data array (* tag × data *) +| Int63 of Uint63.t (* Primitive integer *) | String of string module LargeArray : diff --git a/checker/mod_checking.ml b/checker/mod_checking.ml index 086dd17e39..c33c6d5d09 100644 --- a/checker/mod_checking.ml +++ b/checker/mod_checking.ml @@ -66,6 +66,7 @@ let mk_mtb mp sign delta = let rec check_module env mp mb = Flags.if_verbose Feedback.msg_notice (str " checking module: " ++ str (ModPath.to_string mp)); + let env = Modops.add_retroknowledge mb.mod_retroknowledge env in let (_:module_signature) = check_signature env mb.mod_type mb.mod_mp mb.mod_delta in diff --git a/checker/safe_checking.ml b/checker/safe_checking.ml index 6dc2953060..4a64039e30 100644 --- a/checker/safe_checking.ml +++ b/checker/safe_checking.ml @@ -16,6 +16,7 @@ let import senv clib univs digest = let env = Safe_typing.env_of_safe_env senv in let env = push_context_set ~strict:true mb.mod_constraints env in let env = push_context_set ~strict:true univs env in + let env = Modops.add_retroknowledge mb.mod_retroknowledge env in Mod_checking.check_module env mb.mod_mp mb; let (_,senv) = Safe_typing.import clib univs digest senv in senv diff --git a/checker/validate.ml b/checker/validate.ml index b85944f94f..72cf38ebe6 100644 --- a/checker/validate.ml +++ b/checker/validate.ml @@ -86,6 +86,7 @@ let rec val_gen v ctx o = match v with | Annot (s,v) -> val_gen v (ctx/CtxAnnot s) o | Dyn -> val_dyn ctx o | Proxy { contents = v } -> val_gen v ctx o + | Uint63 -> val_uint63 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 @@ -133,6 +134,10 @@ and val_array v ctx o = val_gen v ctx (Obj.field o i) done +and val_uint63 ctx o = + if not (Uint63.is_uint63 o) then + fail ctx o "not a 63-bit unsigned integer" + let print_frame = function | CtxType t -> t | CtxAnnot t -> t diff --git a/checker/values.ml b/checker/values.ml index 1afe764ca4..7ca2dc8050 100644 --- a/checker/values.ml +++ b/checker/values.ml @@ -34,6 +34,7 @@ type value = | Dyn | Proxy of value ref + | Uint63 let fix (f : value -> value) : value = let self = ref Any in @@ -151,7 +152,8 @@ let rec v_constr = [|v_caseinfo;v_constr;v_constr;Array v_constr|]; (* Case *) [|v_fix|]; (* Fix *) [|v_cofix|]; (* CoFix *) - [|v_proj;v_constr|] (* Proj *) + [|v_proj;v_constr|]; (* Proj *) + [|Uint63|] (* Int *) |]) and v_prec = Tuple ("prec_declaration", @@ -214,9 +216,12 @@ let v_oracle = let v_pol_arity = v_tuple "polymorphic_arity" [|List(Opt v_level);v_univ|] +let v_primitive = + v_enum "primitive" 25 + let v_cst_def = v_sum "constant_def" 0 - [|[|Opt Int|]; [|v_cstr_subst|]; [|v_lazy_constr|]|] + [|[|Opt Int|]; [|v_cstr_subst|]; [|v_lazy_constr|]; [|v_primitive|]|] let v_typing_flags = v_tuple "typing_flags" [|v_bool; v_bool; v_oracle; v_bool; v_bool; v_bool; v_bool|] @@ -288,6 +293,20 @@ 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_type = v_enum "prim_type" 1 + +let v_retro_action = + v_sum "retro_action" 0 [| + [|v_prim_ind; v_ind|]; + [|v_prim_type; v_cst|]; + [|v_cst|]; + |] + +let v_retroknowledge = + v_sum "module_retroknowledge" 1 [|[|List v_retro_action|]|] + let rec v_mae = Sum ("module_alg_expr",0, [|[|v_mp|]; (* SEBident *) @@ -318,7 +337,7 @@ and v_impl = and v_noimpl = v_unit and v_module = Tuple ("module_body", - [|v_mp;v_impl;v_sign;Opt v_mexpr;v_context_set;v_resolver;Any|]) + [|v_mp;v_impl;v_sign;Opt v_mexpr;v_context_set;v_resolver;v_retroknowledge|]) and v_modtype = Tuple ("module_type_body", [|v_mp;v_noimpl;v_sign;Opt v_mexpr;v_context_set;v_resolver;v_unit|]) diff --git a/checker/values.mli b/checker/values.mli index 616b69907f..2ab8da1928 100644 --- a/checker/values.mli +++ b/checker/values.mli @@ -38,6 +38,8 @@ type value = | Proxy of value ref (** Same as the inner value, used to define recursive types *) + | Uint63 + (** 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 3c088b59b5..36014cde73 100644 --- a/checker/votour.ml +++ b/checker/votour.ml @@ -100,6 +100,7 @@ struct init_size seen (fun n -> fold (succ i) (accu + 1 + n) k) os.(i) in fold 0 1 (fun size -> let () = LargeArray.set !sizes p size in k size) + | Int63 _ -> k 0 | String s -> let size = 2 + (String.length s / ws) in let () = LargeArray.set !sizes p size in @@ -116,6 +117,7 @@ struct | Ptr p -> match LargeArray.get !memory p with | Struct (tag, os) -> BLOCK (tag, os) + | Int63 _ -> OTHER (* TODO: pretty-print int63 values *) | String s -> STRING s let input ch = @@ -153,6 +155,7 @@ let rec get_name ?(extra=false) = function |Annot (s,v) -> s^"/"^get_name ~extra v |Dyn -> "<dynamic>" | Proxy v -> get_name ~extra !v + | Uint63 -> "Uint63" (** For tuples, its quite handy to display the inner 1st string (if any). Cf. [structure_body] for instance *) @@ -257,6 +260,7 @@ let rec get_children v o pos = match v with end |Fail s -> raise Forbidden | Proxy v -> get_children !v o pos + | Uint63 -> raise Exit let get_children v o pos = try get_children v o pos diff --git a/dev/ci/docker/bionic_coq/Dockerfile b/dev/ci/docker/bionic_coq/Dockerfile index baf470e021..4cd7faf757 100644 --- a/dev/ci/docker/bionic_coq/Dockerfile +++ b/dev/ci/docker/bionic_coq/Dockerfile @@ -1,4 +1,4 @@ -# CACHEKEY: "bionic_coq-V2018-12-14-V1" +# CACHEKEY: "bionic_coq-V2019-01-28-V1" # ^^ Update when modifying this file. FROM ubuntu:bionic @@ -38,7 +38,7 @@ ENV COMPILER="4.05.0" # `num` does not have a version number as the right version to install varies # with the compiler version. ENV BASE_OPAM="num ocamlfind.1.8.0 dune.1.6.2 ounit.2.0.8 odoc.1.3.0" \ - CI_OPAM="menhir.20180530 elpi.1.1.0 ocamlgraph.1.8.8" + CI_OPAM="menhir.20181113 elpi.1.1.0 ocamlgraph.1.8.8" # BASE switch; CI_OPAM contains Coq's CI dependencies. ENV COQIDE_OPAM="lablgtk.2.18.5 conf-gtksourceview.2" diff --git a/dev/ci/user-overlays/06914-maximedenes-primitive-integers.sh b/dev/ci/user-overlays/06914-maximedenes-primitive-integers.sh new file mode 100644 index 0000000000..6e89741e29 --- /dev/null +++ b/dev/ci/user-overlays/06914-maximedenes-primitive-integers.sh @@ -0,0 +1,9 @@ +if [ "$CI_PULL_REQUEST" = "6914" ] || [ "$CI_BRANCH" = "primitive-bool-list" ]; then + + bignums_CI_REF=primitive-integers + bignums_CI_GITURL=https://github.com/vbgl/bignums + + mtac2_CI_REF=primitive-integers + mtac2_CI_GITURL=https://github.com/vbgl/Mtac2 + +fi diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 8f207d1e0a..2629cf8626 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -292,6 +292,8 @@ let constr_display csr = ^(Array.fold_right (fun x i -> (name_display x)^(if not(i="") then (";"^i) else "")) lna "")^"," ^(array_display bl)^")" + | Int i -> + "Int("^(Uint63.to_string i)^")" and array_display v = "[|"^ @@ -421,6 +423,8 @@ let print_pure_constr csr = print_cut(); done in print_string"{"; print_fix (); print_string"}" + | Int i -> + print_string ("Int("^(Uint63.to_string i)^")") 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 ea126e2756..dc30793a6e 100644 --- a/dev/vm_printers.ml +++ b/dev/vm_printers.ml @@ -82,6 +82,7 @@ and ppwhd whd = | Vcofix _ -> print_string "cofix" | Vconstr_const i -> print_string "C(";print_int i;print_string")" | Vconstr_block b -> ppvblock b + | Vint64 i -> printf "int64(%LiL)" i | Vatom_stk(a,s) -> open_hbox();ppatom a;close_box(); print_string"@";ppstack s diff --git a/doc/sphinx/user-extensions/syntax-extensions.rst b/doc/sphinx/user-extensions/syntax-extensions.rst index ae66791b0c..105b0445fd 100644 --- a/doc/sphinx/user-extensions/syntax-extensions.rst +++ b/doc/sphinx/user-extensions/syntax-extensions.rst @@ -1497,12 +1497,12 @@ Numeral notations for only non-negative integers, and the given numeral is negative. - .. exn:: @ident should go from Decimal.int to @type or (option @type). Instead of Decimal.int, the types Decimal.uint or Z could be used (you may need to require BinNums or Decimal first). + .. exn:: @ident should go from Decimal.int to @type or (option @type). Instead of Decimal.int, the types Decimal.uint or Z or Int63.int could be used (you may need to require BinNums or Decimal or Int63 first). The parsing function given to the :cmd:`Numeral Notation` vernacular is not of the right type. - .. exn:: @ident should go from @type to Decimal.int or (option Decimal.int). Instead of Decimal.int, the types Decimal.uint or Z could be used (you may need to require BinNums or Decimal first). + .. exn:: @ident should go from @type to Decimal.int or (option Decimal.int). Instead of Decimal.int, the types Decimal.uint or Z or Int63.int could be used (you may need to require BinNums or Decimal or Int63 first). The printing function given to the :cmd:`Numeral Notation` vernacular is not of the right type. diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template index c33df52038..7b21b67eea 100644 --- a/doc/stdlib/index-list.html.template +++ b/doc/stdlib/index-list.html.template @@ -259,7 +259,7 @@ through the <tt>Require Import</tt> command.</p> </dd> <dt> <b> Cyclic</b>: - Abstract and 31-bits-based cyclic arithmetic + Abstract and 63-bits-based cyclic arithmetic </dt> <dd> theories/Numbers/Cyclic/Abstract/CyclicAxioms.v @@ -268,11 +268,14 @@ through the <tt>Require Import</tt> command.</p> theories/Numbers/Cyclic/Int31/Cyclic31.v theories/Numbers/Cyclic/Int31/Ring31.v theories/Numbers/Cyclic/Int31/Int31.v + theories/Numbers/Cyclic/Int63/Cyclic63.v + theories/Numbers/Cyclic/Int63/Int63.v + theories/Numbers/Cyclic/Int63/Ring63.v theories/Numbers/Cyclic/ZModulo/ZModulo.v </dd> <dt> <b> Natural</b>: - Abstract and 31-bits-words-based natural arithmetic + Abstract and 63-bits-words-based natural arithmetic </dt> <dd> theories/Numbers/Natural/Abstract/NAdd.v @@ -300,7 +303,7 @@ through the <tt>Require Import</tt> command.</p> </dd> <dt> <b> Integer</b>: - Abstract and concrete (especially 31-bits-words-based) integer + Abstract and concrete (especially 63-bits-words-based) integer arithmetic </dt> <dd> diff --git a/engine/eConstr.ml b/engine/eConstr.ml index 24d161d00a..8493119ee5 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -73,6 +73,7 @@ let mkFix f = of_kind (Fix f) let mkCoFix f = of_kind (CoFix f) let mkProj (p, c) = of_kind (Proj (p, c)) let mkArrow t1 t2 = of_kind (Prod (Anonymous, t1, t2)) +let mkInt i = of_kind (Int i) let mkRef (gr,u) = let open GlobRef in match gr with | ConstRef c -> mkConstU (c,u) @@ -81,6 +82,7 @@ let mkRef (gr,u) = let open GlobRef in match gr with | VarRef x -> mkVar x let applist (f, arg) = mkApp (f, Array.of_list arg) +let applistc f arg = mkApp (f, Array.of_list arg) let isRel sigma c = match kind sigma c with Rel _ -> true | _ -> false let isVar sigma c = match kind sigma c with Var _ -> true | _ -> false @@ -328,7 +330,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 _) -> () + | Construct _ | Int _) -> () | 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 49cbc4d7e5..2f4cf7d5d0 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -124,10 +124,12 @@ val mkCase : case_info * t * t * t array -> t val mkFix : (t, t) pfixpoint -> t val mkCoFix : (t, t) pcofixpoint -> t val mkArrow : t -> t -> t +val mkInt : Uint63.t -> t val mkRef : GlobRef.t * EInstance.t -> t val applist : t * t list -> t +val applistc : t -> t list -> t val mkProd_or_LetIn : rel_declaration -> t -> t val mkLambda_or_LetIn : rel_declaration -> t -> t diff --git a/engine/namegen.ml b/engine/namegen.ml index 018eca1ba2..294b61fd3d 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) with Name id -> id | _ -> assert false) - | Sort _ | Rel _ | Meta _|Evar _|Case (_, _, _, _) -> None + | Sort _ | Rel _ | Meta _|Evar _|Case (_, _, _, _) | Int _ -> None in hdrec c @@ -163,6 +163,7 @@ let hdchar env sigma c = lowercase_first_char id | Evar _ (* We could do better... *) | Meta _ | Case (_, _, _, _) -> "y" + | Int _ -> "i" in hdrec 0 c diff --git a/engine/termops.ml b/engine/termops.ml index 3e902129f3..579100ad4c 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 _) -> c + | Construct _ | Int _) -> 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 _) -> cstr + | Construct _ | Int _) -> 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 13078840ef..8e49800982 100644 --- a/interp/constrextern.ml +++ b/interp/constrextern.ml @@ -967,6 +967,8 @@ let rec extern inctx (custom,scopes as allscopes) vars r = | GCast (c, c') -> CCast (sub_extern true scopes vars c, map_cast_type (extern_typ scopes vars) c') + | GInt i -> + CPrim(Numeral (Uint63.to_string i,true)) in insert_coercion coercion (CAst.make ?loc c) @@ -1312,6 +1314,7 @@ let rec glob_of_pat avoid env sigma pat = DAst.make @@ match pat with Array.map (fun (_,_,ty) -> ty) v, Array.map (fun (_,bd,_) -> bd) v) | PSort s -> GSort s + | PInt i -> GInt i 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/declare.ml b/interp/declare.ml index 6778fa1e7a..ea6ed8321d 100644 --- a/interp/declare.ml +++ b/interp/declare.ml @@ -442,6 +442,9 @@ let assumption_message id = discussion on coqdev: "Chapter 4 of the Reference Manual", 8/10/2015) *) Flags.if_verbose Feedback.msg_info (Id.print id ++ str " is declared") +let register_message id = + Flags.if_verbose Feedback.msg_info (Id.print id ++ str " is registered") + (** Monomorphic universes need to survive sections. *) let input_universe_context : Univ.ContextSet.t -> Libobject.obj = diff --git a/interp/declare.mli b/interp/declare.mli index 468e056909..669657af6f 100644 --- a/interp/declare.mli +++ b/interp/declare.mli @@ -74,6 +74,7 @@ val declare_mind : mutual_inductive_entry -> Libobject.object_name * bool val definition_message : Id.t -> unit val assumption_message : Id.t -> unit +val register_message : Id.t -> unit val fixpoint_message : int array option -> Id.t list -> unit val cofixpoint_message : Id.t list -> unit val recursive_message : bool (** true = fixpoint *) -> diff --git a/interp/dumpglob.ml b/interp/dumpglob.ml index f5be0ddbae..a537b4848c 100644 --- a/interp/dumpglob.ml +++ b/interp/dumpglob.ml @@ -101,6 +101,7 @@ let type_of_logical_kind = function | Property | Proposition | Corollary -> "thm") + | IsPrimitive -> "prim" let type_of_global_ref gr = if Typeclasses.is_class gr then diff --git a/interp/impargs.ml b/interp/impargs.ml index 8a89bcdf26..959455dfd2 100644 --- a/interp/impargs.ml +++ b/interp/impargs.ml @@ -225,7 +225,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 _ -> assert false + | Prod _ | Meta _ | Cast _ | Int _ -> assert false let is_rigid env sigma t = let open Context.Rel.Declaration in diff --git a/interp/notation.ml b/interp/notation.ml index ca27d439fb..bc68d97bb8 100644 --- a/interp/notation.ml +++ b/interp/notation.ml @@ -574,6 +574,7 @@ type target_kind = | Int of int_ty (* Coq.Init.Decimal.int + uint *) | UInt of Names.inductive (* Coq.Init.Decimal.uint *) | Z of z_pos_ty (* Coq.Numbers.BinNums.Z and positive *) + | Int63 (* Coq.Numbers.Cyclic.Int63.Int63.int *) type string_target_kind = | ListByte @@ -637,6 +638,7 @@ let rec constr_of_glob env sigma g = match DAst.get g with let sigma,c = constr_of_glob env sigma gc in let sigma,cl = List.fold_left_map (constr_of_glob env) sigma gcl in sigma,mkApp (c, Array.of_list cl) + | Glob_term.GInt i -> sigma, mkInt i | _ -> raise NotAValidPrimToken @@ -649,6 +651,7 @@ let rec glob_of_constr token_kind ?loc env sigma c = match Constr.kind c with | Const (c, _) -> DAst.make ?loc (Glob_term.GRef (ConstRef c, None)) | Ind (ind, _) -> DAst.make ?loc (Glob_term.GRef (IndRef ind, None)) | Var id -> DAst.make ?loc (Glob_term.GRef (VarRef id, None)) + | Int i -> DAst.make ?loc (Glob_term.GInt i) | _ -> Loc.raise ?loc (PrimTokenNotationError(token_kind,env,sigma,UnexpectedTerm c)) let no_such_prim_token uninterpreted_token_kind ?loc ty = @@ -683,6 +686,16 @@ let uninterp to_raw o (Glob_term.AnyGlobConstr n) = end +(** Conversion from bigint to int63 *) +let rec int63_of_pos_bigint i = + let open Bigint in + if equal i zero then Uint63.of_int 0 + else + let (quo,rem) = div2_with_rest i in + if rem then Uint63.add (Uint63.of_int 1) + (Uint63.mul (Uint63.of_int 2) (int63_of_pos_bigint quo)) + else Uint63.mul (Uint63.of_int 2) (int63_of_pos_bigint quo) + module Numeral = struct (** * Numeral notation *) open PrimTokenNotation @@ -838,6 +851,32 @@ let bigint_of_z z = match Constr.kind z with end | _ -> raise NotAValidPrimToken +(** Now, [Int63] from/to bigint *) + +let int63_of_pos_bigint ?loc n = + let i = int63_of_pos_bigint n in + mkInt i + +let error_negative ?loc = + CErrors.user_err ?loc ~hdr:"interp_int63" (Pp.str "int63 are only non-negative numbers.") + +let error_overflow ?loc n = + CErrors.user_err ?loc ~hdr:"interp_int63" Pp.(str "overflow in int63 literal: " ++ str (Bigint.to_string n)) + +let interp_int63 ?loc n = + let open Bigint in + if is_pos_or_zero n + then + if less_than n (pow two 63) + then int63_of_pos_bigint ?loc n + else error_overflow ?loc n + else error_negative ?loc + +let bigint_of_int63 c = + match Constr.kind c with + | Int i -> Bigint.of_string (Uint63.to_string i) + | _ -> raise NotAValidPrimToken + let big2raw n = if Bigint.is_pos_or_zero n then (Bigint.to_string n, true) else (Bigint.to_string (Bigint.neg n), false) @@ -856,6 +895,7 @@ let interp o ?loc n = | UInt uint_ty when snd n -> coquint_of_rawnum uint_ty (fst n) | UInt _ (* n <= 0 *) -> no_such_prim_token "number" ?loc o.ty_name | Z z_pos_ty -> z_of_bigint z_pos_ty (raw2big n) + | Int63 -> interp_int63 ?loc (raw2big n) in let env = Global.env () in let sigma = Evd.from_env env in @@ -877,6 +917,7 @@ let uninterp o n = | (Int _, c) -> rawnum_of_coqint c | (UInt _, c) -> (rawnum_of_coquint c, true) | (Z _, c) -> big2raw (bigint_of_z c) + | (Int63, c) -> big2raw (bigint_of_int63 c) end o n end diff --git a/interp/notation.mli b/interp/notation.mli index a482e00e81..5dcc96dc29 100644 --- a/interp/notation.mli +++ b/interp/notation.mli @@ -127,6 +127,7 @@ type target_kind = | Int of int_ty (* Coq.Init.Decimal.int + uint *) | UInt of Names.inductive (* Coq.Init.Decimal.uint *) | Z of z_pos_ty (* Coq.Numbers.BinNums.Z and positive *) + | Int63 (* Coq.Numbers.Cyclic.Int63.Int63.int *) type string_target_kind = | ListByte @@ -320,3 +321,6 @@ val entry_has_ident : notation_entry_level -> bool (** Rem: printing rules for primitive token are canonical *) val with_notation_protection : ('a -> 'b) -> 'a -> 'b + +(** Conversion from bigint to int63 *) +val int63_of_pos_bigint : Bigint.bigint -> Uint63.t diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 8d225fe683..890c24e633 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -89,9 +89,11 @@ let rec eq_notation_constr (vars1,vars2 as vars) t1 t2 = match t1, t2 with glob_sort_eq s1 s2 | NCast (t1, c1), NCast (t2, c2) -> (eq_notation_constr vars) t1 t2 && cast_type_eq (eq_notation_constr vars) c1 c2 +| NInt i1, NInt i2 -> + Uint63.equal i1 i2 | (NRef _ | NVar _ | NApp _ | NHole _ | NList _ | NLambda _ | NProd _ | NBinderList _ | NLetIn _ | NCases _ | NLetTuple _ | NIf _ - | NRec _ | NSort _ | NCast _ ), _ -> false + | NRec _ | NSort _ | NCast _ | NInt _), _ -> false (**********************************************************************) (* Re-interpret a notation as a glob_constr, taking care of binders *) @@ -220,6 +222,7 @@ let glob_constr_of_notation_constr_with_binders ?loc g f e nc = | NSort x -> GSort x | NHole (x, naming, arg) -> GHole (x, naming, arg) | NRef x -> GRef (x,None) + | NInt i -> GInt i let glob_constr_of_notation_constr ?loc x = let rec aux () x = @@ -435,6 +438,7 @@ let notation_constr_and_vars_of_glob_constr recvars a = NRec (fk,idl,dll,Array.map aux tl,Array.map aux bl) | GCast (c,k) -> NCast (aux c,map_cast_type aux k) | GSort s -> NSort s + | GInt i -> NInt i | GHole (w,naming,arg) -> if arg != None then has_ltac := true; NHole (w, naming, arg) @@ -623,6 +627,7 @@ let rec subst_notation_constr subst bound raw = NRec (fk,idl,dll',tl',bl') | NSort _ -> raw + | NInt _ -> raw | NHole (knd, naming, solve) -> let nknd = match knd with @@ -1189,6 +1194,7 @@ let rec match_ inner u alp metas sigma a1 a2 = match_cast (match_in u alp metas) (match_in u alp metas sigma t1 t2) c1 c2 | GSort (GType _), NSort (GType _) when not u -> sigma | GSort s1, NSort s2 when glob_sort_eq s1 s2 -> sigma + | GInt i1, NInt i2 when Uint63.equal i1 i2 -> sigma | GPatVar _, NHole _ -> (*Don't hide Metas, they bind in ltac*) raise No_match | a, NHole _ -> sigma @@ -1216,7 +1222,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 _ ), _ -> raise No_match + | GCast _ | GInt _ ), _ -> raise No_match and match_in u = match_ true u diff --git a/interp/notation_term.ml b/interp/notation_term.ml index 0ef1f267f6..6fe20486dc 100644 --- a/interp/notation_term.ml +++ b/interp/notation_term.ml @@ -43,6 +43,7 @@ type notation_constr = notation_constr array * notation_constr array | NSort of glob_sort | NCast of notation_constr * notation_constr cast_type + | NInt of Uint63.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 be2b05da8d..0865487c98 100644 --- a/kernel/byterun/coq_fix_code.c +++ b/kernel/byterun/coq_fix_code.c @@ -36,20 +36,15 @@ void init_arity () { arity[PUSHACC6]=arity[PUSHACC7]=arity[ENVACC1]=arity[ENVACC2]= arity[ENVACC3]=arity[ENVACC4]=arity[PUSHENVACC1]=arity[PUSHENVACC2]= arity[PUSHENVACC3]=arity[PUSHENVACC4]=arity[APPLY1]=arity[APPLY2]= - arity[APPLY3]=arity[RESTART]=arity[OFFSETCLOSUREM2]= + arity[APPLY3]=arity[APPLY4]=arity[RESTART]=arity[OFFSETCLOSUREM2]= arity[OFFSETCLOSURE0]=arity[OFFSETCLOSURE2]=arity[PUSHOFFSETCLOSUREM2]= arity[PUSHOFFSETCLOSURE0]=arity[PUSHOFFSETCLOSURE2]= arity[GETFIELD0]=arity[GETFIELD1]=arity[SETFIELD0]=arity[SETFIELD1]= arity[CONST0]=arity[CONST1]=arity[CONST2]=arity[CONST3]= arity[PUSHCONST0]=arity[PUSHCONST1]=arity[PUSHCONST2]=arity[PUSHCONST3]= arity[ACCUMULATE]=arity[STOP]=arity[MAKEPROD]= - arity[ADDINT31]=arity[ADDCINT31]=arity[ADDCARRYCINT31]= - arity[SUBINT31]=arity[SUBCINT31]=arity[SUBCARRYCINT31]= - arity[MULCINT31]=arity[MULINT31]=arity[COMPAREINT31]= - arity[DIV21INT31]=arity[DIVINT31]=arity[ADDMULDIVINT31]= - arity[HEAD0INT31]=arity[TAIL0INT31]= - arity[COMPINT31]=arity[DECOMPINT31]= - arity[ORINT31]=arity[ANDINT31]=arity[XORINT31]=0; + arity[ADDINT63]=arity[SUBINT63]=arity[LTINT63]=arity[LEINT63]= + arity[ISINT]=arity[AREINT2]=0; /* instruction with one operand */ arity[ACC]=arity[PUSHACC]=arity[POP]=arity[ENVACC]=arity[PUSHENVACC]= arity[PUSH_RETADDR]=arity[APPLY]=arity[APPTERM1]=arity[APPTERM2]= @@ -58,10 +53,20 @@ void init_arity () { arity[MAKEBLOCK1]=arity[MAKEBLOCK2]=arity[MAKEBLOCK3]=arity[MAKEBLOCK4]= arity[MAKEACCU]=arity[CONSTINT]=arity[PUSHCONSTINT]=arity[GRABREC]= arity[PUSHFIELDS]=arity[GETFIELD]=arity[SETFIELD]= - arity[BRANCH]=arity[ISCONST]=arity[ENSURESTACKCAPACITY]=1; + arity[BRANCH]=arity[ENSURESTACKCAPACITY]= + arity[CHECKADDINT63]=arity[CHECKADDCINT63]=arity[CHECKADDCARRYCINT63]= + arity[CHECKSUBINT63]=arity[CHECKSUBCINT63]=arity[CHECKSUBCARRYCINT63]= + arity[CHECKMULINT63]=arity[CHECKMULCINT63]= + arity[CHECKDIVINT63]=arity[CHECKMODINT63]=arity[CHECKDIVEUCLINT63]= + arity[CHECKDIV21INT63]= + arity[CHECKLXORINT63]=arity[CHECKLORINT63]=arity[CHECKLANDINT63]= + arity[CHECKLSLINT63]=arity[CHECKLSRINT63]=arity[CHECKADDMULDIVINT63]= + arity[CHECKLSLINT63CONST1]=arity[CHECKLSRINT63CONST1]= + arity[CHECKEQINT63]=arity[CHECKLTINT63]=arity[CHECKLEINT63]= + arity[CHECKCOMPAREINT63]=arity[CHECKHEAD0INT63]=arity[CHECKTAIL0INT63]=1; /* instruction with two operands */ arity[APPTERM]=arity[MAKEBLOCK]=arity[CLOSURE]= - arity[ARECONST]=arity[PROJ]=2; + arity[PROJ]=2; /* instruction with four operands */ arity[MAKESWITCHBLOCK]=4; /* instruction with arbitrary operands */ diff --git a/kernel/byterun/coq_fix_code.h b/kernel/byterun/coq_fix_code.h index 638d6b5ab5..5a233e6178 100644 --- a/kernel/byterun/coq_fix_code.h +++ b/kernel/byterun/coq_fix_code.h @@ -29,6 +29,7 @@ void init_arity(); value coq_tcode_of_code(value code); value coq_makeaccu (value i); value coq_pushpop (value i); +value coq_accucond (value i); value coq_is_accumulate_code(value code); #endif /* _COQ_FIX_CODE_ */ diff --git a/kernel/byterun/coq_instruct.h b/kernel/byterun/coq_instruct.h index d92e85fdf8..c7abedaed6 100644 --- a/kernel/byterun/coq_instruct.h +++ b/kernel/byterun/coq_instruct.h @@ -26,7 +26,7 @@ enum instructions { ENVACC1, ENVACC2, ENVACC3, ENVACC4, ENVACC, PUSHENVACC1, PUSHENVACC2, PUSHENVACC3, PUSHENVACC4, PUSHENVACC, PUSH_RETADDR, - APPLY, APPLY1, APPLY2, APPLY3, + APPLY, APPLY1, APPLY2, APPLY3, APPLY4, APPTERM, APPTERM1, APPTERM2, APPTERM3, RETURN, RESTART, GRAB, GRABREC, CLOSURE, CLOSUREREC, CLOSURECOFIX, @@ -46,15 +46,21 @@ enum instructions { MAKESWITCHBLOCK, MAKEACCU, MAKEPROD, /* spiwack: */ BRANCH, - ADDINT31, ADDCINT31, ADDCARRYCINT31, - SUBINT31, SUBCINT31, SUBCARRYCINT31, - MULCINT31, MULINT31, DIV21INT31, DIVINT31, - ADDMULDIVINT31, COMPAREINT31, - HEAD0INT31, TAIL0INT31, - ISCONST, ARECONST, - COMPINT31, DECOMPINT31, - ORINT31, ANDINT31, XORINT31, -/* /spiwack */ + CHECKADDINT63, ADDINT63, CHECKADDCINT63, CHECKADDCARRYCINT63, + CHECKSUBINT63, SUBINT63, CHECKSUBCINT63, CHECKSUBCARRYCINT63, + CHECKMULINT63, CHECKMULCINT63, + CHECKDIVINT63, CHECKMODINT63, CHECKDIVEUCLINT63, CHECKDIV21INT63, + CHECKLXORINT63, CHECKLORINT63, CHECKLANDINT63, + CHECKLSLINT63, CHECKLSRINT63, CHECKADDMULDIVINT63, + CHECKLSLINT63CONST1, CHECKLSRINT63CONST1, + + CHECKEQINT63, CHECKLTINT63, LTINT63, CHECKLEINT63, LEINT63, + CHECKCOMPAREINT63, + + CHECKHEAD0INT63, CHECKTAIL0INT63, + + ISINT, AREINT2, + STOP }; diff --git a/kernel/byterun/coq_interp.c b/kernel/byterun/coq_interp.c index a944dbb06c..d2c88bffcc 100644 --- a/kernel/byterun/coq_interp.c +++ b/kernel/byterun/coq_interp.c @@ -23,6 +23,12 @@ #include "coq_memory.h" #include "coq_values.h" +#ifdef ARCH_SIXTYFOUR +#include "coq_uint63_native.h" +#else +#include "coq_uint63_emul.h" +#endif + /* spiwack: I append here a few macros for value/number manipulation */ #define uint32_of_value(val) (((uint32_t)(val)) >> 1) #define value_of_uint32(i) ((value)((((uint32_t)(i)) << 1) | 1)) @@ -155,6 +161,38 @@ if (sp - num_args < coq_stack_threshold) { \ #endif #endif +#define CheckInt1() do{ \ + if (Is_uint63(accu)) 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; \ + } \ + }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 AllocCarry(cond) Alloc_small(accu, 1, (cond)? coq_tag_C1 : coq_tag_C0) +#define AllocPair() Alloc_small(accu, 2, coq_tag_pair) + /* For signal handling, we hijack some code from the caml runtime */ extern intnat caml_signals_are_pending; @@ -372,8 +410,10 @@ value coq_interprete goto check_stack; } Instruct(APPLY1) { - value arg1 = sp[0]; + value arg1; + apply1: print_instr("APPLY1"); + arg1 = sp[0]; sp -= 3; sp[0] = arg1; sp[1] = (value)pc; @@ -388,10 +428,14 @@ value coq_interprete coq_extra_args = 0; goto check_stack; } + Instruct(APPLY2) { - value arg1 = sp[0]; - value arg2 = sp[1]; + value arg1; + value arg2; + apply2: print_instr("APPLY2"); + arg1 = sp[0]; + arg2 = sp[1]; sp -= 3; sp[0] = arg1; sp[1] = arg2; @@ -403,11 +447,16 @@ value coq_interprete coq_extra_args = 1; goto check_stack; } + Instruct(APPLY3) { - value arg1 = sp[0]; - value arg2 = sp[1]; - value arg3 = sp[2]; + value arg1; + value arg2; + value arg3; + apply3: print_instr("APPLY3"); + arg1 = sp[0]; + arg2 = sp[1]; + arg3 = sp[2]; sp -= 3; sp[0] = arg1; sp[1] = arg2; @@ -420,7 +469,28 @@ value coq_interprete coq_extra_args = 2; goto check_stack; } - /* Stack checks */ + + Instruct(APPLY4) { + value arg1 = sp[0]; + value arg2 = sp[1]; + value arg3 = sp[2]; + value arg4 = sp[3]; + print_instr("APPLY4"); + sp -= 3; + sp[0] = arg1; + sp[1] = arg2; + sp[2] = arg3; + sp[3] = arg4; + sp[4] = (value)pc; + sp[5] = coq_env; + sp[6] = Val_long(coq_extra_args); + pc = Code_val(accu); + coq_env = accu; + coq_extra_args = 3; + goto check_stack; + } + + /* Stack checks */ check_stack: print_instr("check_stack"); @@ -1127,7 +1197,6 @@ value coq_interprete Next; } - /* spiwack: code for interpreting compiled integers */ Instruct(BRANCH) { /* unconditional branching */ print_instr("BRANCH"); @@ -1136,338 +1205,339 @@ value coq_interprete Next; } - Instruct(ADDINT31) { + Instruct(CHECKADDINT63){ + print_instr("CHECKADDINT63"); + CheckInt2(); + } + Instruct(ADDINT63) { /* Adds the integer in the accumulator with the one ontop of the stack (which is poped)*/ - print_instr("ADDINT31"); - accu = - (value)((uint32_t) accu + (uint32_t) *sp++ - 1); - /* nota,unlike CaML we don't want - to have a different behavior depending on the - architecture. Thus we cast the operand to uint32 */ + print_instr("ADDINT63"); + accu = uint63_add(accu, *sp++); Next; } - Instruct (ADDCINT31) { - print_instr("ADDCINT31"); + Instruct (CHECKADDCINT63) { + print_instr("CHECKADDCINT63"); + CheckInt2(); /* returns the sum with a carry */ - uint32_t s; - s = (uint32_t)accu + (uint32_t)*sp++ - 1; - if( (uint32_t)s < (uint32_t)accu ) { - /* carry */ - Alloc_small(accu, 1, 2); /* ( _ , arity, tag ) */ - } - else { - /*no carry */ - Alloc_small(accu, 1, 1); - } - Field(accu, 0)=(value)s; + value s; + s = uint63_add(accu, *sp++); + AllocCarry(uint63_lt(s,accu)); + Field(accu, 0) = s; Next; } - Instruct (ADDCARRYCINT31) { - print_instr("ADDCARRYCINT31"); + Instruct (CHECKADDCARRYCINT63) { + print_instr("ADDCARRYCINT63"); + CheckInt2(); /* returns the sum plus one with a carry */ - uint32_t s; - s = (uint32_t)accu + (uint32_t)*sp++ + 1; - if( (uint32_t)s <= (uint32_t)accu ) { - /* carry */ - Alloc_small(accu, 1, 2); /* ( _ , arity, tag ) */ - } - else { - /*no carry */ - Alloc_small(accu, 1, 1); - } - Field(accu, 0)=(value)s; + value s; + s = uint63_addcarry(accu, *sp++); + AllocCarry(uint63_leq(s, accu)); + Field(accu, 0) = s; Next; } - Instruct (SUBINT31) { - print_instr("SUBINT31"); + Instruct (CHECKSUBINT63) { + print_instr("CHECKSUBINT63"); + CheckInt2(); + } + Instruct (SUBINT63) { + print_instr("SUBINT63"); /* returns the subtraction */ - accu = - (value)((uint32_t) accu - (uint32_t) *sp++ + 1); + accu = uint63_sub(accu, *sp++); Next; } - Instruct (SUBCINT31) { - print_instr("SUBCINT31"); + Instruct (CHECKSUBCINT63) { + print_instr("SUBCINT63"); + CheckInt2(); /* returns the subtraction with a carry */ - uint32_t b; - uint32_t s; - b = (uint32_t)*sp++; - s = (uint32_t)accu - b + 1; - if( (uint32_t)accu < b ) { - /* carry */ - Alloc_small(accu, 1, 2); /* ( _ , arity, tag ) */ - } - else { - /*no carry */ - Alloc_small(accu, 1, 1); - } - Field(accu, 0)=(value)s; + value b; + value s; + b = *sp++; + s = uint63_sub(accu,b); + AllocCarry(uint63_lt(accu,b)); + Field(accu, 0) = s; Next; } - Instruct (SUBCARRYCINT31) { - print_instr("SUBCARRYCINT31"); + Instruct (CHECKSUBCARRYCINT63) { + print_instr("SUBCARRYCINT63"); + CheckInt2(); /* returns the subtraction minus one with a carry */ - uint32_t b; - uint32_t s; - b = (uint32_t)*sp++; - s = (value)((uint32_t)accu - b - 1); - if( (uint32_t)accu <= b ) { - /* carry */ - Alloc_small(accu, 1, 2); /* ( _ , arity, tag ) */ - } - else { - /*no carry */ - Alloc_small(accu, 1, 1); - } - Field(accu, 0)=(value)s; + value b; + value s; + b = *sp++; + s = uint63_subcarry(accu,b); + AllocCarry(uint63_leq(accu,b)); + Field(accu, 0) = s; Next; } - Instruct (MULINT31) { + Instruct (CHECKMULINT63) { + print_instr("MULINT63"); + CheckInt2(); /* returns the multiplication */ - print_instr("MULINT31"); - accu = - value_of_uint32((uint32_of_value(accu)) * (uint32_of_value(*sp++))); + accu = uint63_mul(accu,*sp++); Next; } - Instruct (MULCINT31) { - /*returns the multiplication on a double size word - (special case for 0) */ - print_instr("MULCINT31"); - uint64_t p; + Instruct (CHECKMULCINT63) { + /*returns the multiplication on a pair */ + print_instr("MULCINT63"); + CheckInt2(); /*accu = 2v+1, *sp=2w+1 ==> p = 2v*w */ - p = UI64_of_value (accu) * UI64_of_uint32 ((*sp++)^1); - if (p == 0) { - accu = (value)1; + /* TODO: implement + p = I64_mul (UI64_of_value (accu), UI64_of_uint32 ((*sp++)^1)); + AllocPair(); */ + /* Field(accu, 0) = (value)(I64_lsr(p,31)|1) ; */ /*higher part*/ + /* Field(accu, 1) = (value)(I64_to_int32(p)|1); */ /*lower part*/ + value x = accu; + AllocPair(); + Field(accu, 1) = uint63_mulc(x, *sp++, &Field(accu, 0)); + Next; + } + + Instruct(CHECKDIVINT63) { + print_instr("CHEKDIVINT63"); + CheckInt2(); + /* spiwack: a priori no need of the NON_STANDARD_DIV_MOD flag + since it probably only concerns negative number. + needs to be checked at this point */ + value divisor; + divisor = *sp++; + if (uint63_eq0(divisor)) { + accu = divisor; } else { - /* the output type is supposed to have a constant constructor - and a non-constant constructor (in that order), the tag - of the non-constant constructor is then 1 */ - Alloc_small(accu, 2, 1); /* ( _ , arity, tag ) */ - /*unsigned shift*/ - Field(accu, 0) = (value)((p >> 31)|1) ; /*higher part*/ - Field(accu, 1) = (value)((uint32_t)p|1); /*lower part*/ - } + accu = uint63_div(accu, divisor); + } Next; } - Instruct (DIV21INT31) { - print_instr("DIV21INT31"); - /* spiwack: takes three int31 (the two first ones represent an - int62) and performs the euclidian division of the - int62 by the int31 */ - uint64_t bigint; - bigint = UI64_of_value(accu); - bigint = (bigint << 31) | UI64_of_value(*sp++); - uint64_t divisor; - divisor = UI64_of_value(*sp++); - Alloc_small(accu, 2, 1); /* ( _ , arity, tag ) */ - if (divisor == 0) { - Field(accu, 0) = 1; /* 2*0+1 */ - Field(accu, 1) = 1; /* 2*0+1 */ + Instruct(CHECKMODINT63) { + print_instr("CHEKMODINT63"); + CheckInt2(); + value divisor; + divisor = *sp++; + if (uint63_eq0(divisor)) { + accu = divisor; } - else { - uint64_t quo, mod; - quo = bigint / divisor; - mod = bigint % divisor; - Field(accu, 0) = value_of_uint32((uint32_t)(quo)); - Field(accu, 1) = value_of_uint32((uint32_t)(mod)); + else { + accu = uint63_mod(accu,divisor); } Next; } - Instruct (DIVINT31) { - print_instr("DIVINT31"); + Instruct (CHECKDIVEUCLINT63) { + print_instr("DIVEUCLINT63"); + CheckInt2(); /* spiwack: a priori no need of the NON_STANDARD_DIV_MOD flag since it probably only concerns negative number. needs to be checked at this point */ - uint32_t divisor; - divisor = uint32_of_value(*sp++); - if (divisor == 0) { + value divisor; + divisor = *sp++; + if (uint63_eq0(divisor)) { Alloc_small(accu, 2, 1); /* ( _ , arity, tag ) */ - Field(accu, 0) = 1; /* 2*0+1 */ - Field(accu, 1) = 1; /* 2*0+1 */ + Field(accu, 0) = divisor; + Field(accu, 1) = divisor; } else { - uint32_t modulus; - modulus = uint32_of_value(accu); + value modulus; + modulus = accu; Alloc_small(accu, 2, 1); /* ( _ , arity, tag ) */ - Field(accu, 0) = value_of_uint32(modulus/divisor); - Field(accu, 1) = value_of_uint32(modulus%divisor); + Field(accu, 0) = uint63_div(modulus,divisor); + Field(accu, 1) = uint63_mod(modulus,divisor); } Next; } - Instruct (ADDMULDIVINT31) { - print_instr("ADDMULDIVINT31"); - /* higher level shift (does shifts and cycles and such) */ - uint32_t shiftby; - shiftby = uint32_of_value(accu); - if (shiftby > 31) { - if (shiftby < 62) { - sp++; - accu = (value)(((((uint32_t)*sp++)^1) << (shiftby - 31)) | 1); - } - else { - sp+=2; - accu = (value)(1); - } + Instruct (CHECKDIV21INT63) { + print_instr("DIV21INT63"); + CheckInt3(); + /* spiwack: takes three int31 (the two first ones represent an + int62) and performs the euclidian division of the + int62 by the int31 */ + /* TODO: implement this + bigint = UI64_of_value(accu); + bigint = I64_or(I64_lsl(bigint, 31),UI64_of_value(*sp++)); + uint64 divisor; + divisor = UI64_of_value(*sp++); + Alloc_small(accu, 2, 1); */ /* ( _ , arity, tag ) */ + /* if (I64_is_zero (divisor)) { + Field(accu, 0) = 1; */ /* 2*0+1 */ + /* Field(accu, 1) = 1; */ /* 2*0+1 */ + /* } + else { + uint64 quo, mod; + I64_udivmod(bigint, divisor, &quo, &mod); + Field(accu, 0) = value_of_uint32(I64_to_int32(quo)); + Field(accu, 1) = value_of_uint32(I64_to_int32(mod)); + } */ + value bigint; /* TODO: fix */ + bigint = *sp++; /* TODO: take accu into account */ + value divisor; + divisor = *sp++; + if (uint63_eq0(divisor)) { + Alloc_small(accu, 2, 1); + Field(accu, 0) = divisor; + Field(accu, 1) = divisor; } - else{ - /* *sp = 2*x+1 --> accu = 2^(shiftby+1)*x */ - accu = (value)((((uint32_t)*sp++)^1) << shiftby); - /* accu = 2^(shiftby+1)*x --> 2^(shifby+1)*x+2*y/2^(31-shiftby)+1 */ - accu = (value)((accu | (((uint32_t)(*sp++)) >> (31-shiftby)))|1); + else { + value quo, mod; + mod = uint63_div21(accu, bigint, divisor, &quo); + Alloc_small(accu, 2, 1); + Field(accu, 0) = quo; + Field(accu, 1) = mod; } Next; } - Instruct (COMPAREINT31) { - /* returns Eq if equal, Lt if accu is less than *sp, Gt otherwise */ - /* assumes Inductive _ : _ := Eq | Lt | Gt */ - print_instr("COMPAREINT31"); - if ((uint32_t)accu == (uint32_t)*sp) { - accu = 1; /* 2*0+1 */ - sp++; - } - else{if ((uint32_t)accu < (uint32_t)(*sp++)) { - accu = 3; /* 2*1+1 */ - } - else{ - accu = 5; /* 2*2+1 */ - }} + Instruct(CHECKLXORINT63) { + print_instr("CHECKLXORINT63"); + CheckInt2(); + accu = uint63_lxor(accu,*sp++); Next; } - - Instruct (HEAD0INT31) { - int r = 0; - uint32_t x; - print_instr("HEAD0INT31"); - x = (uint32_t) accu; - if (!(x & 0xFFFF0000)) { x <<= 16; r += 16; } - if (!(x & 0xFF000000)) { x <<= 8; r += 8; } - if (!(x & 0xF0000000)) { x <<= 4; r += 4; } - if (!(x & 0xC0000000)) { x <<= 2; r += 2; } - if (!(x & 0x80000000)) { x <<=1; r += 1; } - if (!(x & 0x80000000)) { r += 1; } - accu = value_of_uint32(r); + + Instruct(CHECKLORINT63) { + print_instr("CHECKLORINT63"); + CheckInt2(); + accu = uint63_lor(accu,*sp++); Next; } - - Instruct (TAIL0INT31) { - int r = 0; - uint32_t x; - print_instr("TAIL0INT31"); - x = (((uint32_t) accu >> 1) | 0x80000000); - if (!(x & 0xFFFF)) { x >>= 16; r += 16; } - if (!(x & 0x00FF)) { x >>= 8; r += 8; } - if (!(x & 0x000F)) { x >>= 4; r += 4; } - if (!(x & 0x0003)) { x >>= 2; r += 2; } - if (!(x & 0x0001)) { x >>=1; r += 1; } - if (!(x & 0x0001)) { r += 1; } - accu = value_of_uint32(r); + + Instruct(CHECKLANDINT63) { + print_instr("CHECKLANDINT63"); + CheckInt2(); + accu = uint63_land(accu,*sp++); Next; } - Instruct (ISCONST) { - /* Branches if the accu does not contain a constant - (i.e., a non-block value) */ - print_instr("ISCONST"); - if ((accu & 1) == 0) /* last bit is 0 -> it is a block */ - pc += *pc; - else - pc++; + Instruct(CHECKLSLINT63) { + print_instr("CHECKLSLINT63"); + CheckInt2(); + value p = *sp++; + accu = uint63_lsl(accu,p); Next; + } + Instruct(CHECKLSRINT63) { + print_instr("CHECKLSRINT63"); + CheckInt2(); + value p = *sp++; + accu = uint63_lsr(accu,p); + Next; } - Instruct (ARECONST) { - /* Branches if the n first values on the stack are not - all constansts */ - print_instr("ARECONST"); - int i, n, ok; - ok = 1; - n = *pc++; - for(i=0; i < n; i++) { - if ((sp[i] & 1) == 0) { - ok = 0; - break; - } + Instruct(CHECKLSLINT63CONST1) { + print_instr("CHECKLSLINT63CONST1"); + if (Is_uint63(accu)) { + pc++; + accu = uint63_lsl1(accu); + Next; + } else { + *--sp = uint63_one(); + *--sp = accu; + accu = Field(coq_global_data, *pc++); + goto apply2; } - if(ok) pc++; else pc += *pc; + } + + Instruct(CHECKLSRINT63CONST1) { + print_instr("CHECKLSRINT63CONST1"); + if (Is_uint63(accu)) { + pc++; + accu = uint63_lsr1(accu); + Next; + } else { + *--sp = uint63_one(); + *--sp = accu; + accu = Field(coq_global_data, *pc++); + goto apply2; + } + } + + Instruct (CHECKADDMULDIVINT63) { + print_instr("CHECKADDMULDIVINT63"); + CheckInt3(); + value x; + value y; + x = *sp++; + y = *sp++; + accu = uint63_addmuldiv(accu,x,y); Next; } - Instruct (COMPINT31) { - /* makes an 31-bit integer out of the accumulator and - the 30 first values of the stack - and put it in the accumulator (the accumulator then the - topmost get to be the heavier bits) */ - print_instr("COMPINT31"); - int i; - /*accu=accu or accu = (value)((unsigned long)1-accu) if bool - is used for the bits */ - for(i=0; i < 30; i++) { - accu = (value) ((((uint32_t)accu-1) << 1) | *sp++); - /* -1 removes the tag bit, << 1 multiplies the value by 2, - | *sp++ pops the last value and add it (no carry involved) - not that it reintroduces a tag bit */ - /* alternative, if bool is used for the bits : - accu = (value) ((((unsigned long)accu) << 1) & !*sp++); */ + Instruct (CHECKEQINT63) { + print_instr("CHECKEQINT63"); + CheckInt2(); + accu = uint63_eq(accu,*sp++) ? coq_true : coq_false; + Next; + } + + Instruct (CHECKLTINT63) { + print_instr("CHECKLTINT63"); + CheckInt2(); + } + Instruct (LTINT63) { + print_instr("LTINT63"); + accu = uint63_lt(accu,*sp++) ? coq_true : coq_false; + Next; + } + + Instruct (CHECKLEINT63) { + print_instr("CHECKLEINT63"); + CheckInt2(); + } + Instruct (LEINT63) { + print_instr("LEINT63"); + accu = uint63_leq(accu,*sp++) ? coq_true : coq_false; + Next; + } + + Instruct (CHECKCOMPAREINT63) { + /* returns Eq if equal, Lt if accu is less than *sp, Gt otherwise */ + /* assumes Inductive _ : _ := Eq | Lt | Gt */ + print_instr("CHECKCOMPAREINT63"); + CheckInt2(); + if (uint63_eq(accu,*sp)) { + accu = coq_Eq; + sp++; } + else accu = uint63_lt(accu,*sp++) ? coq_Lt : coq_Gt; Next; } - Instruct (DECOMPINT31) { - /* builds a block out of a 31-bit integer (from the accumulator), - used before cases */ - int i; - value block; - print_instr("DECOMPINT31"); - Alloc_small(block, 31, 1); // Alloc_small(*, size, tag) - for(i = 30; i >= 0; i--) { - Field(block, i) = (value)(accu & 3); /* two last bits of the accumulator */ - //Field(block, i) = 3; - accu = (value) ((uint32_t)accu >> 1) | 1; /* last bit must be a one */ - }; - accu = block; + Instruct (CHECKHEAD0INT63) { + print_instr("CHECKHEAD0INT63"); + CheckInt1(); + accu = uint63_head0(accu); Next; } - Instruct (ORINT31) { - /* returns the bitwise or */ - print_instr("ORINT31"); - accu = - value_of_uint32((uint32_of_value(accu)) | (uint32_of_value(*sp++))); - Next; + Instruct (CHECKTAIL0INT63) { + print_instr("CHECKTAIL0INT63"); + CheckInt1(); + accu = uint63_tail0(accu); + Next; } - Instruct (ANDINT31) { - /* returns the bitwise and */ - print_instr("ANDINT31"); - accu = - value_of_uint32((uint32_of_value(accu)) & (uint32_of_value(*sp++))); + Instruct (ISINT){ + print_instr("ISINT"); + accu = (Is_uint63(accu)) ? coq_true : coq_false; Next; } - Instruct (XORINT31) { - /* returns the bitwise xor */ - print_instr("XORINT31"); - accu = - value_of_uint32((uint32_of_value(accu)) ^ (uint32_of_value(*sp++))); + Instruct (AREINT2){ + print_instr("AREINT2"); + accu = (Is_uint63(accu) && Is_uint63(sp[0])) ? coq_true : coq_false; + sp++; Next; } - /* /spiwack */ - - /* Debugging and machine control */ diff --git a/kernel/byterun/coq_uint63_emul.h b/kernel/byterun/coq_uint63_emul.h new file mode 100644 index 0000000000..5499f124a2 --- /dev/null +++ b/kernel/byterun/coq_uint63_emul.h @@ -0,0 +1,97 @@ +# pragma once + +# include <caml/callback.h> +# include <caml/stack.h> + + +#define Is_uint63(v) (Tag_val(v) == Custom_tag) + +# define DECLARE_NULLOP(name) \ +value uint63_##name() { \ + static value* cb = 0; \ + CAMLparam0(); \ + if (!cb) cb = caml_named_value("uint63 " #name); \ + CAMLreturn(*cb); \ + } + +# define DECLARE_UNOP(name) \ +value uint63_##name(value x) { \ + static value* cb = 0; \ + CAMLparam1(x); \ + if (!cb) cb = caml_named_value("uint63 " #name); \ + CAMLreturn(caml_callback(*cb, x)); \ + } + +# define DECLARE_PREDICATE(name) \ +value uint63_##name(value x) { \ + static value* cb = 0; \ + CAMLparam1(x); \ + if (!cb) cb = caml_named_value("uint63 " #name); \ + CAMLreturn(Int_val(caml_callback(*cb, x))); \ + } + +# define DECLARE_BINOP(name) \ +value uint63_##name(value x, value y) { \ + static value* cb = 0; \ + CAMLparam2(x, y); \ + if (!cb) cb = caml_named_value("uint63 " #name); \ + CAMLreturn(caml_callback2(*cb, x, y)); \ + } + +# define DECLARE_RELATION(name) \ +value uint63_##name(value x, value y) { \ + static value* cb = 0; \ + CAMLparam2(x, y); \ + if (!cb) cb = caml_named_value("uint63 " #name); \ + CAMLreturn(Int_val(caml_callback2(*cb, x, y))); \ + } + +# define DECLARE_TEROP(name) \ +value uint63_##name(value x, value y, value z) { \ + static value* cb = 0; \ + CAMLparam3(x, y, z); \ + if (!cb) cb = caml_named_value("uint63 " #name); \ + CAMLreturn(caml_callback3(*cb, x, y, z)); \ + } + + +DECLARE_NULLOP(one) +DECLARE_BINOP(add) +DECLARE_BINOP(addcarry) +DECLARE_TEROP(addmuldiv) +DECLARE_BINOP(div) +DECLARE_TEROP(div21_ml) +DECLARE_RELATION(eq) +DECLARE_PREDICATE(eq0) +DECLARE_UNOP(head0) +DECLARE_BINOP(land) +DECLARE_RELATION(leq) +DECLARE_BINOP(lor) +DECLARE_BINOP(lsl) +DECLARE_UNOP(lsl1) +DECLARE_BINOP(lsr) +DECLARE_UNOP(lsr1) +DECLARE_RELATION(lt) +DECLARE_BINOP(lxor) +DECLARE_BINOP(mod) +DECLARE_BINOP(mul) +DECLARE_BINOP(mulc_ml) +DECLARE_BINOP(sub) +DECLARE_BINOP(subcarry) +DECLARE_UNOP(tail0) + +value uint63_div21(value x, value y, value z, value* q) { + CAMLparam3(x, y, z); + CAMLlocal1(qr); + qr = uint63_div21_ml(x, y, z); + *q = Field(qr, 0); + CAMLreturn(Field(qr, 1)); +} + +value uint63_mulc(value x, value y, value* h) { + CAMLparam2(x, y); + CAMLlocal1(hl); + hl = uint63_mulc_ml(x, y); + *h = Field(hl, 0); + CAMLreturn(Field(hl, 1)); +} diff --git a/kernel/byterun/coq_uint63_native.h b/kernel/byterun/coq_uint63_native.h new file mode 100644 index 0000000000..92f4dc79bc --- /dev/null +++ b/kernel/byterun/coq_uint63_native.h @@ -0,0 +1,125 @@ +#define Is_uint63(v) (Is_long(v)) + +#define uint63_of_value(val) ((uint64_t)(val) >> 1) + +/* 2^63 * y + x as a value */ +//#define Val_intint(x,y) ((value)(((uint64_t)(x)) << 1 + ((uint64_t)(y) << 64))) + +#define uint63_zero ((value) 1) /* 2*0 + 1 */ +#define uint63_one() ((value) 3) /* 2*1 + 1 */ + +#define uint63_eq(x,y) ((x) == (y)) +#define uint63_eq0(x) ((x) == (uint64_t)1) +#define uint63_lt(x,y) ((uint64_t) (x) < (uint64_t) (y)) +#define uint63_leq(x,y) ((uint64_t) (x) <= (uint64_t) (y)) + +#define uint63_add(x,y) ((value)((uint64_t) (x) + (uint64_t) (y) - 1)) +#define uint63_addcarry(x,y) ((value)((uint64_t) (x) + (uint64_t) (y) + 1)) +#define uint63_sub(x,y) ((value)((uint64_t) (x) - (uint64_t) (y) + 1)) +#define uint63_subcarry(x,y) ((value)((uint64_t) (x) - (uint64_t) (y) - 1)) +#define uint63_mul(x,y) (Val_long(uint63_of_value(x) * uint63_of_value(y))) +#define uint63_div(x,y) (Val_long(uint63_of_value(x) / uint63_of_value(y))) +#define uint63_mod(x,y) (Val_long(uint63_of_value(x) % uint63_of_value(y))) + +#define uint63_lxor(x,y) ((value)(((uint64_t)(x) ^ (uint64_t)(y)) | 1)) +#define uint63_lor(x,y) ((value)((uint64_t)(x) | (uint64_t)(y))) +#define uint63_land(x,y) ((value)((uint64_t)(x) & (uint64_t)(y))) + +/* TODO: is + or | better? OCAML uses + */ +/* TODO: is - or ^ better? */ +#define uint63_lsl(x,y) ((y) < (uint64_t) 127 ? ((value)((((uint64_t)(x)-1) << (uint63_of_value(y))) | 1)) : uint63_zero) +#define uint63_lsr(x,y) ((y) < (uint64_t) 127 ? ((value)(((uint64_t)(x) >> (uint63_of_value(y))) | 1)) : uint63_zero) +#define uint63_lsl1(x) ((value)((((uint64_t)(x)-1) << 1) +1)) +#define uint63_lsr1(x) ((value)(((uint64_t)(x) >> 1) |1)) + +/* addmuldiv(p,x,y) = x * 2^p + y / 2 ^ (63 - p) */ +/* (modulo 2^63) for p <= 63 */ +value uint63_addmuldiv(uint64_t p, uint64_t x, uint64_t y) { + uint64_t shiftby = uint63_of_value(p); + value r = (value)((uint64_t)(x^1) << shiftby); + r |= ((uint64_t)y >> (63-shiftby)) | 1; + return r; +} + +value uint63_head0(uint64_t x) { + int r = 0; + if (!(x & 0xFFFFFFFF00000000)) { x <<= 32; r += 32; } + if (!(x & 0xFFFF000000000000)) { x <<= 16; r += 16; } + if (!(x & 0xFF00000000000000)) { x <<= 8; r += 8; } + if (!(x & 0xF000000000000000)) { x <<= 4; r += 4; } + if (!(x & 0xC000000000000000)) { x <<= 2; r += 2; } + if (!(x & 0x8000000000000000)) { x <<=1; r += 1; } + return Val_int(r); +} + +value uint63_tail0(value x) { + int r = 0; + x = (uint64_t)x >> 1; + if (!(x & 0xFFFFFFFF)) { x >>= 32; r += 32; } + if (!(x & 0x0000FFFF)) { x >>= 16; r += 16; } + if (!(x & 0x000000FF)) { x >>= 8; r += 8; } + if (!(x & 0x0000000F)) { x >>= 4; r += 4; } + if (!(x & 0x00000003)) { x >>= 2; r += 2; } + if (!(x & 0x00000001)) { x >>=1; r += 1; } + return Val_int(r); +} + +value uint63_mulc(value x, value y, value* h) { + x = (uint64_t)x >> 1; + y = (uint64_t)y >> 1; + uint64_t lx = x & 0xFFFFFFFF; + uint64_t ly = y & 0xFFFFFFFF; + uint64_t hx = x >> 32; + uint64_t hy = y >> 32; + uint64_t hr = hx * hy; + uint64_t lr = lx * ly; + lx *= hy; + ly *= hx; + hr += (lx >> 32) + (ly >> 32); + lx <<= 32; + lr += lx; + if (lr < lx) { hr++; } + ly <<= 32; + lr += ly; + if (lr < ly) { hr++; } + hr = (hr << 1) | (lr >> 63); + *h = Val_int(hr); + return Val_int(lr); +} + +#define lt128(xh,xl,yh,yl) (uint63_lt(xh,yh) || (uint63_eq(xh,yh) && uint63_lt(xl,yl))) +#define le128(xh,xl,yh,yl) (uint63_lt(xh,yh) || (uint63_eq(xh,yh) && uint63_leq(xl,yl))) + +value uint63_div21(value xh, value xl, value y, value* q) { + xh = (uint64_t)xh >> 1; + xl = ((uint64_t)xl >> 1) | ((uint64_t)xh << 63); + xh = (uint64_t)xh >> 1; + uint64_t maskh = 0; + uint64_t maskl = 1; + uint64_t dh = 0; + uint64_t dl = (uint64_t)y >> 1; + int cmp = 1; + while (dh >= 0 && cmp) { + cmp = lt128(dh,dl,xh,xl); + dh = (dh << 1) | (dl >> 63); + dl = dl << 1; + maskh = (maskh << 1) | (maskl >> 63); + maskl = maskl << 1; + } + uint64_t remh = xh; + uint64_t reml = xl; + uint64_t quotient = 0; + while (maskh | maskl) { + if (le128(dh,dl,remh,reml)) { + quotient = quotient | maskl; + if (uint63_lt(reml,dl)) {remh = remh - dh - 1;} else {remh = remh - dh;} + reml = reml - dl; + } + maskl = (maskl >> 1) | (maskh << 63); + maskh = maskh >> 1; + dl = (dl >> 1) | (dh << 63); + dh = dh >> 1; + } + *q = Val_int(quotient); + return Val_int(reml); +} diff --git a/kernel/byterun/coq_values.h b/kernel/byterun/coq_values.h index bb0f0eb5e4..0cf6ccf532 100644 --- a/kernel/byterun/coq_values.h +++ b/kernel/byterun/coq_values.h @@ -28,6 +28,16 @@ /* 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 coq_tag_C1 2 +#define coq_tag_C0 1 +#define coq_tag_pair 1 +#define coq_true Val_int(0) +#define coq_false Val_int(1) +#define coq_Eq Val_int(0) +#define coq_Lt Val_int(1) +#define coq_Gt Val_int(2) + #endif /* _COQ_VALUES_ */ diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml index 196bb16f32..5fec55fea1 100644 --- a/kernel/cClosure.ml +++ b/kernel/cClosure.ml @@ -28,8 +28,9 @@ open Util open Pp open Names open Constr -open Vars +open Declarations open Environ +open Vars open Esubst let stats = ref false @@ -303,6 +304,7 @@ and fterm = | FProd of Name.t * fconstr * constr * fconstr subs | FLetIn of Name.t * fconstr * fconstr * constr * fconstr subs | FEvar of existential * fconstr subs + | FInt of Uint63.t | FLIFT of int * fconstr | FCLOS of constr * fconstr subs | FLOCKED @@ -335,19 +337,22 @@ type clos_infos = { i_flags : reds; i_cache : infos_cache } -type clos_tab = fconstr KeyTable.t +type clos_tab = fconstr constant_def KeyTable.t let info_flags info = info.i_flags let info_env info = info.i_cache.i_env (**********************************************************************) (* The type of (machine) stacks (= lambda-bar-calculus' contexts) *) +type 'a next_native_args = (CPrimitives.arg_kind * 'a) list type stack_member = | Zapp of fconstr array | ZcaseT of case_info * constr * constr array * fconstr subs | Zproj of Projection.Repr.t | Zfix of fconstr * stack + | Zprimitive of CPrimitives.t * pconstant * fconstr list * fconstr next_native_args + (* operator, constr def, arguments already seen (in rev order), next arguments *) | Zshift of int | Zupdate of fconstr @@ -358,7 +363,7 @@ let append_stack v s = if Int.equal (Array.length v) 0 then s else match s with | Zapp l :: s -> Zapp (Array.append v l) :: s - | (ZcaseT _ | Zproj _ | Zfix _ | Zshift _ | Zupdate _) :: _ | [] -> + | (ZcaseT _ | Zproj _ | Zfix _ | Zshift _ | Zupdate _ | Zprimitive _) :: _ | [] -> Zapp v :: s (* Collapse the shifts in the stack *) @@ -366,20 +371,20 @@ let zshift n s = match (n,s) with (0,_) -> s | (_,Zshift(k)::s) -> Zshift(n+k)::s - | (_,(ZcaseT _ | Zproj _ | Zfix _ | Zapp _ | Zupdate _) :: _) | _,[] -> Zshift(n)::s + | (_,(ZcaseT _ | Zproj _ | Zfix _ | Zapp _ | Zupdate _ | Zprimitive _) :: _) | _,[] -> Zshift(n)::s let rec stack_args_size = function | Zapp v :: s -> Array.length v + stack_args_size s | Zshift(_)::s -> stack_args_size s | Zupdate(_)::s -> stack_args_size s - | (ZcaseT _ | Zproj _ | Zfix _) :: _ | [] -> 0 + | (ZcaseT _ | Zproj _ | Zfix _ | Zprimitive _) :: _ | [] -> 0 (* Lifting. Preserves sharing (useful only for cell with norm=Red). lft_fconstr always create a new cell, while lift_fconstr avoids it when the lift is 0. *) let rec lft_fconstr n ft = match ft.term with - | (FInd _|FConstruct _|FFlex(ConstKey _|VarKey _)) -> ft + | (FInd _|FConstruct _|FFlex(ConstKey _|VarKey _)|FInt _) -> ft | FRel i -> {norm=Norm;term=FRel(i+n)} | FLambda(k,tys,f,e) -> {norm=Cstr; term=FLambda(k,tys,f,subs_shft(n,e))} | FFix(fx,e) -> {norm=Cstr; term=FFix(fx,subs_shft(n,e))} @@ -411,7 +416,7 @@ let compact_stack head stk = (** The stack contains [Zupdate] marks only if in sharing mode *) let _ = update ~share:true m h'.norm h'.term in strip_rec depth s - | ((ZcaseT _ | Zproj _ | Zfix _ | Zapp _) :: _ | []) as stk -> zshift depth stk + | ((ZcaseT _ | Zproj _ | Zfix _ | Zapp _ | Zprimitive _) :: _ | []) as stk -> zshift depth stk in strip_rec 0 stk @@ -447,6 +452,7 @@ let mk_clos e t = | Meta _ | Sort _ -> { norm = Norm; term = FAtom t } | Ind kn -> { norm = Norm; term = FInd kn } | Construct kn -> { norm = Cstr; term = FConstruct kn } + | Int i -> {norm = Cstr; term = FInt i} | (CoFix _|Lambda _|Fix _|Prod _|Evar _|App _|Case _|Cast _|LetIn _|Proj _) -> {norm = Red; term = FCLOS(t,e)} @@ -465,32 +471,34 @@ let mk_clos_vect env v = match v with let ref_value_cache ({ i_cache = cache; _ }) tab ref = try - Some (KeyTable.find tab ref) + KeyTable.find tab ref with Not_found -> - try - let body = - match ref with - | RelKey n -> - let open! Context.Rel.Declaration in - let i = n - 1 in - let (d, _) = - try Range.get cache.i_env.env_rel_context.env_rel_map i - with Invalid_argument _ -> raise Not_found - in - begin match d with - | LocalAssum _ -> raise Not_found - | LocalDef (_, t, _) -> lift n t - end - | VarKey id -> assoc_defined id cache.i_env - | ConstKey cst -> constant_value_in cache.i_env cst + let v = + try + let body = + match ref with + | RelKey n -> + let open! Context.Rel.Declaration in + let i = n - 1 in + let (d, _) = + try Range.get cache.i_env.env_rel_context.env_rel_map i + with Invalid_argument _ -> raise Not_found + in + begin match d with + | LocalAssum _ -> raise Not_found + | LocalDef (_, t, _) -> lift n t + end + | VarKey id -> assoc_defined id cache.i_env + | ConstKey cst -> constant_value_in cache.i_env cst + in + Def (inject body) + with + | NotEvaluableConst (IsPrimitive op) (* Const *) -> Primitive op + | Not_found (* List.assoc *) + | NotEvaluableConst _ (* Const *) + -> Undef None in - let v = inject body in - KeyTable.add tab ref v; - Some v - with - | Not_found (* List.assoc *) - | NotEvaluableConst _ (* Const *) - -> None + KeyTable.add tab ref v; v (* The inverse of mk_clos: move back to constr *) let rec to_constr lfts v = @@ -559,6 +567,10 @@ let rec to_constr lfts v = let subs = comp_subs lfts env in mkEvar(ev,Array.map (fun a -> subst_constr subs a) args) | FLIFT (k,a) -> to_constr (el_shft k lfts) a + + | FInt i -> + Constr.mkInt i + | FCLOS (t,env) -> if is_subs_id env && is_lift_id lfts then t else @@ -607,6 +619,10 @@ let rec zip m stk = | Zupdate(rf)::s -> (** The stack contains [Zupdate] marks only if in sharing mode *) zip (update ~share:true rf m.norm m.term) s + | Zprimitive(_op,c,rargs,kargs)::s -> + let args = List.rev_append rargs (m::List.map snd kargs) in + let f = {norm = Red;term = FFlex (ConstKey c)} in + zip {norm=neutr m.norm; term = FApp (f, Array.of_list args)} s let fapp_stack (m,stk) = zip m stk @@ -628,7 +644,7 @@ let strip_update_shift_app_red head stk = | Zupdate(m)::s -> (** The stack contains [Zupdate] marks only if in sharing mode *) strip_rec rstk (update ~share:true m h.norm h.term) depth s - | ((ZcaseT _ | Zproj _ | Zfix _) :: _ | []) as stk -> + | ((ZcaseT _ | Zproj _ | Zfix _ | Zprimitive _) :: _ | []) as stk -> (depth,List.rev rstk, stk) in strip_rec [] head 0 stk @@ -656,7 +672,7 @@ let get_nth_arg head n stk = | Zupdate(m)::s -> (** The stack contains [Zupdate] mark only if in sharing mode *) strip_rec rstk (update ~share:true m h.norm h.term) n s - | ((ZcaseT _ | Zproj _ | Zfix _) :: _ | []) as s -> (None, List.rev rstk @ s) in + | ((ZcaseT _ | Zproj _ | Zfix _ | Zprimitive _) :: _ | []) as s -> (None, List.rev rstk @ s) in strip_rec [] head n stk (* Beta reduction: look for an applied argument in the stack. @@ -678,24 +694,70 @@ let rec get_args n tys f e = function else (* more lambdas *) let etys = List.skipn na tys in get_args (n-na) etys f (subs_cons(l,e)) s - | ((ZcaseT _ | Zproj _ | Zfix _) :: _ | []) as stk -> + | ((ZcaseT _ | Zproj _ | Zfix _ | Zprimitive _) :: _ | []) as stk -> (Inr {norm=Cstr;term=FLambda(n,tys,f,e)}, stk) (* Eta expansion: add a reference to implicit surrounding lambda at end of stack *) let rec eta_expand_stack = function | (Zapp _ | Zfix _ | ZcaseT _ | Zproj _ - | Zshift _ | Zupdate _ as e) :: s -> + | Zshift _ | Zupdate _ | Zprimitive _ as e) :: s -> e :: eta_expand_stack s | [] -> [Zshift 1; Zapp [|{norm=Norm; term= FRel 1}|]] +(* Get the arguments of a native operator *) +let rec skip_native_args rargs nargs = + match nargs with + | (kd, a) :: nargs' -> + if kd = CPrimitives.Kwhnf then rargs, nargs + else skip_native_args (a::rargs) nargs' + | [] -> rargs, [] + +let get_native_args op c stk = + let kargs = CPrimitives.kind op in + let rec get_args rnargs kargs args = + match kargs, args with + | kd::kargs, a::args -> get_args ((kd,a)::rnargs) kargs args + | _, _ -> rnargs, kargs, args in + let rec strip_rec rnargs h depth kargs = function + | Zshift k :: s -> + strip_rec (List.map (fun (kd,f) -> kd,lift_fconstr k f) rnargs) + (lift_fconstr k h) (depth+k) kargs s + | Zapp args :: s' -> + begin match get_args rnargs kargs (Array.to_list args) with + | rnargs, [], [] -> + (skip_native_args [] (List.rev rnargs), s') + | rnargs, [], eargs -> + (skip_native_args [] (List.rev rnargs), + Zapp (Array.of_list eargs) :: s') + | rnargs, kargs, _ -> + strip_rec rnargs {norm = h.norm;term=FApp(h, args)} depth kargs s' + end + | Zupdate(m) :: s -> + strip_rec rnargs (update ~share:true m h.norm h.term) depth kargs s + | (Zprimitive _ | ZcaseT _ | Zproj _ | Zfix _) :: _ | [] -> assert false + in strip_rec [] {norm = Red;term = FFlex(ConstKey c)} 0 kargs stk + +let get_native_args1 op c stk = + match get_native_args op c stk with + | ((rargs, (kd,a):: nargs), stk) -> + assert (kd = CPrimitives.Kwhnf); + (rargs, a, nargs, stk) + | _ -> assert false + +let check_native_args op stk = + let nargs = CPrimitives.arity op in + let rargs = stack_args_size stk in + nargs <= rargs + + (* Iota reduction: extract the arguments to be passed to the Case branches *) let rec reloc_rargs_rec depth = function | Zapp args :: s -> Zapp (lift_fconstr_vect depth args) :: reloc_rargs_rec depth s | Zshift(k)::s -> if Int.equal k depth then s else reloc_rargs_rec (depth-k) s - | ((ZcaseT _ | Zproj _ | Zfix _ | Zupdate _) :: _ | []) as stk -> stk + | ((ZcaseT _ | Zproj _ | Zfix _ | Zupdate _ | Zprimitive _) :: _ | []) as stk -> stk let reloc_rargs depth stk = if Int.equal depth 0 then stk else reloc_rargs_rec depth stk @@ -712,7 +774,7 @@ let rec try_drop_parameters depth n = function | [] -> if Int.equal n 0 then [] else raise Not_found - | (ZcaseT _ | Zproj _ | Zfix _ | Zupdate _) :: _ -> assert false + | (ZcaseT _ | Zproj _ | Zfix _ | Zupdate _ | Zprimitive _) :: _ -> assert false (* strip_update_shift_app only produces Zapp and Zshift items *) let drop_parameters depth n argstk = @@ -757,7 +819,7 @@ let rec project_nth_arg n = function let q = Array.length args in if n >= q then project_nth_arg (n - q) s else (* n < q *) args.(n) - | (ZcaseT _ | Zproj _ | Zfix _ | Zupdate _ | Zshift _) :: _ | [] -> assert false + | (ZcaseT _ | Zproj _ | Zfix _ | Zupdate _ | Zshift _ | Zprimitive _) :: _ | [] -> assert false (* After drop_parameters we have a purely applicative stack *) @@ -814,7 +876,7 @@ let rec knh info m stk = (* cases where knh stops *) | (FFlex _|FLetIn _|FConstruct _|FEvar _| - FCoFix _|FLambda _|FRel _|FAtom _|FInd _|FProd _) -> + FCoFix _|FLambda _|FRel _|FAtom _|FInd _|FProd _|FInt _) -> (m, stk) (* The same for pure terms *) @@ -828,7 +890,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 { norm = Red; term = FProj (p, mk_clos e c) } stk - | (Ind _|Const _|Construct _|Var _|Meta _ | Sort _) -> (mk_clos e t, stk) + | (Ind _|Const _|Construct _|Var _|Meta _ | Sort _ | Int _) -> (mk_clos e t, stk) | CoFix cfx -> { norm = Cstr; term = FCoFix (cfx,e) }, stk | Lambda _ -> { norm = Cstr; term = mk_lambda e t }, stk | Prod (n, t, c) -> @@ -837,6 +899,162 @@ and knht info e t stk = { norm = Red; term = FLetIn (n, mk_clos e b, mk_clos e t, c, e) }, stk | Evar ev -> { norm = Red; term = FEvar (ev, e) }, stk +let inject c = mk_clos (subs_id 0) c + +(************************************************************************) +(* Reduction of Native operators *) + +open Primred + +module FNativeEntries = + struct + type elem = fconstr + type args = fconstr array + type evd = unit + + let get = Array.get + + let get_int () e = + match [@ocaml.warning "-4"] e.term with + | FInt i -> i + | _ -> raise Primred.NativeDestKO + + let dummy = {norm = Norm; term = FRel 0} + + let current_retro = ref Retroknowledge.empty + let defined_int = ref false + let fint = ref dummy + + let init_int retro = + match retro.Retroknowledge.retro_int63 with + | Some c -> + defined_int := true; + fint := { norm = Norm; term = FFlex (ConstKey (Univ.in_punivs c)) } + | None -> defined_int := false + + let defined_bool = ref false + let ftrue = ref dummy + let ffalse = ref dummy + + let init_bool retro = + match retro.Retroknowledge.retro_bool with + | Some (ct,cf) -> + defined_bool := true; + ftrue := { norm = Cstr; term = FConstruct (Univ.in_punivs ct) }; + ffalse := { norm = Cstr; term = FConstruct (Univ.in_punivs cf) } + | None -> defined_bool :=false + + let defined_carry = ref false + let fC0 = ref dummy + let fC1 = ref dummy + + let init_carry retro = + match retro.Retroknowledge.retro_carry with + | Some(c0,c1) -> + defined_carry := true; + fC0 := { norm = Cstr; term = FConstruct (Univ.in_punivs c0) }; + fC1 := { norm = Cstr; term = FConstruct (Univ.in_punivs c1) } + | None -> defined_carry := false + + let defined_pair = ref false + let fPair = ref dummy + + let init_pair retro = + match retro.Retroknowledge.retro_pair with + | Some c -> + defined_pair := true; + fPair := { norm = Cstr; term = FConstruct (Univ.in_punivs c) } + | None -> defined_pair := false + + let defined_cmp = ref false + let fEq = ref dummy + let fLt = ref dummy + let fGt = ref dummy + + let init_cmp retro = + match retro.Retroknowledge.retro_cmp with + | Some (cEq, cLt, cGt) -> + defined_cmp := true; + fEq := { norm = Cstr; term = FConstruct (Univ.in_punivs cEq) }; + fLt := { norm = Cstr; term = FConstruct (Univ.in_punivs cLt) }; + fGt := { norm = Cstr; term = FConstruct (Univ.in_punivs cGt) } + | None -> defined_cmp := false + + let defined_refl = ref false + + let frefl = ref dummy + + let init_refl retro = + match retro.Retroknowledge.retro_refl with + | Some crefl -> + defined_refl := true; + frefl := { norm = Cstr; term = FConstruct (Univ.in_punivs crefl) } + | None -> defined_refl := false + + let init env = + current_retro := env.retroknowledge; + init_int !current_retro; + init_bool !current_retro; + init_carry !current_retro; + init_pair !current_retro; + init_cmp !current_retro; + init_refl !current_retro + + let check_env env = + if not (!current_retro == env.retroknowledge) then init env + + let check_int env = + check_env env; + assert (!defined_int) + + let check_bool env = + check_env env; + assert (!defined_bool) + + let check_carry env = + check_env env; + assert (!defined_carry && !defined_int) + + let check_pair env = + check_env env; + assert (!defined_pair && !defined_int) + + let check_cmp env = + check_env env; + assert (!defined_cmp) + + let mkInt env i = + check_int env; + { norm = Norm; term = FInt i } + + let mkBool env b = + check_bool env; + if b then !ftrue else !ffalse + + let mkCarry env b e = + check_carry env; + {norm = Cstr; + term = FApp ((if b then !fC1 else !fC0),[|!fint;e|])} + + let mkIntPair env e1 e2 = + check_pair env; + { norm = Cstr; term = FApp(!fPair, [|!fint;!fint;e1;e2|]) } + + let mkLt env = + check_cmp env; + !fLt + + let mkEq env = + check_cmp env; + !fEq + + let mkGt env = + check_cmp env; + !fGt + + end + +module FredNative = RedNative(FNativeEntries) (************************************************************************) @@ -849,16 +1067,21 @@ let rec knr info tab m stk = | Inr lam, s -> (lam,s)) | FFlex(ConstKey (kn,_ as c)) when red_set info.i_flags (fCONST kn) -> (match ref_value_cache info tab (ConstKey c) with - Some v -> kni info tab v stk - | None -> (set_norm m; (m,stk))) + | Def v -> kni info tab v stk + | Primitive op when check_native_args op stk -> + let rargs, a, nargs, stk = get_native_args1 op c stk in + kni info tab a (Zprimitive(op,c,rargs,nargs)::stk) + | Undef _ | OpaqueDef _ | Primitive _ -> (set_norm m; (m,stk))) | FFlex(VarKey id) when red_set info.i_flags (fVAR id) -> (match ref_value_cache info tab (VarKey id) with - Some v -> kni info tab v stk - | None -> (set_norm m; (m,stk))) + | Def v -> kni info tab v stk + | Primitive _ -> assert false + | OpaqueDef _ | Undef _ -> (set_norm m; (m,stk))) | FFlex(RelKey k) when red_set info.i_flags fDELTA -> (match ref_value_cache info tab (RelKey k) with - Some v -> kni info tab v stk - | None -> (set_norm m; (m,stk))) + | Def v -> kni info tab v stk + | Primitive _ -> assert false + | OpaqueDef _ | Undef _ -> (set_norm m; (m,stk))) | FConstruct((_ind,c),_u) -> let use_match = red_set info.i_flags fMATCH in let use_fix = red_set info.i_flags fFIX in @@ -884,13 +1107,32 @@ let rec knr info tab m stk = | (_, args, (((ZcaseT _|Zproj _)::_) as stk')) -> let (fxe,fxbd) = contract_fix_vect m.term in knit info tab fxe fxbd (args@stk') - | (_,args, ((Zapp _ | Zfix _ | Zshift _ | Zupdate _) :: _ | [] as s)) -> (m,args@s)) + | (_,args, ((Zapp _ | Zfix _ | Zshift _ | Zupdate _ | Zprimitive _) :: _ | [] as s)) -> (m,args@s)) | FLetIn (_,v,_,bd,e) when red_set info.i_flags fZETA -> knit info tab (subs_cons([|v|],e)) bd stk | FEvar(ev,env) -> (match info.i_cache.i_sigma ev with Some c -> knit info tab env c stk | None -> (m,stk)) + | FInt _ -> + (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 + begin match nargs with + | [] -> + let args = Array.of_list (List.rev rargs) in + begin match FredNative.red_prim (info_env info) () op args with + | Some m -> kni info tab m s + | None -> + let f = {norm = Whnf; term = FFlex (ConstKey c)} in + let m = {norm = Whnf; term = FApp(f,args)} in + (m,s) + end + | (kd,a)::nargs -> + assert (kd = CPrimitives.Kwhnf); + kni info tab a (Zprimitive(op,c,rargs,nargs)::s) + end + | (_, _, s) -> (m, s)) | FLOCKED | FRel _ | FAtom _ | FFlex (RelKey _ | ConstKey _ | VarKey _) | FInd _ | FApp _ | FProj _ | FFix _ | FCoFix _ | FCaseT _ | FLambda _ | FProd _ | FLetIn _ | FLIFT _ | FCLOS _ -> (m, stk) @@ -927,6 +1169,12 @@ let rec zip_term zfun m stk = zip_term zfun (lift n m) s | Zupdate(_rf)::s -> zip_term zfun m s + | Zprimitive(_,c,rargs, kargs)::s -> + let kargs = List.map (fun (_,a) -> zfun a) kargs in + let args = + List.fold_left (fun args a -> zfun a ::args) (m::kargs) rargs in + let h = mkApp (mkConstU c, Array.of_list args) in + zip_term zfun h s (* Computes the strong normal form of a term. 1- Calls kni @@ -972,7 +1220,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 _ -> term_of_fconstr m + | FApp _ | FCaseT _ | FLIFT _ | FCLOS _ | FInt _ -> term_of_fconstr m (* Initialization and then normalization *) @@ -1015,9 +1263,9 @@ let unfold_reference info tab key = | ConstKey (kn,_) -> if red_set info.i_flags (fCONST kn) then ref_value_cache info tab key - else None + else Undef None | VarKey i -> if red_set info.i_flags (fVAR i) then ref_value_cache info tab key - else None + else Undef None | RelKey _ -> ref_value_cache info tab key diff --git a/kernel/cClosure.mli b/kernel/cClosure.mli index 46be1bb279..bd04677374 100644 --- a/kernel/cClosure.mli +++ b/kernel/cClosure.mli @@ -10,6 +10,7 @@ open Names open Constr +open Declarations open Environ open Esubst @@ -117,6 +118,7 @@ type fterm = | FProd of Name.t * fconstr * constr * fconstr subs | FLetIn of Name.t * fconstr * fconstr * constr * fconstr subs | FEvar of existential * fconstr subs + | FInt of Uint63.t | FLIFT of int * fconstr | FCLOS of constr * fconstr subs | FLOCKED @@ -124,12 +126,15 @@ type fterm = (*********************************************************************** s A [stack] is a context of arguments, arguments are pushed by [append_stack] one array at a time *) +type 'a next_native_args = (CPrimitives.arg_kind * 'a) list type stack_member = | Zapp of fconstr array | ZcaseT of case_info * constr * constr array * fconstr subs | Zproj of Projection.Repr.t | Zfix of fconstr * stack + | Zprimitive of CPrimitives.t * pconstant * fconstr list * fconstr next_native_args + (* operator, constr def, reduced arguments rev, next arguments *) | Zshift of int | Zupdate of fconstr @@ -138,6 +143,10 @@ and stack = stack_member list val empty_stack : stack val append_stack : fconstr array -> stack -> stack +val check_native_args : CPrimitives.t -> stack -> bool +val get_native_args1 : CPrimitives.t -> pconstant -> stack -> + fconstr list * fconstr * fconstr next_native_args * stack + val stack_args_size : stack -> int val eta_expand_stack : stack -> stack @@ -201,7 +210,7 @@ val eta_expand_ind_stack : env -> inductive -> fconstr -> stack -> (** Conversion auxiliary functions to do step by step normalisation *) (** [unfold_reference] unfolds references in a [fconstr] *) -val unfold_reference : clos_infos -> clos_tab -> table_key -> fconstr option +val unfold_reference : clos_infos -> clos_tab -> table_key -> fconstr constant_def (*********************************************************************** i This is for lazy debug *) diff --git a/kernel/cPrimitives.ml b/kernel/cPrimitives.ml index 5b91a9b572..da5c4fb07b 100644 --- a/kernel/cPrimitives.ml +++ b/kernel/cPrimitives.ml @@ -9,85 +9,147 @@ (************************************************************************) type t = - | Int31head0 - | Int31tail0 - | Int31add - | Int31sub - | Int31mul - | Int31div - | Int31mod -(* - | Int31lsr - | Int31lsl - *) - | Int31land - | Int31lor - | Int31lxor - | Int31addc - | Int31subc - | Int31addcarryc - | Int31subcarryc - | Int31mulc - | Int31diveucl - | Int31div21 - | Int31addmuldiv - | Int31eq - | Int31lt - | Int31le - | Int31compare + | Int63head0 + | Int63tail0 + | Int63add + | Int63sub + | Int63mul + | Int63div + | Int63mod + | Int63lsr + | Int63lsl + | Int63land + | Int63lor + | Int63lxor + | Int63addc + | Int63subc + | Int63addCarryC + | Int63subCarryC + | Int63mulc + | Int63diveucl + | Int63div21 + | Int63addMulDiv + | Int63eq + | Int63lt + | Int63le + | Int63compare + +let equal (p1 : t) (p2 : t) = + p1 == p2 let hash = function - | Int31head0 -> 1 - | Int31tail0 -> 2 - | Int31add -> 3 - | Int31sub -> 4 - | Int31mul -> 5 - | Int31div -> 6 - | Int31mod -> 7 -(* - | Int31lsr -> 8 - | Int31lsl -> 9 - *) - | Int31land -> 10 - | Int31lor -> 11 - | Int31lxor -> 12 - | Int31addc -> 13 - | Int31subc -> 14 - | Int31addcarryc -> 15 - | Int31subcarryc -> 16 - | Int31mulc -> 17 - | Int31diveucl -> 18 - | Int31div21 -> 19 - | Int31addmuldiv -> 20 - | Int31eq -> 21 - | Int31lt -> 22 - | Int31le -> 23 - | Int31compare -> 24 + | Int63head0 -> 1 + | Int63tail0 -> 2 + | Int63add -> 3 + | Int63sub -> 4 + | Int63mul -> 5 + | Int63div -> 6 + | Int63mod -> 7 + | Int63lsr -> 8 + | Int63lsl -> 9 + | Int63land -> 10 + | Int63lor -> 11 + | Int63lxor -> 12 + | Int63addc -> 13 + | Int63subc -> 14 + | Int63addCarryC -> 15 + | Int63subCarryC -> 16 + | Int63mulc -> 17 + | Int63diveucl -> 18 + | Int63div21 -> 19 + | Int63addMulDiv -> 20 + | Int63eq -> 21 + | Int63lt -> 22 + | Int63le -> 23 + | Int63compare -> 24 +(* Should match names in nativevalues.ml *) let to_string = function - | Int31head0 -> "head0" - | Int31tail0 -> "tail0" - | Int31add -> "add" - | Int31sub -> "sub" - | Int31mul -> "mul" - | Int31div -> "div" - | Int31mod -> "mod" -(* - | Int31lsr -> "l_sr" - | Int31lsl -> "l_sl" - *) - | Int31land -> "l_and" - | Int31lor -> "l_or" - | Int31lxor -> "l_xor" - | Int31addc -> "addc" - | Int31subc -> "subc" - | Int31addcarryc -> "addcarryc" - | Int31subcarryc -> "subcarryc" - | Int31mulc -> "mulc" - | Int31diveucl -> "diveucl" - | Int31div21 -> "div21" - | Int31addmuldiv -> "addmuldiv" - | Int31eq -> "eq" - | Int31lt -> "lt" - | Int31le -> "le" - | Int31compare -> "compare" + | Int63head0 -> "head0" + | Int63tail0 -> "tail0" + | Int63add -> "add" + | Int63sub -> "sub" + | Int63mul -> "mul" + | Int63div -> "div" + | Int63mod -> "rem" + | Int63lsr -> "l_sr" + | Int63lsl -> "l_sl" + | Int63land -> "l_and" + | Int63lor -> "l_or" + | Int63lxor -> "l_xor" + | Int63addc -> "addc" + | Int63subc -> "subc" + | Int63addCarryC -> "addCarryC" + | Int63subCarryC -> "subCarryC" + | Int63mulc -> "mulc" + | Int63diveucl -> "diveucl" + | Int63div21 -> "div21" + | Int63addMulDiv -> "addMulDiv" + | Int63eq -> "eq" + | Int63lt -> "lt" + | Int63le -> "le" + | Int63compare -> "compare" + +type arg_kind = + | Kparam (* not needed for the evaluation of the primitive when it reduces *) + | Kwhnf (* need to be reduced in whnf before reducing the primitive *) + | Karg (* no need to be reduced in whnf. example: [v] in [Array.set t i v] *) + +type args_red = arg_kind list + +(* Invariant only argument of type int63 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] + + | 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 + +(** 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 + | PIT_bool -> "bool" + | PIT_carry -> "carry" + | PIT_pair -> "pair" + | PIT_cmp -> "cmp" + +let prim_type_to_string = function + | PT_int63 -> "int63" + +let op_or_type_to_string = function + | OT_op op -> to_string op + | OT_type t -> prim_type_to_string t diff --git a/kernel/cPrimitives.mli b/kernel/cPrimitives.mli index 1e99a69d2f..3f8174bd7b 100644 --- a/kernel/cPrimitives.mli +++ b/kernel/cPrimitives.mli @@ -9,33 +9,62 @@ (************************************************************************) type t = - | Int31head0 - | Int31tail0 - | Int31add - | Int31sub - | Int31mul - | Int31div - | Int31mod -(* - | Int31lsr - | Int31lsl - *) - | Int31land - | Int31lor - | Int31lxor - | Int31addc - | Int31subc - | Int31addcarryc - | Int31subcarryc - | Int31mulc - | Int31diveucl - | Int31div21 - | Int31addmuldiv - | Int31eq - | Int31lt - | Int31le - | Int31compare + | Int63head0 + | Int63tail0 + | Int63add + | Int63sub + | Int63mul + | Int63div + | Int63mod + | Int63lsr + | Int63lsl + | Int63land + | Int63lor + | Int63lxor + | Int63addc + | Int63subc + | Int63addCarryC + | Int63subCarryC + | Int63mulc + | Int63diveucl + | Int63div21 + | Int63addMulDiv + | Int63eq + | Int63lt + | Int63le + | Int63compare + +val equal : t -> t -> bool + +type arg_kind = + | Kparam (* not needed for the elavuation of the primitive*) + | Kwhnf (* need to be reduced in whnf before reducing the primitive *) + | Karg (* no need to be reduced in whnf *) + +type args_red = arg_kind list val hash : t -> int val to_string : t -> string + +val arity : t -> int + +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 + +type op_or_type = + | OT_op of t + | OT_type of prim_type + +val prim_ind_to_string : prim_ind -> string +val op_or_type_to_string : op_or_type -> string diff --git a/kernel/cbytecodes.ml b/kernel/cbytecodes.ml index c63795b295..7570004fe5 100644 --- a/kernel/cbytecodes.ml +++ b/kernel/cbytecodes.ml @@ -16,6 +16,7 @@ open Names open Vmvalues +open Constr module Label = struct @@ -26,7 +27,6 @@ module Label = let reset_label_counter () = counter := no end - type instruction = | Klabel of Label.t | Kacc of int @@ -59,46 +59,9 @@ type instruction = | Ksequence of bytecodes * bytecodes | Kproj of Projection.Repr.t | Kensurestackcapacity of int -(* spiwack: instructions concerning integers *) | Kbranch of Label.t (* jump to label *) - | Kaddint31 (* adds the int31 in the accu - and the one ontop of the stack *) - | Kaddcint31 (* makes the sum and keeps the carry *) - | Kaddcarrycint31 (* sum +1, keeps the carry *) - | Ksubint31 (* subtraction modulo *) - | Ksubcint31 (* subtraction, keeps the carry *) - | Ksubcarrycint31 (* subtraction -1, keeps the carry *) - | Kmulint31 (* multiplication modulo *) - | Kmulcint31 (* multiplication, result in two - int31, for exact computation *) - | Kdiv21int31 (* divides a double size integer - (represented by an int31 in the - accumulator and one on the top of - the stack) by an int31. The result - is a pair of the quotient and the - rest. - If the divisor is 0, it returns - 0. *) - | Kdivint31 (* euclidian division (returns a pair - quotient,rest) *) - | Kaddmuldivint31 (* generic operation for shifting and - cycling. Takes 3 int31 i j and s, - and returns x*2^s+y/(2^(31-s) *) - | Kcompareint31 (* unsigned comparison of int31 - cf COMPAREINT31 in - kernel/byterun/coq_interp.c - for more info *) - | Khead0int31 (* Give the numbers of 0 in head of a in31*) - | Ktail0int31 (* Give the numbers of 0 in tail of a in31 - ie low bits *) - | Kisconst of Label.t (* conditional jump *) - | Kareconst of int*Label.t (* conditional jump *) - | Kcompint31 (* dynamic compilation of int31 *) - | Kdecompint31 (* dynamic decompilation of int31 *) - | Klorint31 (* bitwise operations: or and xor *) - | Klandint31 - | Klxorint31 -(* /spiwack *) + | Kprim of CPrimitives.t * pconstant option + | Kareint of int and bytecodes = instruction list @@ -110,53 +73,6 @@ type fv_elem = type fv = fv_elem array -(* spiwack: this exception is expected to be raised by function expecting - closed terms. *) -exception NotClosed - - -module Fv_elem = -struct -type t = fv_elem - -let compare e1 e2 = match e1, e2 with -| FVnamed id1, FVnamed id2 -> Id.compare id1 id2 -| FVnamed _, (FVrel _ | FVuniv_var _ | FVevar _) -> -1 -| FVrel _, FVnamed _ -> 1 -| FVrel r1, FVrel r2 -> Int.compare r1 r2 -| FVrel _, (FVuniv_var _ | FVevar _) -> -1 -| FVuniv_var i1, FVuniv_var i2 -> Int.compare i1 i2 -| FVuniv_var _i1, (FVnamed _ | FVrel _) -> 1 -| FVuniv_var _i1, FVevar _ -> -1 -| FVevar _, (FVnamed _ | FVrel _ | FVuniv_var _) -> 1 -| FVevar e1, FVevar e2 -> Evar.compare e1 e2 - -end - -module FvMap = Map.Make(Fv_elem) - -(*spiwack: both type have been moved from Cbytegen because I needed then - for the retroknowledge *) -type vm_env = { - size : int; (* longueur de la liste [n] *) - fv_rev : fv_elem list; (* [fvn; ... ;fv1] *) - fv_fwd : int FvMap.t; (* reverse mapping *) - } - - -type comp_env = { - arity : int; (* arity of the current function, 0 if none *) - nb_uni_stack : int ; (* number of universes on the stack, *) - (* universes are always at the bottom. *) - nb_stack : int; (* number of variables on the stack *) - in_stack : int list; (* position in the stack *) - nb_rec : int; (* number of mutually recursive functions *) - pos_rec : instruction list; (* instruction d'acces pour les variables *) - (* de point fix ou de cofix *) - offset : int; - in_env : vm_env ref (* The free variables of the expression *) - } - (* --- Pretty print *) open Pp open Util @@ -228,28 +144,10 @@ let rec pp_instr i = | Kensurestackcapacity size -> str "growstack " ++ int size - | Kaddint31 -> str "addint31" - | Kaddcint31 -> str "addcint31" - | Kaddcarrycint31 -> str "addcarrycint31" - | Ksubint31 -> str "subint31" - | Ksubcint31 -> str "subcint31" - | Ksubcarrycint31 -> str "subcarrycint31" - | Kmulint31 -> str "mulint31" - | Kmulcint31 -> str "mulcint31" - | Kdiv21int31 -> str "div21int31" - | Kdivint31 -> str "divint31" - | Kcompareint31 -> str "compareint31" - | Khead0int31 -> str "head0int31" - | Ktail0int31 -> str "tail0int31" - | Kaddmuldivint31 -> str "addmuldivint31" - | Kisconst lbl -> str "isconst " ++ int lbl - | Kareconst(n,lbl) -> str "areconst " ++ int n ++ spc () ++ int lbl - | Kcompint31 -> str "compint31" - | Kdecompint31 -> str "decompint" - | Klorint31 -> str "lorint31" - | Klandint31 -> str "landint31" - | Klxorint31 -> str "lxorint31" + | Kprim (op, id) -> str (CPrimitives.to_string op) ++ str " " ++ + (match id with Some (id,_u) -> Constant.print id | None -> str "") + | Kareint n -> str "areint " ++ int n and pp_bytecodes c = match c with diff --git a/kernel/cbytecodes.mli b/kernel/cbytecodes.mli index 9c04c166a2..423e7bbe65 100644 --- a/kernel/cbytecodes.mli +++ b/kernel/cbytecodes.mli @@ -12,6 +12,7 @@ open Names open Vmvalues +open Constr module Label : sig @@ -57,48 +58,15 @@ type instruction = | Kproj of Projection.Repr.t | Kensurestackcapacity of int -(** spiwack: instructions concerning integers *) | Kbranch of Label.t (** jump to label, is it needed ? *) - | Kaddint31 (** adds the int31 in the accu - and the one ontop of the stack *) - | Kaddcint31 (** makes the sum and keeps the carry *) - | Kaddcarrycint31 (** sum +1, keeps the carry *) - | Ksubint31 (** subtraction modulo *) - | Ksubcint31 (** subtraction, keeps the carry *) - | Ksubcarrycint31 (** subtraction -1, keeps the carry *) - | Kmulint31 (** multiplication modulo *) - | Kmulcint31 (** multiplication, result in two - int31, for exact computation *) - | Kdiv21int31 (** divides a double size integer - (represented by an int31 in the - accumulator and one on the top of - the stack) by an int31. The result - is a pair of the quotient and the - rest. - If the divisor is 0, it returns - 0. *) - | Kdivint31 (** euclidian division (returns a pair - quotient,rest) *) - | Kaddmuldivint31 (** generic operation for shifting and - cycling. Takes 3 int31 i j and s, - and returns x*2^s+y/(2^(31-s) *) - | Kcompareint31 (** unsigned comparison of int31 - cf COMPAREINT31 in - kernel/byterun/coq_interp.c - for more info *) - | Khead0int31 (** Give the numbers of 0 in head of a in31*) - | Ktail0int31 (** Give the numbers of 0 in tail of a in31 - ie low bits *) - | Kisconst of Label.t (** conditional jump *) - | Kareconst of int*Label.t (** conditional jump *) - | Kcompint31 (** dynamic compilation of int31 *) - | Kdecompint31 (** dynamix decompilation of int31 *) - | Klorint31 (** bitwise operations: or and xor *) - | Klandint31 - | Klxorint31 + | Kprim of CPrimitives.t * pconstant option + + | Kareint of int and bytecodes = instruction list +val pp_bytecodes : bytecodes -> Pp.t + type fv_elem = FVnamed of Id.t | FVrel of int @@ -107,34 +75,4 @@ type fv_elem = type fv = fv_elem array - -(** spiwack: this exception is expected to be raised by function expecting - closed terms. *) -exception NotClosed - -module FvMap : Map.S with type key = fv_elem - -(*spiwack: both type have been moved from Cbytegen because I needed them - for the retroknowledge *) -type vm_env = { - size : int; (** length of the list [n] *) - fv_rev : fv_elem list; (** [fvn; ... ;fv1] *) - fv_fwd : int FvMap.t; (** reverse mapping *) - } - - -type comp_env = { - arity : int; (* arity of the current function, 0 if none *) - nb_uni_stack : int ; (** number of universes on the stack *) - nb_stack : int; (** number of variables on the stack *) - in_stack : int list; (** position in the stack *) - nb_rec : int; (** number of mutually recursive functions *) - (** (= nbr) *) - pos_rec : instruction list; (** instruction d'acces pour les variables *) - (** de point fix ou de cofix *) - offset : int; - in_env : vm_env ref (** the variables that are accessed *) - } - -val pp_bytecodes : bytecodes -> Pp.t val pp_fv_elem : fv_elem -> Pp.t diff --git a/kernel/cbytegen.ml b/kernel/cbytegen.ml index 73620ae578..b95a940c14 100644 --- a/kernel/cbytegen.ml +++ b/kernel/cbytegen.ml @@ -17,7 +17,6 @@ open Names open Vmvalues open Cbytecodes open Cemitcodes -open Cinstr open Clambda open Constr open Declarations @@ -97,6 +96,48 @@ open Environ (* In Cfxe_t accumulators, we need to store [fcofixi] for testing *) (* conversion of cofixpoints (which is intentional). *) +module Fv_elem = +struct +type t = fv_elem + +let compare e1 e2 = match e1, e2 with +| FVnamed id1, FVnamed id2 -> Id.compare id1 id2 +| FVnamed _, (FVrel _ | FVuniv_var _ | FVevar _) -> -1 +| FVrel _, FVnamed _ -> 1 +| FVrel r1, FVrel r2 -> Int.compare r1 r2 +| FVrel _, (FVuniv_var _ | FVevar _) -> -1 +| FVuniv_var i1, FVuniv_var i2 -> Int.compare i1 i2 +| FVuniv_var _, (FVnamed _ | FVrel _) -> 1 +| FVuniv_var _, FVevar _ -> -1 +| FVevar _, (FVnamed _ | FVrel _ | FVuniv_var _) -> 1 +| FVevar e1, FVevar e2 -> Evar.compare e1 e2 + +end + +module FvMap = Map.Make(Fv_elem) + +(*spiwack: both type have been moved from Cbytegen because I needed then + for the retroknowledge *) +type vm_env = { + size : int; (* longueur de la liste [n] *) + fv_rev : fv_elem list; (* [fvn; ... ;fv1] *) + fv_fwd : int FvMap.t; (* reverse mapping *) + } + + +type comp_env = { + arity : int; (* arity of the current function, 0 if none *) + nb_uni_stack : int ; (* number of universes on the stack, *) + (* universes are always at the bottom. *) + nb_stack : int; (* number of variables on the stack *) + in_stack : int list; (* position in the stack *) + nb_rec : int; (* number of mutually recursive functions *) + pos_rec : instruction list; (* instruction d'acces pour les variables *) + (* de point fix ou de cofix *) + offset : int; + in_env : vm_env ref (* The free variables of the expression *) + } + module Config = struct let stack_threshold = 256 (* see byterun/coq_memory.h *) let stack_safety_margin = 15 @@ -391,7 +432,6 @@ let init_fun_code () = fun_code := [] (* Compilation of constructors and inductive types *) -(* Inv: arity > 0 *) (* If [tag] hits the OCaml limitation for non constant constructors, we switch to @@ -437,7 +477,7 @@ let comp_app comp_fun comp_arg cenv f args sz cont = comp_fun cenv f (sz + nargs) (Kappterm(nargs, k + nargs) :: (discard_dead_code cont))) | None -> - if nargs < 4 then + if nargs <= 4 then comp_args comp_arg cenv args sz (Kpush :: (comp_fun cenv f (sz+nargs) (Kapply nargs :: cont))) else @@ -476,15 +516,6 @@ let rec get_alias env kn = | BCalias kn' -> get_alias env kn' | _ -> kn) -(* spiwack: additional function which allow different part of compilation of the - 31-bit integers *) - -let make_areconst n else_lbl cont = - if n <= 0 then - cont - else - Kareconst (n, else_lbl)::cont - (* sz is the size of the local stack *) let rec compile_lam env cenv lam sz cont = set_max_stack_size sz; @@ -495,6 +526,8 @@ let rec compile_lam env cenv lam sz cont = | Lval v -> compile_structured_constant cenv (Const_val v) sz cont + | Luint i -> compile_structured_constant cenv (Const_uint i) sz cont + | Lproj (p,arg) -> compile_lam env cenv arg sz (Kproj p :: cont) @@ -629,6 +662,17 @@ let rec compile_lam env cenv lam sz cont = compile_fv cenv fv.fv_rev sz (Kclosurecofix(fv.size, init, lbl_types, lbl_bodies) :: cont) + | Lif(t, bt, bf) -> + let branch, cont = make_branch cont in + let lbl_true = Label.create() in + let lbl_false = Label.create() in + compile_lam env cenv t sz + (Kswitch([|lbl_true;lbl_false|],[||]) :: + Klabel lbl_false :: + compile_lam env cenv bf sz + (branch :: + Klabel lbl_true :: + compile_lam env cenv bt sz cont)) | Lcase(ci,rtbl,t,a,branches) -> let ind = ci.ci_ind in @@ -729,41 +773,9 @@ let rec compile_lam env cenv lam sz cont = let cont = code_makeblock ~stack_size:(sz+arity-1) ~arity ~tag cont in comp_args (compile_lam env) cenv args sz cont - | Lprim (kn, ar, op, args) -> - op_compilation env ar op kn cenv args sz cont - - | Luint v -> - (match v with - | UintVal i -> compile_structured_constant cenv (Const_b0 (Uint31.to_int i)) sz cont - | UintDigits ds -> - let nargs = Array.length ds in - if Int.equal nargs 31 then - let (escape,labeled_cont) = make_branch cont in - let else_lbl = Label.create() in - comp_args (compile_lam env) cenv ds sz - ( Kisconst else_lbl::Kareconst(30,else_lbl)::Kcompint31::escape::Klabel else_lbl::Kmakeblock(31, 1)::labeled_cont) - else - let code_construct cont = (* spiwack: variant of the global code_construct - which handles dynamic compilation of - integers *) - let f_cont = - let else_lbl = Label.create () in - [Kacc 0; Kpop 1; Kisconst else_lbl; Kareconst(30,else_lbl); - Kcompint31; Kreturn 0; Klabel else_lbl; Kmakeblock(31, 1); Kreturn 0] - in - let lbl = Label.create() in - fun_code := [Ksequence (add_grab 31 lbl f_cont,!fun_code)]; - Kclosure(lbl,0) :: cont - in - comp_app (fun _ _ _ cont -> code_construct cont) - (compile_lam env) cenv () ds sz cont - | UintDecomp t -> - let escape_lbl, labeled_cont = label_code cont in - compile_lam env cenv t sz ((Kisconst escape_lbl)::Kdecompint31::labeled_cont)) - - -(* spiwack : compilation of constants with their arguments. - Makes a special treatment with 31-bit integer addition *) + | Lprim (kn, op, args) -> + comp_args (compile_lam env) cenv args sz (Kprim(op, kn)::cont) + and compile_get_global cenv (kn,u) sz cont = set_max_stack_size sz; if Univ.Instance.is_empty u then @@ -800,43 +812,6 @@ and compile_constant env cenv kn u args sz cont = comp_app (fun _ _ _ cont -> Kgetglobal kn :: cont) compile_arg cenv () all sz cont -(*template for n-ary operation, invariant: n>=1, - the operations does the following : - 1/ checks if all the arguments are constants (i.e. non-block values) - 2/ if they are, uses the "op" instruction to execute - 3/ if at least one is not, branches to the normal behavior: - Kgetglobal (get_alias !global_env kn) *) -and op_compilation env n op = - let code_construct cenv kn sz cont = - let f_cont = - let else_lbl = Label.create () in - Kareconst(n, else_lbl):: Kacc 0:: Kpop 1:: - op:: Kreturn 0:: Klabel else_lbl:: - (* works as comp_app with nargs = n and tailcall cont [Kreturn 0]*) - compile_get_global cenv kn sz ( - Kappterm(n, n):: []) (* = discard_dead_code [Kreturn 0] *) - in - let lbl = Label.create () in - fun_code := [Ksequence (add_grab n lbl f_cont, !fun_code)]; - Kclosure(lbl, 0)::cont - in - fun kn cenv args sz cont -> - let nargs = Array.length args in - if Int.equal nargs n then (*if it is a fully applied addition*) - let (escape, labeled_cont) = make_branch cont in - let else_lbl = Label.create () in - assert (n < 4); - comp_args (compile_lam env) cenv args sz - (Kisconst else_lbl::(make_areconst (n-1) else_lbl - (*Kaddint31::escape::Klabel else_lbl::Kpush::*) - (op::escape::Klabel else_lbl::Kpush:: - (* works as comp_app with nargs < 4 and non-tailcall cont*) - compile_get_global cenv kn (sz+n) (Kapply n::labeled_cont)))) - else - comp_app (fun cenv _ sz cont -> code_construct cenv kn sz cont) - (compile_lam env) cenv () args sz cont - - let is_univ_copy max u = let u = Univ.Instance.to_array u in if Array.length u = max then @@ -903,7 +878,7 @@ let compile ~fail_on_error ?universes:(universes=0) env c = fn msg; None let compile_constant_body ~fail_on_error env univs = function - | Undef _ | OpaqueDef _ -> Some BCconstant + | Undef _ | OpaqueDef _ | Primitive _ -> Some BCconstant | Def sb -> let body = Mod_subst.force_constr sb in let instance_size = @@ -923,41 +898,3 @@ let compile_constant_body ~fail_on_error env univs = function (* Shortcut of the previous function used during module strengthening *) let compile_alias kn = BCalias (Constant.make1 (Constant.canonical kn)) - -(*(* template compilation for 2ary operation, it probably possible - to make a generic such function with arity abstracted *) -let op2_compilation op = - let code_construct normal cont = (*kn cont =*) - let f_cont = - let else_lbl = Label.create () in - Kareconst(2, else_lbl):: Kacc 0:: Kpop 1:: - op:: Kreturn 0:: Klabel else_lbl:: - (* works as comp_app with nargs = 2 and tailcall cont [Kreturn 0]*) - (*Kgetglobal (get_alias !global_env kn):: *) - normal:: - Kappterm(2, 2):: [] (* = discard_dead_code [Kreturn 0] *) - in - let lbl = Label.create () in - fun_code := [Ksequence (add_grab 2 lbl f_cont, !fun_code)]; - Kclosure(lbl, 0)::cont - in - fun normal fc _ cenv args sz cont -> - if not fc then raise Not_found else - let nargs = Array.length args in - if nargs=2 then (*if it is a fully applied addition*) - let (escape, labeled_cont) = make_branch cont in - let else_lbl = Label.create () in - comp_args compile_constr cenv args sz - (Kisconst else_lbl::(make_areconst 1 else_lbl - (*Kaddint31::escape::Klabel else_lbl::Kpush::*) - (op::escape::Klabel else_lbl::Kpush:: - (* works as comp_app with nargs = 2 and non-tailcall cont*) - (*Kgetglobal (get_alias !global_env kn):: *) - normal:: - Kapply 2::labeled_cont))) - else if nargs=0 then - code_construct normal cont - else - comp_app (fun _ _ _ cont -> code_construct normal cont) - compile_constr cenv () args sz cont *) - diff --git a/kernel/cbytegen.mli b/kernel/cbytegen.mli index 57d3e6fc27..b1b55aef48 100644 --- a/kernel/cbytegen.mli +++ b/kernel/cbytegen.mli @@ -20,7 +20,8 @@ val compile : fail_on_error:bool -> (** init, fun, fv *) val compile_constant_body : fail_on_error:bool -> - env -> constant_universes -> constant_def -> body_code option + env -> constant_universes -> Constr.t Mod_subst.substituted constant_def -> + body_code option (** Shortcut of the previous function used during module strengthening *) diff --git a/kernel/cemitcodes.ml b/kernel/cemitcodes.ml index 50f5607e31..a84a7c0cf9 100644 --- a/kernel/cemitcodes.ml +++ b/kernel/cemitcodes.ml @@ -18,6 +18,7 @@ open Vmvalues open Cbytecodes open Copcodes open Mod_subst +open CPrimitives type emitcodes = String.t @@ -129,6 +130,8 @@ let out_word env b1 b2 b3 b4 = let out env opcode = out_word env opcode 0 0 0 +let is_immed i = Uint63.le (Uint63.of_int i) Uint63.maxuint31 + let out_int env n = out_word env n (n asr 8) (n asr 16) (n asr 24) @@ -198,6 +201,39 @@ let slot_for_proj_name env p = (* Emission of one instruction *) +let nocheck_prim_op = function + | Int63add -> opADDINT63 + | Int63sub -> opSUBINT63 + | Int63lt -> opLTINT63 + | Int63le -> opLEINT63 + | _ -> assert false + + +let check_prim_op = function + | Int63head0 -> opCHECKHEAD0INT63 + | Int63tail0 -> opCHECKTAIL0INT63 + | Int63add -> opCHECKADDINT63 + | Int63sub -> opCHECKSUBINT63 + | Int63mul -> opCHECKMULINT63 + | Int63div -> opCHECKDIVINT63 + | Int63mod -> opCHECKMODINT63 + | Int63lsr -> opCHECKLSRINT63 + | Int63lsl -> opCHECKLSLINT63 + | Int63land -> opCHECKLANDINT63 + | Int63lor -> opCHECKLORINT63 + | Int63lxor -> opCHECKLXORINT63 + | Int63addc -> opCHECKADDCINT63 + | Int63subc -> opCHECKSUBCINT63 + | Int63addCarryC -> opCHECKADDCARRYCINT63 + | Int63subCarryC -> opCHECKSUBCARRYCINT63 + | Int63mulc -> opCHECKMULCINT63 + | Int63diveucl -> opCHECKDIVEUCLINT63 + | Int63div21 -> opCHECKDIV21INT63 + | Int63addMulDiv -> opCHECKADDMULDIVINT63 + | Int63eq -> opCHECKEQINT63 + | Int63lt -> opCHECKLTINT63 + | Int63le -> opCHECKLEINT63 + | Int63compare -> opCHECKCOMPAREINT63 let emit_instr env = function | Klabel lbl -> define_label env lbl @@ -218,7 +254,7 @@ let emit_instr env = function | Kpush_retaddr lbl -> out env opPUSH_RETADDR; out_label env lbl | Kapply n -> - if n < 4 then out env(opAPPLY1 + n - 1) else (out env opAPPLY; out_int env n) + if n <= 4 then out env(opAPPLY1 + n - 1) else (out env opAPPLY; out_int env n) | Kappterm(n, sz) -> if n < 4 then (out env(opAPPTERM1 + n - 1); out_int env sz) else (out env opAPPTERM; out_int env n; out_int env sz) @@ -250,7 +286,7 @@ let emit_instr env = function Array.iter (out_label_with_orig env org) lbl_bodies | Kgetglobal q -> out env opGETGLOBAL; slot_for_getglobal env q - | Kconst (Const_b0 i) -> + | Kconst (Const_b0 i) when is_immed i -> if i >= 0 && i <= 3 then out env (opCONST0 + i) else (out env opCONSTINT; out_int env i) @@ -287,32 +323,20 @@ let emit_instr env = function | Ksequence _ -> invalid_arg "Cemitcodes.emit_instr" | Kproj p -> out env opPROJ; out_int env (Projection.Repr.arg p); slot_for_proj_name env p | Kensurestackcapacity size -> out env opENSURESTACKCAPACITY; out_int env size - (* spiwack *) | Kbranch lbl -> out env opBRANCH; out_label env lbl - | Kaddint31 -> out env opADDINT31 - | Kaddcint31 -> out env opADDCINT31 - | Kaddcarrycint31 -> out env opADDCARRYCINT31 - | Ksubint31 -> out env opSUBINT31 - | Ksubcint31 -> out env opSUBCINT31 - | Ksubcarrycint31 -> out env opSUBCARRYCINT31 - | Kmulint31 -> out env opMULINT31 - | Kmulcint31 -> out env opMULCINT31 - | Kdiv21int31 -> out env opDIV21INT31 - | Kdivint31 -> out env opDIVINT31 - | Kaddmuldivint31 -> out env opADDMULDIVINT31 - | Kcompareint31 -> out env opCOMPAREINT31 - | Khead0int31 -> out env opHEAD0INT31 - | Ktail0int31 -> out env opTAIL0INT31 - | Kisconst lbl -> out env opISCONST; out_label env lbl - | Kareconst(n,lbl) -> out env opARECONST; out_int env n; out_label env lbl - | Kcompint31 -> out env opCOMPINT31 - | Kdecompint31 -> out env opDECOMPINT31 - | Klorint31 -> out env opORINT31 - | Klandint31 -> out env opANDINT31 - | Klxorint31 -> out env opXORINT31 - (*/spiwack *) - | Kstop -> - out env opSTOP + | Kprim (op,None) -> + out env (nocheck_prim_op op) + + | Kprim(op,Some (q,_u)) -> + out env (check_prim_op op); + slot_for_getglobal env q + + | Kareint 1 -> out env opISINT + | Kareint 2 -> out env opAREINT2; + + | Kstop -> out env opSTOP + + | Kareint _ -> assert false (* Emission of a current list and remaining lists of instructions. Include some peephole optimization. *) @@ -337,7 +361,7 @@ let rec emit env insns remaining = match insns with emit env c remaining | Kpush :: Kgetglobal id :: c -> out env opPUSHGETGLOBAL; slot_for_getglobal env id; emit env c remaining - | Kpush :: Kconst (Const_b0 i) :: c -> + | Kpush :: Kconst (Const_b0 i) :: c when is_immed i -> if i >= 0 && i <= 3 then out env (opPUSHCONST0 + i) else (out env opPUSHCONSTINT; out_int env i); @@ -360,7 +384,7 @@ type to_patch = emitcodes * patches * fv (* Substitution *) let subst_strcst s sc = match sc with - | Const_sort _ | Const_b0 _ | Const_univ_level _ | Const_val _ -> sc + | Const_sort _ | Const_b0 _ | Const_univ_level _ | Const_val _ | Const_uint _ -> 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/cemitcodes.mli b/kernel/cemitcodes.mli index 39ddf4a047..41cc641dc8 100644 --- a/kernel/cemitcodes.mli +++ b/kernel/cemitcodes.mli @@ -1,3 +1,10 @@ +(************************************************************************) +(* 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 Vmvalues open Cbytecodes diff --git a/kernel/cinstr.mli b/kernel/cinstr.mli deleted file mode 100644 index dca1757b7d..0000000000 --- a/kernel/cinstr.mli +++ /dev/null @@ -1,52 +0,0 @@ -(************************************************************************) -(* * 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) *) -(************************************************************************) -open Names -open Constr -open Vmvalues -open Cbytecodes - -(** This file defines the lambda code for the bytecode compiler. It has been -extracted from Clambda.ml because of the retroknowledge architecture. *) - -type uint = - | UintVal of Uint31.t - | UintDigits of lambda array - | UintDecomp of lambda - -and lambda = - | Lrel of Name.t * int - | Lvar of Id.t - | Levar of Evar.t * lambda array - | Lprod of lambda * lambda - | Llam of Name.t array * lambda - | Llet of Name.t * lambda * lambda - | Lapp of lambda * lambda array - | Lconst of pconstant - | Lprim of pconstant * int (* arity *) * instruction * lambda array - | Lcase of case_info * reloc_table * lambda * lambda * lam_branches - | Lfix of (int array * int) * fix_decl - | Lcofix of int * fix_decl (* must be in eta-expanded form *) - | Lmakeblock of int * lambda array - | Lval of structured_values - | Lsort of Sorts.t - | Lind of pinductive - | Lproj of Projection.Repr.t * lambda - | Lint of int - | Luint of uint - -(* Cofixpoints have to be in eta-expanded form for their call-by-need evaluation -to be correct. Otherwise, memoization of previous evaluations will be applied -again to extra arguments (see #7333). *) - -and lam_branches = - { constant_branches : lambda array; - nonconstant_branches : (Name.t array * lambda) array } - -and fix_decl = Name.t array * lambda array * lambda array diff --git a/kernel/clambda.ml b/kernel/clambda.ml index 1e4dbfd418..5c21a5ec25 100644 --- a/kernel/clambda.ml +++ b/kernel/clambda.ml @@ -5,13 +5,44 @@ open Term open Constr open Declarations open Vmvalues -open Cbytecodes -open Cinstr open Environ open Pp let pr_con sp = str(Names.Label.to_string (Constant.label sp)) +type lambda = + | Lrel of Name.t * int + | Lvar of Id.t + | Levar of Evar.t * lambda array + | Lprod of lambda * lambda + | Llam of Name.t array * lambda + | Llet of Name.t * lambda * lambda + | Lapp of lambda * lambda array + | Lconst of pconstant + | Lprim of pconstant option * CPrimitives.t * lambda array + (* No check if None *) + | Lcase of case_info * reloc_table * lambda * lambda * lam_branches + | Lif of lambda * lambda * lambda + | Lfix of (int array * int) * fix_decl + | Lcofix of int * fix_decl + | Lint of int + | Lmakeblock of int * lambda array + | Luint of Uint63.t + | Lval of structured_values + | Lsort of Sorts.t + | Lind of pinductive + | Lproj of Projection.Repr.t * lambda + +(* We separate branches for constant and non-constant constructors. If the OCaml + limitation on non-constant constructors is reached, remaining branches are + stored in [extra_branches]. *) +and lam_branches = + { constant_branches : lambda array; + nonconstant_branches : (Name.t array * lambda) array } +(* extra_branches : (name array * lambda) array } *) + +and fix_decl = Name.t array * lambda array * lambda array + (** Printing **) let pp_names ids = @@ -77,6 +108,10 @@ let rec pp_lam lam = pp_names ids ++ str " => " ++ pp_lam c) (Array.to_list branches.nonconstant_branches))) ++ cut() ++ str "end") + | Lif (t, bt, bf) -> + v 0 (str "(if " ++ pp_lam t ++ + cut () ++ str "then " ++ pp_lam bt ++ + cut() ++ str "else " ++ pp_lam bf ++ str ")") | Lfix((t,i),(lna,tl,bl)) -> let fixl = Array.mapi (fun i id -> (id,t.(i),tl.(i),bl.(i))) lna in hov 1 @@ -104,22 +139,26 @@ let rec pp_lam lam = (str "(makeblock " ++ int tag ++ spc() ++ prlist_with_sep spc pp_lam (Array.to_list args) ++ str")") + | Luint i -> str (Uint63.to_string i) | Lval _ -> str "values" | Lsort s -> pp_sort s | Lind ((mind,i), _) -> MutInd.print mind ++ str"#" ++ int i - | Lprim((kn,_u),_ar,_op,args) -> - hov 1 - (str "(PRIM " ++ pr_con kn ++ spc() ++ - prlist_with_sep spc pp_lam (Array.to_list args) ++ - str")") + | Lprim(Some (kn,_u),_op,args) -> + hov 1 + (str "(PRIM " ++ pr_con kn ++ spc() ++ + prlist_with_sep spc pp_lam (Array.to_list args) ++ + str")") + | Lprim(None,op,args) -> + hov 1 + (str "(PRIM_NC " ++ str (CPrimitives.to_string op) ++ spc() ++ + prlist_with_sep spc pp_lam (Array.to_list args) ++ + str")") | Lproj(p,arg) -> hov 1 (str "(proj " ++ Projection.Repr.print p ++ str "(" ++ pp_lam arg ++ str ")") | Lint i -> Pp.(str "(int:" ++ int i ++ str ")") - | Luint _ -> - str "(uint)" (*s Constructors *) @@ -151,9 +190,9 @@ let shift subst = subs_shft (1, subst) (* A generic map function *) -let rec map_lam_with_binders g f n lam = +let map_lam_with_binders g f n lam = match lam with - | Lrel _ | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Lint _ -> lam + | Lrel _ | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Lint _ | Luint _ -> lam | Levar (evk, args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Levar (evk, args') @@ -192,6 +231,11 @@ let rec map_lam_with_binders g f n lam = in if t == t' && a == a' && branches == branches' then lam else Lcase(ci,rtbl,t',a',branches') + | Lif(t,bt,bf) -> + let t' = f n t in + let bt' = f n bt in + let bf' = f n bf in + if t == t' && bt == bt' && bf == bf' then lam else Lif(t',bt',bf') | Lfix(init,(ids,ltypes,lbodies)) -> let ltypes' = Array.Smart.map (f n) ltypes in let lbodies' = Array.Smart.map (f (g (Array.length ids) n)) lbodies in @@ -205,25 +249,12 @@ let rec map_lam_with_binders g f n lam = | Lmakeblock(tag,args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Lmakeblock(tag,args') - | Lprim(kn,ar,op,args) -> + | Lprim(kn,op,args) -> let args' = Array.Smart.map (f n) args in - if args == args' then lam else Lprim(kn,ar,op,args') + if args == args' then lam else Lprim(kn,op,args') | Lproj(p,arg) -> let arg' = f n arg in if arg == arg' then lam else Lproj(p,arg') - | Luint u -> - let u' = map_uint g f n u in - if u == u' then lam else Luint u' - -and map_uint _g f n u = - match u with - | UintVal _ -> u - | UintDigits(args) -> - let args' = Array.Smart.map (f n) args in - if args == args' then u else UintDigits(args') - | UintDecomp(a) -> - let a' = f n a in - if a == a' then u else UintDecomp(a') (*s Lift and substitution *) @@ -271,28 +302,58 @@ let lam_subst_args subst args = let can_subst lam = match lam with - | Lrel _ | Lvar _ | Lconst _ + | Lrel _ | Lvar _ | Lconst _ | Luint _ | Lval _ | Lsort _ | Lind _ -> true | _ -> false + +let can_merge_if bt bf = + match bt, bf with + | Llam(_idst,_), Llam(_idsf,_) -> true + | _ -> false + +let merge_if t bt bf = + let (idst,bodyt) = decompose_Llam bt in + let (idsf,bodyf) = decompose_Llam bf in + let nt = Array.length idst in + let nf = Array.length idsf in + let common,idst,idsf = + if nt = nf then idst, [||], [||] + else + if nt < nf then idst,[||], Array.sub idsf nt (nf - nt) + else idsf, Array.sub idst nf (nt - nf), [||] in + Llam(common, + Lif(lam_lift (Array.length common) t, + mkLlam idst bodyt, + mkLlam idsf bodyf)) + + let rec simplify subst lam = match lam with | Lrel(id,i) -> lam_subst_rel lam id i subst | Llet(id,def,body) -> - let def' = simplify subst def in - if can_subst def' then simplify (cons def' subst) body - else - let body' = simplify (lift subst) body in - if def == def' && body == body' then lam - else Llet(id,def',body') + let def' = simplify subst def in + if can_subst def' then simplify (cons def' subst) body + else + let body' = simplify (lift subst) body in + if def == def' && body == body' then lam + else Llet(id,def',body') | Lapp(f,args) -> - begin match simplify_app subst f subst args with + begin match simplify_app subst f subst args with | Lapp(f',args') when f == f' && args == args' -> lam | lam' -> lam' - end + end + | Lif(t,bt,bf) -> + let t' = simplify subst t in + let bt' = simplify subst bt in + let bf' = simplify subst bf in + if can_merge_if bt' bf' then merge_if t' bt' bf' + else + if t == t' && bt == bt' && bf == bf' then lam + else Lif(t',bt',bf') | _ -> map_lam_with_binders liftn simplify subst lam and simplify_app substf f substa args = @@ -352,7 +413,7 @@ 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 _ -> kind + | Lvar _ | Lconst _ | Lval _ | Lsort _ | Lind _ | Lint _ | Luint _ -> kind | Levar (_, args) -> occurrence_args k kind args | Lprod(dom, codom) -> @@ -363,7 +424,7 @@ let rec occurrence k kind lam = occurrence (k+1) (occurrence k kind def) body | Lapp(f, args) -> occurrence_args k (occurrence k kind f) args - | Lprim(_,_,_,args) | Lmakeblock(_,args) -> + | Lprim(_,_,args) | Lmakeblock(_,args) -> occurrence_args k kind args | Lcase(_ci,_rtbl,t,a,branches) -> let kind = occurrence k (occurrence k kind t) a in @@ -374,6 +435,9 @@ let rec occurrence k kind lam = in Array.iter on_b branches.nonconstant_branches; !r + | Lif (t, bt, bf) -> + let kind = occurrence k kind t in + kind && occurrence k kind bt && occurrence k kind bf | Lfix(_,(ids,ltypes,lbodies)) | Lcofix(_,(ids,ltypes,lbodies)) -> let kind = occurrence_args k kind ltypes in @@ -381,17 +445,10 @@ let rec occurrence k kind lam = kind | Lproj(_,arg) -> occurrence k kind arg - | Luint u -> occurrence_uint k kind u and occurrence_args k kind args = Array.fold_left (occurrence k) kind args -and occurrence_uint k kind u = - match u with - | UintVal _ -> kind - | UintDigits args -> occurrence_args k kind args - | UintDecomp t -> occurrence k kind t - let occur_once lam = try let _ = occurrence 1 true lam in true with Not_found -> false @@ -439,11 +496,12 @@ let check_compilable ib = let is_value lc = match lc with - | Lval _ | Lint _ -> true + | Lval _ | Lint _ | Luint _ -> true | _ -> false let get_value lc = match lc with + | Luint i -> val_of_uint i | Lval v -> v | Lint i -> val_of_int i | _ -> raise Not_found @@ -491,26 +549,18 @@ let rec get_alias env kn = (* Compilation of primitive *) -let _h = Name(Id.of_string "f") +let prim kn p args = + Lprim(Some kn, p, args) -(* let expand_prim kn op arity = let ids = Array.make arity Anonymous in let args = make_args arity 1 in Llam(ids, prim kn op args) -*) -let compile_prim n op kn fc args = - if not fc then raise Not_found - else - Lprim(kn, n, op, args) - - (* - let (nparams, arity) = CPrimitives.arity op in - let expected = nparams + arity in - if Array.length args >= expected then prim kn op args - else mkLapp (expand_prim kn op expected) args -*) +let lambda_of_prim kn op args = + let arity = CPrimitives.arity op in + if Array.length args >= arity then prim kn op args + else mkLapp (expand_prim kn op arity) args (*i Global environment *) @@ -661,13 +711,6 @@ let rec lambda_of_constr env c = (* translation of the argument *) let la = lambda_of_constr env a in - let gr = GlobRef.IndRef ind in - let la = - try - Retroknowledge.get_vm_before_match_info env.global_env.retroknowledge - gr la - with Not_found -> la - in (* translation of the type *) let lt = lambda_of_constr env t in (* translation of branches *) @@ -713,88 +756,30 @@ let rec lambda_of_constr env c = let lc = lambda_of_constr env c in Lproj (Projection.repr p,lc) + | Int i -> Luint i + and lambda_of_app env f args = match Constr.kind f with - | Const (kn,_u as c) -> - let kn = get_alias env.global_env kn in - (* spiwack: checks if there is a specific way to compile the constant - if there is not, Not_found is raised, and the function - falls back on its normal behavior *) - (try - (* We delay the compilation of arguments to avoid an exponential behavior *) - let f = Retroknowledge.get_vm_compiling_info env.global_env.retroknowledge - (GlobRef.ConstRef kn) in - let args = lambda_of_args env 0 args in - f args - with Not_found -> - let cb = lookup_constant kn env.global_env in - begin match cb.const_body with + | Const (kn,u as c) -> + let kn = get_alias env.global_env kn in + let cb = lookup_constant kn env.global_env in + begin match cb.const_body with + | Primitive op -> lambda_of_prim (kn,u) op (lambda_of_args env 0 args) | Def csubst when cb.const_inline_code -> - lambda_of_app env (Mod_subst.force_constr csubst) args + lambda_of_app env (Mod_subst.force_constr csubst) args | Def _ | OpaqueDef _ | Undef _ -> mkLapp (Lconst c) (lambda_of_args env 0 args) - end) + end | Construct (c,_) -> - let tag, nparams, arity = Renv.get_construct_info env c in - let nargs = Array.length args in - let gr = GlobRef.ConstructRef c in - if Int.equal (nparams + arity) nargs then (* fully applied *) - (* spiwack: *) - (* 1/ tries to compile the constructor in an optimal way, - it is supposed to work only if the arguments are - all fully constructed, fails with Cbytecodes.NotClosed. - it can also raise Not_found when there is no special - treatment for this constructor - for instance: tries to to compile an integer of the - form I31 D1 D2 ... D31 to [D1D2...D31] as - a processor number (a caml number actually) *) - try - try - Retroknowledge.get_vm_constant_static_info - env.global_env.retroknowledge - gr args - with NotClosed -> - (* 2/ if the arguments are not all closed (this is - expectingly (and it is currently the case) the only - reason why this exception is raised) tries to - give a clever, run-time behavior to the constructor. - Raises Not_found if there is no special treatment - for this integer. - this is done in a lazy fashion, using the constructor - Bspecial because it needs to know the continuation - and such, which can't be done at this time. - for instance, for int31: if one of the digit is - not closed, it's not impossible that the number - gets fully instantiated at run-time, thus to ensure - uniqueness of the representation in the vm - it is necessary to try and build a caml integer - during the execution *) - let rargs = Array.sub args nparams arity in - let args = lambda_of_args env nparams rargs in - Retroknowledge.get_vm_constant_dynamic_info - env.global_env.retroknowledge - gr args - with Not_found -> - (* 3/ if no special behavior is available, then the compiler - falls back to the normal behavior *) + let tag, nparams, arity = Renv.get_construct_info env c in + let nargs = Array.length args in + if nparams < nargs then (* got all parameters *) let args = lambda_of_args env nparams args in makeblock tag 0 arity args - else - let args = lambda_of_args env nparams args in - (* spiwack: tries first to apply the run-time compilation - behavior of the constructor, as in 2/ above *) - (try - (Retroknowledge.get_vm_constant_dynamic_info - env.global_env.retroknowledge - gr) args - with Not_found -> - if nparams <= nargs then (* got all parameters *) - makeblock tag 0 arity args - else (* still expecting some parameters *) - makeblock tag (nparams - nargs) arity empty_args) + else makeblock tag (nparams - nargs) arity empty_args | _ -> - let f = lambda_of_constr env f in - let args = lambda_of_args env 0 args in - mkLapp f args + let f = lambda_of_constr env f in + let args = lambda_of_args env 0 args in + mkLapp f args and lambda_of_args env start args = let nargs = Array.length args in @@ -822,43 +807,3 @@ let lambda_of_constr ~optimize genv c = if !dump_lambda then Feedback.msg_debug (pp_lam lam); lam - -(** Retroknowledge, to be removed once we move to primitive machine integers *) -let compile_structured_int31 fc args = - if not fc then raise Not_found else - Luint (UintVal - (Uint31.of_int (Array.fold_left - (fun temp_i -> fun t -> match kind t with - | Construct ((_,d),_) -> 2*temp_i+d-1 - | _ -> raise NotClosed) - 0 args))) - -let dynamic_int31_compilation fc args = - if not fc then raise Not_found else - Luint (UintDigits args) - -let d0 = Lint 0 -let d1 = Lint 1 - -(* We are relying here on the tags of digits constructors *) -let digits_from_uint i = - let digits = Array.make 31 d0 in - for k = 0 to 30 do - if Int.equal ((Uint31.to_int i lsr k) land 1) 1 then - digits.(30-k) <- d1 - done; - digits - -let int31_escape_before_match fc t = - if not fc then - raise Not_found - else - match t with - | Luint (UintVal i) -> - let digits = digits_from_uint i in - Lmakeblock (1, digits) - | Luint (UintDigits args) -> - Lmakeblock (1,args) - | Luint (UintDecomp _) -> - assert false - | _ -> Luint (UintDecomp t) diff --git a/kernel/clambda.mli b/kernel/clambda.mli index 8ff10b4549..4d921fd45e 100644 --- a/kernel/clambda.mli +++ b/kernel/clambda.mli @@ -1,7 +1,37 @@ open Names -open Cinstr +open Constr +open Vmvalues open Environ +type lambda = + | Lrel of Name.t * int + | Lvar of Id.t + | Levar of Evar.t * lambda array + | Lprod of lambda * lambda + | Llam of Name.t array * lambda + | Llet of Name.t * lambda * lambda + | Lapp of lambda * lambda array + | Lconst of pconstant + | Lprim of pconstant option * CPrimitives.t * lambda array + (* No check if None *) + | Lcase of case_info * reloc_table * lambda * lambda * lam_branches + | Lif of lambda * lambda * lambda + | Lfix of (int array * int) * fix_decl + | Lcofix of int * fix_decl + | Lint of int + | Lmakeblock of int * lambda array + | Luint of Uint63.t + | Lval of structured_values + | Lsort of Sorts.t + | Lind of pinductive + | Lproj of Projection.Repr.t * lambda + +and lam_branches = + { constant_branches : lambda array; + nonconstant_branches : (Name.t array * lambda) array } + +and fix_decl = Name.t array * lambda array * lambda array + exception TooLargeInductive of Pp.t val lambda_of_constr : optimize:bool -> env -> Constr.t -> lambda @@ -10,22 +40,5 @@ val decompose_Llam : lambda -> Name.t array * lambda val get_alias : env -> Constant.t -> Constant.t -val compile_prim : int -> Cbytecodes.instruction -> Constr.pconstant -> bool -> lambda array -> lambda - -(** spiwack: this function contains the information needed to perform - the static compilation of int31 (trying and obtaining - a 31-bit integer in processor representation at compile time) *) -val compile_structured_int31 : bool -> Constr.t array -> lambda - -(** this function contains the information needed to perform - the dynamic compilation of int31 (trying and obtaining a - 31-bit integer in processor representation at runtime when - it failed at compile time *) -val dynamic_int31_compilation : bool -> lambda array -> lambda - -(*spiwack: compiling function to insert dynamic decompilation before - matching integers (in case they are in processor representation) *) -val int31_escape_before_match : bool -> lambda -> lambda - (** Dump the VM lambda code after compilation (for debugging purposes) *) val dump_lambda : bool ref diff --git a/kernel/constr.ml b/kernel/constr.ml index d67d15b23b..c392494e95 100644 --- a/kernel/constr.ml +++ b/kernel/constr.ml @@ -101,6 +101,7 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term = | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint | Proj of Projection.t * 'constr + | Int of Uint63.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 @@ -233,6 +234,9 @@ let mkRef (gr,u) = let open GlobRef in match gr with | ConstructRef c -> mkConstructU (c,u) | VarRef x -> mkVar x +(* Constructs a primitive integer *) +let mkInt i = Int i + (************************************************************************) (* kind_of_term = constructions as seen by the user *) (************************************************************************) @@ -438,7 +442,7 @@ let decompose_appvect c = let fold f acc c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _) -> acc + | Construct _ | Int _) -> 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 @@ -458,7 +462,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 _) -> () + | Construct _ | Int _) -> () | Cast (c,_,t) -> f c; f t | Prod (_,t,c) -> f t; f c | Lambda (_,t,c) -> f t; f c @@ -478,7 +482,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 _) -> () + | Construct _ | Int _) -> () | 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 @@ -504,7 +508,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 _) -> acc + | Construct _ | Int _) -> 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 @@ -600,7 +604,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 _) -> c + | Construct _ | Int _) -> c | Cast (b,k,t) -> let b' = f b in let t' = f t in @@ -665,7 +669,7 @@ let map = map_gen false let fold_map f accu c = match kind c with | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _) -> accu, c + | Construct _ | Int _) -> accu, c | Cast (b,k,t) -> let accu, b' = f accu b in let accu, t' = f accu t in @@ -725,7 +729,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 _) -> c0 + | Construct _ | Int _) -> c0 | Cast (c, k, t) -> let c' = f l c in let t' = f l t in @@ -802,7 +806,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 _ -> acc + | Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ | Int _ -> 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 @@ -843,6 +847,7 @@ let compare_head_gen_leq_with kind1 kind2 leq_universes leq_sorts eq leq nargs t | Rel n1, Rel n2 -> Int.equal n1 n2 | Meta m1, Meta m2 -> Int.equal m1 m2 | Var id1, Var id2 -> Id.equal id1 id2 + | Int i1, Int i2 -> Uint63.equal i1 i2 | 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 @@ -869,7 +874,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 _), _ -> false + | CoFix _ | Int _), _ -> 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, @@ -1044,6 +1049,8 @@ let constr_ord_int f t1 t2 = ln1 ln2 tl1 tl2 bl1 bl2 | CoFix _, _ -> -1 | _, CoFix _ -> 1 | 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 let rec compare m n= constr_ord_int compare m n @@ -1127,9 +1134,10 @@ let hasheq t1 t2 = && array_eqeq lna1 lna2 && array_eqeq tl1 tl2 && array_eqeq bl1 bl2 + | Int i1, Int i2 -> i1 == i2 | (Rel _ | Meta _ | Var _ | Sort _ | Cast _ | Prod _ | Lambda _ | LetIn _ | App _ | Proj _ | Evar _ | Const _ | Ind _ | Construct _ | Case _ - | Fix _ | CoFix _), _ -> false + | Fix _ | CoFix _ | Int _), _ -> 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 *) @@ -1190,10 +1198,6 @@ let hashcons (sh_sort,sh_ci,sh_construct,sh_ind,sh_con,sh_na,sh_id) = | Evar (e,l) -> let l, hl = hash_term_array l in (Evar (e,l), combinesmall 8 (combine (Evar.hash e) hl)) - | Proj (p,c) -> - let c, hc = sh_rec c in - let p' = Projection.hcons p in - (Proj (p', c), combinesmall 17 (combine (Projection.SyntacticOrd.hash p') hc)) | Const (c,u) -> let c' = sh_con c in let u', hu = sh_instance u in @@ -1232,6 +1236,13 @@ let hashcons (sh_sort,sh_ci,sh_construct,sh_ind,sh_con,sh_na,sh_id) = (t, combinesmall 15 n) | Rel n -> (t, combinesmall 16 n) + | Proj (p,c) -> + let c, hc = sh_rec c in + let p' = Projection.hcons p in + (Proj (p', c), combinesmall 17 (combine (Projection.SyntacticOrd.hash p') hc)) + | Int i -> + let (h,l) = Uint63.to_int2 i in + (t, combinesmall 18 (combine h l)) and sh_rec t = let (y, h) = hash_term t in @@ -1277,8 +1288,6 @@ let rec hash t = | App (Cast(c, _, _),l) -> hash (mkApp (c,l)) | App (c,l) -> combinesmall 7 (combine (hash_term_array l) (hash c)) - | Proj (p,c) -> - combinesmall 17 (combine (Projection.hash p) (hash c)) | Evar (e,l) -> combinesmall 8 (combine (Evar.hash e) (hash_term_array l)) | Const (c,u) -> @@ -1295,6 +1304,9 @@ let rec hash t = combinesmall 14 (combine (hash_term_array bl) (hash_term_array tl)) | Meta n -> combinesmall 15 n | Rel n -> combinesmall 16 n + | Proj (p,c) -> + combinesmall 17 (combine (Projection.hash p) (hash c)) + | Int i -> combinesmall 18 (Uint63.hash i) and hash_term_array t = Array.fold_left (fun acc t -> combine (hash t) acc) 0 t @@ -1425,3 +1437,4 @@ let rec debug_print c = Name.print na ++ str":" ++ debug_print ty ++ cut() ++ str":=" ++ debug_print bd) (Array.to_list fixl)) ++ str"}") + | Int i -> str"Int("++str (Uint63.to_string i) ++ str")" diff --git a/kernel/constr.mli b/kernel/constr.mli index f2cedcdabb..fdc3296a6a 100644 --- a/kernel/constr.mli +++ b/kernel/constr.mli @@ -72,6 +72,9 @@ val mkRel : int -> constr (** Constructs a Variable *) val mkVar : Id.t -> constr +(** Constructs a machine integer *) +val mkInt : Uint63.t -> constr + (** Constructs an patvar named "?n" *) val mkMeta : metavariable -> constr @@ -228,6 +231,7 @@ type ('constr, 'types, 'sort, 'univs) kind_of_term = | Fix of ('constr, 'types) pfixpoint | CoFix of ('constr, 'types) pcofixpoint | Proj of Projection.t * 'constr + | Int of Uint63.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/cooking.ml b/kernel/cooking.ml index f4b4834d98..88586352f6 100644 --- a/kernel/cooking.ml +++ b/kernel/cooking.ml @@ -155,7 +155,7 @@ type recipe = { from : constant_body; info : Opaqueproof.cooking_info } type inline = bool type result = { - cook_body : constant_def; + cook_body : constr Mod_subst.substituted constant_def; cook_type : types; cook_universes : constant_universes; cook_private_univs : Univ.ContextSet.t option; @@ -169,6 +169,7 @@ let on_body ml hy f = function | OpaqueDef o -> OpaqueDef (Opaqueproof.discharge_direct_opaque ~cook_constr:f { Opaqueproof.modlist = ml; abstract = hy } o) + | Primitive _ -> CErrors.anomaly (Pp.str "Primitives cannot be cooked") let expmod_constr_subst cache modlist subst c = let subst = Univ.make_instance_subst subst in diff --git a/kernel/cooking.mli b/kernel/cooking.mli index 7ff4b657d3..07c6f37fab 100644 --- a/kernel/cooking.mli +++ b/kernel/cooking.mli @@ -18,7 +18,7 @@ type recipe = { from : constant_body; info : Opaqueproof.cooking_info } type inline = bool type result = { - cook_body : constant_def; + cook_body : constr Mod_subst.substituted constant_def; cook_type : types; cook_universes : constant_universes; cook_private_univs : Univ.ContextSet.t option; diff --git a/kernel/declarations.ml b/kernel/declarations.ml index 016b63be09..1008492825 100644 --- a/kernel/declarations.ml +++ b/kernel/declarations.ml @@ -47,10 +47,11 @@ type inline = int option transparent body, or an opaque one *) (* Global declarations (i.e. constants) can be either: *) -type constant_def = +type 'a constant_def = | Undef of inline (** a global assumption *) - | Def of constr Mod_subst.substituted (** or a transparent global definition *) + | Def of 'a (** or a transparent global definition *) | OpaqueDef of Opaqueproof.opaque (** or an opaque global definition *) + | Primitive of CPrimitives.t (** or a primitive operation *) type constant_universes = | Monomorphic_const of Univ.ContextSet.t @@ -88,7 +89,7 @@ type typing_flags = { * the OpaqueDef *) type constant_body = { const_hyps : Constr.named_context; (** New: younger hyp at top *) - const_body : constant_def; + const_body : Constr.t Mod_subst.substituted constant_def; const_type : types; const_body_code : Cemitcodes.to_patch_substituted option; const_universes : constant_universes; diff --git a/kernel/declareops.ml b/kernel/declareops.ml index 707c46048b..5686c4071d 100644 --- a/kernel/declareops.ml +++ b/kernel/declareops.ml @@ -56,8 +56,9 @@ let constant_is_polymorphic cb = | Monomorphic_const _ -> false | Polymorphic_const _ -> true + let constant_has_body cb = match cb.const_body with - | Undef _ -> false + | Undef _ | Primitive _ -> false | Def _ | OpaqueDef _ -> true let constant_polymorphic_context cb = @@ -67,7 +68,7 @@ let constant_polymorphic_context cb = let is_opaque cb = match cb.const_body with | OpaqueDef _ -> true - | Undef _ | Def _ -> false + | Undef _ | Def _ | Primitive _ -> false (** {7 Constant substitutions } *) @@ -83,7 +84,7 @@ let subst_const_type sub arity = (** No need here to check for physical equality after substitution, at least for Def due to the delayed substitution [subst_constr_subst]. *) let subst_const_def sub def = match def with - | Undef _ -> def + | Undef _ | Primitive _ -> def | Def c -> Def (subst_constr sub c) | OpaqueDef o -> OpaqueDef (Opaqueproof.subst_opaque sub o) @@ -119,6 +120,7 @@ let hcons_rel_context l = List.Smart.map hcons_rel_decl l let hcons_const_def = function | Undef inl -> Undef inl + | Primitive p -> Primitive p | Def l_constr -> let constr = force_constr l_constr in Def (from_val (Constr.hcons constr)) diff --git a/kernel/dune b/kernel/dune index 01abdb8f67..79161519ba 100644 --- a/kernel/dune +++ b/kernel/dune @@ -3,7 +3,7 @@ (synopsis "The Coq Kernel") (public_name coq.kernel) (wrapped false) - (modules_without_implementation cinstr nativeinstr) + (modules (:standard \ uint63_x86 uint63_amd64 write_uint63)) (libraries lib byterun)) (rule @@ -11,6 +11,16 @@ (deps (:h-file byterun/coq_instruct.h) make-opcodes) (action (run ./make_opcodes.sh %{h-file} %{targets}))) +(executable + (name write_uint63) + (modules write_uint63) + (libraries unix)) + +(rule + (targets uint63.ml) + (deps (:gen ./write_uint63.exe) uint63_x86.ml uint63_amd64.ml) + (action (run %{gen}))) + (documentation (package coq)) diff --git a/kernel/entries.ml b/kernel/entries.ml index 58bb782f15..013327599d 100644 --- a/kernel/entries.ml +++ b/kernel/entries.ml @@ -87,9 +87,16 @@ type inline = int option (* inlining level, None for no inlining *) type parameter_entry = Constr.named_context option * types in_constant_universes_entry * inline +type primitive_entry = { + prim_entry_type : types option; + prim_entry_univs : Univ.ContextSet.t; (* always monomorphic *) + prim_entry_content : CPrimitives.op_or_type; +} + type 'a constant_entry = | DefinitionEntry of 'a definition_entry | ParameterEntry of parameter_entry + | PrimitiveEntry of primitive_entry (** {6 Modules } *) diff --git a/kernel/environ.ml b/kernel/environ.ml index 77820a301e..02f38e7214 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -119,7 +119,7 @@ let empty_env = { env_universes = UGraph.initial_universes; env_engagement = PredicativeSet }; env_typing_flags = Declareops.safe_flags Conv_oracle.empty; - retroknowledge = Retroknowledge.initial_retroknowledge; + retroknowledge = Retroknowledge.empty; indirect_pterms = Opaqueproof.empty_opaquetab } @@ -450,7 +450,10 @@ let constant_type env (kn,u) = let csts = Univ.AUContext.instantiate u uctx in (subst_instance_constr u cb.const_type, csts) -type const_evaluation_result = NoBody | Opaque +type const_evaluation_result = + | NoBody + | Opaque + | IsPrimitive of CPrimitives.t exception NotEvaluableConst of const_evaluation_result @@ -461,14 +464,14 @@ let constant_value_and_type env (kn, u) = let b' = match cb.const_body with | Def l_body -> Some (subst_instance_constr u (Mod_subst.force_constr l_body)) | OpaqueDef _ -> None - | Undef _ -> None + | Undef _ | Primitive _ -> None in b', subst_instance_constr u cb.const_type, cst let body_of_constant_body env cb = let otab = opaque_tables env in match cb.const_body with - | Undef _ -> + | Undef _ | Primitive _ -> None | Def c -> Some (Mod_subst.force_constr c, Declareops.constant_polymorphic_context cb) @@ -492,6 +495,7 @@ let constant_value_in env (kn,u) = subst_instance_constr u b | OpaqueDef _ -> raise (NotEvaluableConst Opaque) | Undef _ -> raise (NotEvaluableConst NoBody) + | Primitive p -> raise (NotEvaluableConst (IsPrimitive p)) let constant_opt_value_in env cst = try Some (constant_value_in env cst) @@ -503,7 +507,13 @@ let evaluable_constant kn env = match cb.const_body with | Def _ -> true | OpaqueDef _ -> false - | Undef _ -> false + | Undef _ | Primitive _ -> false + +let is_primitive env c = + let cb = lookup_constant c env in + match cb.Declarations.const_body with + | Declarations.Primitive _ -> true + | _ -> false let polymorphic_constant cst env = Declareops.constant_is_polymorphic (lookup_constant cst env) @@ -743,29 +753,4 @@ let is_type_in_type env r = | IndRef ind -> type_in_type_ind ind env | ConstructRef cstr -> type_in_type_ind (inductive_of_constructor cstr) env -(*spiwack: the following functions assemble the pieces of the retroknowledge - note that the "consistent" register function is available in the module - Safetyping, Environ only synchronizes the proactive and the reactive parts*) - -open Retroknowledge - -(* lifting of the "get" functions works also for "mem"*) -let retroknowledge f env = - f env.retroknowledge - -let registered env field = - retroknowledge mem env field - -let register_one env field entry = - { env with retroknowledge = Retroknowledge.add_field env.retroknowledge field entry } - -(* [register env field entry] may register several fields when needed *) -let register env field gr = - match field with - | KInt31 Int31Type -> - let i31c = match gr with - | GlobRef.IndRef i31t -> GlobRef.ConstructRef (i31t, 1) - | _ -> anomaly ~label:"Environ.register" (Pp.str "should be an inductive type.") - in - register_one (register_one env (KInt31 Int31Constructor) i31c) field gr - | field -> register_one env field gr +let set_retroknowledge env r = { env with retroknowledge = r } diff --git a/kernel/environ.mli b/kernel/environ.mli index 6d4d3b282b..8d5bd85b94 100644 --- a/kernel/environ.mli +++ b/kernel/environ.mli @@ -195,11 +195,15 @@ val type_in_type_constant : Constant.t -> env -> bool (** {6 ... } *) (** [constant_value env c] raises [NotEvaluableConst Opaque] if - [c] is opaque and [NotEvaluableConst NoBody] if it has no - body and [NotEvaluableConst IsProj] if [c] is a projection + [c] is opaque, [NotEvaluableConst NoBody] if it has no + body, [NotEvaluableConst IsProj] if [c] is a projection, + [NotEvaluableConst (IsPrimitive p)] if [c] is primitive [p] and [Not_found] if it does not exist in [env] *) -type const_evaluation_result = NoBody | Opaque +type const_evaluation_result = + | NoBody + | Opaque + | IsPrimitive of CPrimitives.t exception NotEvaluableConst of const_evaluation_result val constant_type : env -> Constant.t puniverses -> types constrained @@ -223,6 +227,8 @@ val constant_value_in : env -> Constant.t puniverses -> constr val constant_type_in : env -> Constant.t puniverses -> types val constant_opt_value_in : env -> Constant.t puniverses -> constr option +val is_primitive : env -> Constant.t -> bool + (** {6 Primitive projections} *) (** Checks that the number of parameters is correct. *) @@ -334,13 +340,8 @@ val is_polymorphic : env -> Names.GlobRef.t -> bool val is_template_polymorphic : env -> GlobRef.t -> bool val is_type_in_type : env -> GlobRef.t -> bool -open Retroknowledge -(** functions manipulating the retroknowledge - @author spiwack *) - -val registered : env -> field -> bool - -val register : env -> field -> GlobRef.t -> env - (** Native compiler *) val no_link_info : link_info + +(** Primitives *) +val set_retroknowledge : env -> Retroknowledge.retroknowledge -> env diff --git a/kernel/inductive.ml b/kernel/inductive.ml index 05c5c0e821..c62d0e7ded 100644 --- a/kernel/inductive.ml +++ b/kernel/inductive.ml @@ -804,7 +804,7 @@ let rec subterm_specif renv stack t = | Not_subterm -> Not_subterm) | Var _ | Sort _ | Cast _ | Prod _ | LetIn _ | App _ | Const _ | Ind _ - | Construct _ | CoFix _ -> Not_subterm + | Construct _ | CoFix _ | Int _ -> Not_subterm (* Other terms are not subterms *) @@ -1008,7 +1008,7 @@ let check_one_fix renv recpos trees def = check_rec_call renv stack (Term.applist(c,l)) end - | Sort _ -> + | Sort _ | Int _ -> assert (List.is_empty l) (* l is not checked because it is considered as the meta's context *) @@ -1194,7 +1194,8 @@ 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 _ -> raise (CoFixGuardError (env,NotGuardedForm t)) in + | Ind _ | Fix _ | Proj _ | Int _ -> + raise (CoFixGuardError (env,NotGuardedForm t)) in let ((mind, _),_) = codomain_is_coind env deftype in let vlra = lookup_subterms env mind in diff --git a/kernel/kernel.mllib b/kernel/kernel.mllib index 0b10e788b6..5108744bde 100644 --- a/kernel/kernel.mllib +++ b/kernel/kernel.mllib @@ -1,6 +1,7 @@ Names TransparentState -Uint31 +Uint63 +CPrimitives Univ UGraph Esubst @@ -19,11 +20,11 @@ Opaqueproof Declarations Entries Nativevalues -CPrimitives Declareops Retroknowledge Conv_oracle Environ +Primred CClosure Reduction Clambda diff --git a/kernel/mod_subst.ml b/kernel/mod_subst.ml index 52fb39e1d0..cd675653cb 100644 --- a/kernel/mod_subst.ml +++ b/kernel/mod_subst.ml @@ -330,6 +330,19 @@ let subst_proj_repr sub p = let subst_proj sub p = Projection.map (subst_mind sub) p +let subst_retro_action subst action = + let open Retroknowledge in + match action with + | Register_ind(prim,ind) -> + let ind' = subst_ind subst ind in + if ind == ind' then action else Register_ind(prim, ind') + | Register_type(prim,c) -> + let c' = subst_constant subst c in + if c == c' then action else Register_type(prim, c') + | Register_inline(c) -> + let c' = subst_constant subst c in + if c == c' then action else Register_inline(c') + (* Here the semantics is completely unclear. What does "Hint Unfold t" means when "t" is a parameter? Does the user mean "Unfold X.t" or does she mean "Unfold y" diff --git a/kernel/mod_subst.mli b/kernel/mod_subst.mli index ea391b3de7..8ab3d04402 100644 --- a/kernel/mod_subst.mli +++ b/kernel/mod_subst.mli @@ -144,6 +144,8 @@ val subst_constant : val subst_proj_repr : substitution -> Projection.Repr.t -> Projection.Repr.t val subst_proj : substitution -> Projection.t -> Projection.t +val subst_retro_action : substitution -> Retroknowledge.action -> Retroknowledge.action + (** Here the semantics is completely unclear. What does "Hint Unfold t" means when "t" is a parameter? Does the user mean "Unfold X.t" or does she mean "Unfold y" diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml index d63dc057b4..f68dd158c2 100644 --- a/kernel/mod_typing.ml +++ b/kernel/mod_typing.ml @@ -68,12 +68,12 @@ let rec check_with_def env struc (idl,(c,ctx)) mp equiv = if List.is_empty idl then (* Toplevel definition *) let cb = match spec with - | SFBconst cb -> cb - | _ -> error_not_a_constant lab + | SFBconst cb -> cb + | _ -> error_not_a_constant lab in (* In the spirit of subtyping.check_constant, we accept any implementations of parameters and opaques terms, - as long as they have the right type *) + as long as they have the right type *) let c', univs, ctx' = match cb.const_universes, ctx with | Monomorphic_const _, None -> @@ -87,6 +87,8 @@ let rec check_with_def env struc (idl,(c,ctx)) mp equiv = | Def cs -> let c' = Mod_subst.force_constr cs in c, Reduction.infer_conv env' (Environ.universes env') c c' + | Primitive _ -> + error_incorrect_with_constraint lab in c', Monomorphic_const Univ.ContextSet.empty, cst | Polymorphic_const uctx, Some ctx -> @@ -95,52 +97,54 @@ let rec check_with_def env struc (idl,(c,ctx)) mp equiv = error_incorrect_with_constraint lab in (** Terms are compared in a context with De Bruijn universe indices *) - let env' = Environ.push_context ~strict:false (Univ.AUContext.repr uctx) env in - let cst = match cb.const_body with - | Undef _ | OpaqueDef _ -> - let j = Typeops.infer env' c in - let typ = cb.const_type in - let cst' = Reduction.infer_conv_leq env' (Environ.universes env') - j.uj_type typ in - cst' - | Def cs -> - let c' = Mod_subst.force_constr cs in - let cst' = Reduction.infer_conv env' (Environ.universes env') c c' in - cst' - in - if not (Univ.Constraint.is_empty cst) then - error_incorrect_with_constraint lab; - c, Polymorphic_const ctx, Univ.Constraint.empty + let env' = Environ.push_context ~strict:false (Univ.AUContext.repr uctx) env in + let cst = match cb.const_body with + | Undef _ | OpaqueDef _ -> + let j = Typeops.infer env' c in + let typ = cb.const_type in + let cst' = Reduction.infer_conv_leq env' (Environ.universes env') + j.uj_type typ in + cst' + | Def cs -> + let c' = Mod_subst.force_constr cs in + let cst' = Reduction.infer_conv env' (Environ.universes env') c c' in + cst' + | Primitive _ -> + error_incorrect_with_constraint lab + in + if not (Univ.Constraint.is_empty cst) then + error_incorrect_with_constraint lab; + c, Polymorphic_const ctx, Univ.Constraint.empty | _ -> error_incorrect_with_constraint lab in let def = Def (Mod_subst.from_val c') in -(* let ctx' = Univ.UContext.make (newus, cst) in *) + (* let ctx' = Univ.UContext.make (newus, cst) in *) let cb' = - { cb with - const_body = def; + { cb with + const_body = def; const_universes = univs ; - const_body_code = Option.map Cemitcodes.from_val - (Cbytegen.compile_constant_body ~fail_on_error:false env' cb.const_universes def) } + const_body_code = Option.map Cemitcodes.from_val + (Cbytegen.compile_constant_body ~fail_on_error:false env' cb.const_universes def) } in before@(lab,SFBconst(cb'))::after, c', ctx' else (* Definition inside a sub-module *) let mb = match spec with - | SFBmodule mb -> mb - | _ -> error_not_a_module (Label.to_string lab) + | SFBmodule mb -> mb + | _ -> error_not_a_module (Label.to_string lab) in begin match mb.mod_expr with - | Abstract -> - let struc = Modops.destr_nofunctor mb.mod_type in - let struc',c',cst = - check_with_def env' struc (idl,(c,ctx)) (MPdot(mp,lab)) mb.mod_delta - in - let mb' = { mb with - mod_type = NoFunctor struc'; - mod_type_alg = None } - in - before@(lab,SFBmodule mb')::after, c', cst - | _ -> error_generative_module_expected lab + | Abstract -> + let struc = Modops.destr_nofunctor mb.mod_type in + let struc',c',cst = + check_with_def env' struc (idl,(c,ctx)) (MPdot(mp,lab)) mb.mod_delta + in + let mb' = { mb with + mod_type = NoFunctor struc'; + mod_type_alg = None } + in + before@(lab,SFBmodule mb')::after, c', cst + | _ -> error_generative_module_expected lab end with | Not_found -> error_no_such_label lab diff --git a/kernel/modops.ml b/kernel/modops.ml index 97ac3cdebb..1dc8eec0da 100644 --- a/kernel/modops.ml +++ b/kernel/modops.ml @@ -197,9 +197,18 @@ let rec subst_structure sub do_delta sign = in List.Smart.map subst_body sign +and subst_retro : type a. Mod_subst.substitution -> a module_retroknowledge -> a module_retroknowledge = + fun subst retro -> + match retro with + | ModTypeRK as r -> r + | ModBodyRK l as r -> + let l' = List.Smart.map (subst_retro_action subst) l in + if l == l' then r else ModBodyRK l + and subst_body : 'a. _ -> _ -> (_ -> 'a -> 'a) -> _ -> 'a generic_module_body -> 'a generic_module_body = fun is_mod sub subst_impl do_delta mb -> - let { mod_mp=mp; mod_expr=me; mod_type=ty; mod_type_alg=aty; _ } = mb in + let { mod_mp=mp; mod_expr=me; mod_type=ty; mod_type_alg=aty; + mod_retroknowledge=retro; _ } = mb in let mp' = subst_mp sub mp in let sub = if ModPath.equal mp mp' then sub @@ -209,8 +218,10 @@ and subst_body : 'a. _ -> _ -> (_ -> 'a -> 'a) -> _ -> 'a generic_module_body -> let ty' = subst_signature sub do_delta ty in let me' = subst_impl sub me in let aty' = Option.Smart.map (subst_expression sub id_delta) aty in + let retro' = subst_retro sub retro in let delta' = do_delta mb.mod_delta sub in - if mp==mp' && me==me' && ty==ty' && aty==aty' && delta'==mb.mod_delta + if mp==mp' && me==me' && ty==ty' && aty==aty' + && retro==retro' && delta'==mb.mod_delta then mb else { mb with @@ -218,7 +229,9 @@ and subst_body : 'a. _ -> _ -> (_ -> 'a -> 'a) -> _ -> 'a generic_module_body -> mod_expr = me'; mod_type = ty'; mod_type_alg = aty'; - mod_delta = delta' } + mod_retroknowledge = retro'; + mod_delta = delta'; + } and subst_module sub do_delta mb = subst_body true sub subst_impl do_delta mb @@ -259,32 +272,12 @@ let do_delta_dom_codom reso sub = subst_dom_codom_delta_resolver sub reso let subst_signature subst = subst_signature subst do_delta_codom let subst_structure subst = subst_structure subst do_delta_codom -(** {6 Retroknowledge } *) - -(* spiwack: here comes the function which takes care of importing - the retroknowledge declared in the library *) -(* lclrk : retroknowledge_action list, rkaction : retroknowledge action *) -let add_retroknowledge = - let perform rkaction env = match rkaction with - | Retroknowledge.RKRegister (f, ((GlobRef.ConstRef _ | GlobRef.IndRef _) as e)) -> - Environ.register env f e - | _ -> - CErrors.anomaly ~label:"Modops.add_retroknowledge" - (Pp.str "had to import an unsupported kind of term.") - in - fun (ModBodyRK lclrk) env -> - (* The order of the declaration matters, for instance (and it's at the - time this comment is being written, the only relevent instance) the - int31 type registration absolutely needs int31 bits to be registered. - Since the local_retroknowledge is stored in reverse order (each new - registration is added at the top of the list) we need a fold_right - for things to go right (the pun is not intented). So we lose - tail recursivity, but the world will have exploded before any module - imports 10 000 retroknowledge registration.*) - List.fold_right perform lclrk env - (** {6 Adding a module in the environment } *) +let add_retroknowledge r env = + match r with + | ModBodyRK l -> List.fold_left Primred.add_retroknowledge env l + let rec add_structure mp sign resolver linkinfo env = let add_one env (l,elem) = match elem with |SFBconst cb -> @@ -399,7 +392,7 @@ let inline_delta_resolver env inl mp mbid mtb delta = let constant = lookup_constant con env in let l = make_inline delta r in match constant.const_body with - | Undef _ | OpaqueDef _ -> l + | Undef _ | OpaqueDef _ | Primitive _ -> l | Def body -> let constr = Mod_subst.force_constr body in let ctx = Declareops.constant_polymorphic_context constant in diff --git a/kernel/modops.mli b/kernel/modops.mli index 0acd09fb12..bb97f0171e 100644 --- a/kernel/modops.mli +++ b/kernel/modops.mli @@ -57,6 +57,8 @@ val add_linked_module : module_body -> link_info -> env -> env (** same, for a module type *) val add_module_type : ModPath.t -> module_type_body -> env -> env +val add_retroknowledge : module_implementation module_retroknowledge -> env -> env + (** {6 Strengthening } *) val strengthen : module_type_body -> ModPath.t -> module_type_body diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml index 482a2f3a3c..c32bdb85d6 100644 --- a/kernel/nativecode.ml +++ b/kernel/nativecode.ml @@ -14,7 +14,6 @@ open Constr open Declarations open Util open Nativevalues -open Nativeinstr open Nativelambda open Environ @@ -286,8 +285,6 @@ type primitive = | Mk_int | Mk_bool | Val_to_int - | Mk_I31_accu - | Decomp_uint | Mk_meta | Mk_evar | MLand @@ -305,7 +302,7 @@ type primitive = | MLmagic | MLarrayget | Mk_empty_instance - | Coq_primitive of CPrimitives.t * (prefix * Constant.t) option + | Coq_primitive of CPrimitives.t * (prefix * pconstant) option let eq_primitive p1 p2 = match p1, p2 with @@ -351,29 +348,27 @@ let primitive_hash = function | Mk_int -> 16 | Mk_bool -> 17 | Val_to_int -> 18 - | Mk_I31_accu -> 19 - | Decomp_uint -> 20 - | Mk_meta -> 21 - | Mk_evar -> 22 - | MLand -> 23 - | MLle -> 24 - | MLlt -> 25 - | MLinteq -> 26 - | MLlsl -> 27 - | MLlsr -> 28 - | MLland -> 29 - | MLlor -> 30 - | MLlxor -> 31 - | MLadd -> 32 - | MLsub -> 33 - | MLmul -> 34 - | MLmagic -> 35 - | Coq_primitive (prim, None) -> combinesmall 36 (CPrimitives.hash prim) - | Coq_primitive (prim, Some (prefix,kn)) -> - combinesmall 37 (combine3 (String.hash prefix) (Constant.hash kn) (CPrimitives.hash prim)) - | Mk_proj -> 38 - | MLarrayget -> 39 - | Mk_empty_instance -> 40 + | Mk_meta -> 19 + | Mk_evar -> 20 + | MLand -> 21 + | MLle -> 22 + | MLlt -> 23 + | MLinteq -> 24 + | MLlsl -> 25 + | MLlsr -> 26 + | MLland -> 27 + | MLlor -> 28 + | MLlxor -> 29 + | MLadd -> 30 + | MLsub -> 31 + | MLmul -> 32 + | MLmagic -> 33 + | Coq_primitive (prim, None) -> combinesmall 34 (CPrimitives.hash prim) + | Coq_primitive (prim, Some (prefix,(kn,_))) -> + combinesmall 35 (combine3 (String.hash prefix) (Constant.hash kn) (CPrimitives.hash prim)) + | Mk_proj -> 36 + | MLarrayget -> 37 + | Mk_empty_instance -> 38 type mllambda = | MLlocal of lname @@ -389,7 +384,7 @@ type mllambda = | MLconstruct of string * constructor * mllambda array (* prefix, constructor name, arguments *) | MLint of int - | MLuint of Uint31.t + | MLuint of Uint63.t | MLsetref of string * mllambda | MLsequence of mllambda * mllambda | MLarray of mllambda array @@ -455,7 +450,7 @@ let rec eq_mllambda gn1 gn2 n env1 env2 t1 t2 = | MLint i1, MLint i2 -> Int.equal i1 i2 | MLuint i1, MLuint i2 -> - Uint31.equal i1 i2 + Uint63.equal i1 i2 | MLsetref (id1, ml1), MLsetref (id2, ml2) -> String.equal id1 id2 && eq_mllambda gn1 gn2 n env1 env2 ml1 ml2 @@ -534,7 +529,7 @@ let rec hash_mllambda gn n env t = | MLint i -> combinesmall 11 i | MLuint i -> - combinesmall 12 (Uint31.to_int i) + combinesmall 12 (Uint63.hash i) | MLsetref (id, ml) -> let hid = String.hash id in let hml = hash_mllambda gn n env ml in @@ -947,9 +942,10 @@ let merge_branches t = Array.iter (fun (c,args,body) -> insert (c,args) body newt) t; Array.of_list (to_list newt) +let app_prim p args = MLapp(MLprimitive p, args) -type prim_aux = - | PAprim of string * Constant.t * CPrimitives.t * prim_aux array +type prim_aux = + | PAprim of string * pconstant * CPrimitives.t * prim_aux array | PAml of mllambda let add_check cond args = @@ -962,97 +958,67 @@ let add_check cond args = | _ -> cond in Array.fold_left aux cond args - + let extract_prim ml_of l = let decl = ref [] in let cond = ref [] in - let rec aux l = + let rec aux l = match l with | Lprim(prefix,kn,p,args) -> - let args = Array.map aux args in - cond := add_check !cond args; - PAprim(prefix,kn,p,args) + let args = Array.map aux args in + cond := add_check !cond args; + PAprim(prefix,kn,p,args) | Lrel _ | Lvar _ | Luint _ | Lval _ | Lconst _ -> PAml (ml_of l) - | _ -> - let x = fresh_lname Anonymous in - decl := (x,ml_of l)::!decl; - PAml (MLlocal x) in + | _ -> + let x = fresh_lname Anonymous in + decl := (x,ml_of l)::!decl; + PAml (MLlocal x) in let res = aux l in (!decl, !cond, res) -let app_prim p args = MLapp(MLprimitive p, args) - -let to_int v = +let cast_to_int v = match v with - | MLapp(MLprimitive Mk_uint, t) -> - begin match t.(0) with - | MLuint i -> MLint (Uint31.to_int i) - | _ -> MLapp(MLprimitive Val_to_int, [|v|]) - end - | MLapp(MLprimitive Mk_int, t) -> t.(0) - | _ -> MLapp(MLprimitive Val_to_int, [|v|]) - -let of_int v = - match v with - | MLapp(MLprimitive Val_to_int, t) -> t.(0) - | _ -> MLapp(MLprimitive Mk_int,[|v|]) + | MLint _ -> v + | _ -> MLapp(MLprimitive Val_to_int, [|v|]) let compile_prim decl cond paux = -(* - let args_to_int args = - for i = 0 to Array.length args - 1 do - args.(i) <- to_int args.(i) - done; - args in - *) + let rec opt_prim_aux paux = match paux with | PAprim(_prefix, _kn, op, args) -> - let args = Array.map opt_prim_aux args in - app_prim (Coq_primitive(op,None)) args -(* - TODO: check if this inlining was useful - begin match op with - | Int31lt -> - if Sys.word_size = 64 then - app_prim Mk_bool [|(app_prim MLlt (args_to_int args))|] - else app_prim (Coq_primitive (CPrimitives.Int31lt,None)) args - | Int31le -> - if Sys.word_size = 64 then - app_prim Mk_bool [|(app_prim MLle (args_to_int args))|] - else app_prim (Coq_primitive (CPrimitives.Int31le, None)) args - | Int31lsl -> of_int (mk_lsl (args_to_int args)) - | Int31lsr -> of_int (mk_lsr (args_to_int args)) - | Int31land -> of_int (mk_land (args_to_int args)) - | Int31lor -> of_int (mk_lor (args_to_int args)) - | Int31lxor -> of_int (mk_lxor (args_to_int args)) - | Int31add -> of_int (mk_add (args_to_int args)) - | Int31sub -> of_int (mk_sub (args_to_int args)) - | Int31mul -> of_int (mk_mul (args_to_int args)) - | _ -> app_prim (Coq_primitive(op,None)) args - end *) - | PAml ml -> ml - and naive_prim_aux paux = + let args = Array.map opt_prim_aux args in + app_prim (Coq_primitive(op,None)) args + | PAml ml -> ml + + and naive_prim_aux paux = match paux with | PAprim(prefix, kn, op, args) -> - app_prim (Coq_primitive(op, Some (prefix, kn))) (Array.map naive_prim_aux args) - | PAml ml -> ml in + app_prim (Coq_primitive(op, Some (prefix,kn))) (Array.map naive_prim_aux args) + | PAml ml -> ml + in - let compile_cond cond paux = + let compile_cond cond paux = match cond with - | [] -> opt_prim_aux paux + | [] -> opt_prim_aux paux | [c1] -> - MLif(app_prim Is_int [|c1|], opt_prim_aux paux, naive_prim_aux paux) + 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; to_int c|]) - (app_prim MLland [|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 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 add_decl decl body = List.fold_left (fun body (x,d) -> MLlet(x,d,body)) body decl in - add_decl decl (compile_cond cond paux) + + (* The optimizations done for checking if integer values are closed are valid + only on 64-bit architectures. So on 32-bit architectures, we fall back to less optimized checks. *) + if max_int = 1073741823 (* 32-bits *) then + add_decl decl (naive_prim_aux paux) + else + add_decl decl (compile_cond cond paux) let ml_of_instance instance u = let ml_of_level l = @@ -1089,6 +1055,11 @@ let ml_of_instance instance u = | Llam(ids,body) -> let lnames,env = push_rels env ids in MLlam(lnames, ml_of_lam env l body) + | Lrec(id,body) -> + let ids,body = decompose_Llam body in + let lname, env = push_rel env id in + let lnames, env = push_rels env ids in + MLletrec([|lname, lnames, ml_of_lam env l body|], MLlocal lname) | Llet(id,def,body) -> let def = ml_of_lam env l def in let lname, env = push_rel env id in @@ -1101,8 +1072,8 @@ let ml_of_instance instance u = mkMLapp (MLglobal(Gconstant (prefix, c))) args | Lproj (prefix, ind, i) -> MLglobal(Gproj (prefix, ind, i)) | Lprim _ -> - let decl,cond,paux = extract_prim (ml_of_lam env l) t in - compile_prim decl cond paux + let decl,cond,paux = extract_prim (ml_of_lam env l) t in + compile_prim decl cond paux | Lcase (annot,p,a,bs) -> (* let predicate_uid fv_pred = compilation of p let rec case_uid fv a_uid = @@ -1311,18 +1282,7 @@ let ml_of_instance instance u = | Lconstruct (prefix, (cn,u)) -> let uargs = ml_of_instance env.env_univ u in mkMLapp (MLglobal (Gconstruct (prefix, cn))) uargs - | Luint v -> - (match v with - | UintVal i -> MLapp(MLprimitive Mk_uint, [|MLuint i|]) - | UintDigits (prefix,cn,ds) -> - let c = MLglobal (Gconstruct (prefix, cn)) in - let ds = Array.map (ml_of_lam env l) ds in - let i31 = MLapp (MLprimitive Mk_I31_accu, [|c|]) in - MLapp(i31, ds) - | UintDecomp (prefix,cn,t) -> - let c = MLglobal (Gconstruct (prefix, cn)) in - let t = ml_of_lam env l t in - MLapp (MLprimitive Decomp_uint, [|c;t|])) + | Luint i -> MLapp(MLprimitive Mk_uint, [|MLuint i|]) | Lval v -> let i = push_symbol (SymbValue v) in get_value_code i | Lsort s -> @@ -1646,7 +1606,7 @@ let pp_mllam fmt l = Format.fprintf fmt "@[(Obj.magic (%s%a) : Nativevalues.t)@]" (string_of_construct prefix c) pp_cargs args | MLint i -> pp_int fmt i - | MLuint i -> Format.fprintf fmt "(Uint31.of_int %a)" pp_int (Uint31.to_int i) + | MLuint i -> Format.fprintf fmt "(%s)" (Uint63.compile i) | MLsetref (s, body) -> Format.fprintf fmt "@[%s@ :=@\n %a@]" s pp_mllam body | MLsequence(l1,l2) -> @@ -1766,8 +1726,6 @@ let pp_mllam fmt l = | Mk_int -> Format.fprintf fmt "mk_int" | Mk_bool -> Format.fprintf fmt "mk_bool" | Val_to_int -> Format.fprintf fmt "val_to_int" - | Mk_I31_accu -> Format.fprintf fmt "mk_I31_accu" - | Decomp_uint -> Format.fprintf fmt "decomp_uint" | Mk_meta -> Format.fprintf fmt "mk_meta_accu" | Mk_evar -> Format.fprintf fmt "mk_evar_accu" | MLand -> Format.fprintf fmt "(&&)" @@ -1787,9 +1745,9 @@ let pp_mllam fmt l = | Mk_empty_instance -> Format.fprintf fmt "Univ.Instance.empty" | Coq_primitive (op,None) -> Format.fprintf fmt "no_check_%s" (CPrimitives.to_string op) - | Coq_primitive (op, Some (prefix,kn)) -> + | Coq_primitive (op, Some (prefix,(c,_))) -> Format.fprintf fmt "%s %a" (CPrimitives.to_string op) - pp_mllam (MLglobal (Gconstant (prefix, kn))) + pp_mllam (MLglobal (Gconstant (prefix,c))) in Format.fprintf fmt "@[%a@]" pp_mllam l @@ -1903,7 +1861,7 @@ let compile_constant env sigma prefix ~interactive con cb = let t = Mod_subst.force_constr t in let code = lambda_of_constr env sigma t in if !Flags.debug then Feedback.msg_debug (Pp.str "Generated lambda code"); - let is_lazy = is_lazy env prefix t in + let is_lazy = is_lazy t in let code = if is_lazy then mk_lazy code else code in let name = if interactive then LinkedInteractive prefix diff --git a/kernel/nativeconv.ml b/kernel/nativeconv.ml index f5d7ab3c9d..baa290367f 100644 --- a/kernel/nativeconv.ml +++ b/kernel/nativeconv.ml @@ -34,6 +34,8 @@ let rec conv_val env pb lvl v1 v2 cu = conv_accu env pb lvl k1 k2 cu | Vconst i1, Vconst i2 -> if Int.equal i1 i2 then cu else raise NotConvertible + | Vint64 i1, Vint64 i2 -> + if Int64.equal i1 i2 then cu else raise NotConvertible | Vblock b1, Vblock b2 -> let n1 = block_size b1 in let n2 = block_size b2 in @@ -47,7 +49,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 _, _ | Vblock _, _ -> raise NotConvertible + | Vaccu _, _ | Vconst _, _ | Vint64 _, _ | Vblock _, _ -> raise NotConvertible and conv_accu env pb lvl k1 k2 cu = let n1 = accu_nargs k1 in diff --git a/kernel/nativeinstr.mli b/kernel/nativeinstr.mli deleted file mode 100644 index 2d8e2ba2f0..0000000000 --- a/kernel/nativeinstr.mli +++ /dev/null @@ -1,59 +0,0 @@ -(************************************************************************) -(* * 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) *) -(************************************************************************) -open Names -open Constr -open Nativevalues - -(** This file defines the lambda code for the native compiler. It has been -extracted from Nativelambda.ml because of the retroknowledge architecture. *) - -type prefix = string - -type uint = - | UintVal of Uint31.t - | UintDigits of prefix * constructor * lambda array - | UintDecomp of prefix * constructor * lambda - -and lambda = - | Lrel of Name.t * int - | Lvar of Id.t - | Lmeta of metavariable * lambda (* type *) - | Levar of Evar.t * lambda array (* arguments *) - | Lprod of lambda * lambda - | Llam of Name.t array * lambda - | Llet of Name.t * lambda * lambda - | Lapp of lambda * lambda array - | Lconst of prefix * pconstant - | Lproj of prefix * inductive * int (* prefix, inductive, index starting from 0 *) - | Lprim of prefix * Constant.t * CPrimitives.t * lambda array - | Lcase of annot_sw * lambda * lambda * lam_branches - (* annotations, term being matched, accu, branches *) - | Lif of lambda * lambda * lambda - | Lfix of (int array * (string * inductive) array * int) * fix_decl - | Lcofix of int * fix_decl (* must be in eta-expanded form *) - | Lmakeblock of prefix * pconstructor * int * lambda array - (* prefix, constructor name, constructor tag, arguments *) - (* A fully applied constructor *) - | Lconstruct of prefix * pconstructor - (* A partially applied constructor *) - | Luint of uint - | Lval of Nativevalues.t - | Lsort of Sorts.t - | Lind of prefix * pinductive - | Llazy - | Lforce - -(* Cofixpoints have to be in eta-expanded form for their call-by-need evaluation -to be correct. Otherwise, memoization of previous evaluations will be applied -again to extra arguments (see #7333). *) - -and lam_branches = (constructor * Name.t array * lambda) array - -and fix_decl = Name.t array * lambda array * lambda array diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml index 70cb8691c6..0869f94042 100644 --- a/kernel/nativelambda.ml +++ b/kernel/nativelambda.ml @@ -14,12 +14,46 @@ open Constr open Declarations open Environ open Nativevalues -open Nativeinstr module RelDecl = Context.Rel.Declaration - -exception NotClosed +(** This file defines the lambda code generation phase of the native compiler *) +type prefix = string + +type lambda = + | Lrel of Name.t * int + | Lvar of Id.t + | Lmeta of metavariable * lambda (* type *) + | Levar of Evar.t * lambda array (* arguments *) + | Lprod of lambda * lambda + | Llam of Name.t array * lambda + | Lrec of Name.t * lambda + | Llet of Name.t * lambda * lambda + | Lapp of lambda * lambda array + | Lconst of prefix * pconstant + | Lproj of prefix * inductive * int (* prefix, inductive, index starting from 0 *) + | Lprim of prefix * pconstant * CPrimitives.t * lambda array + (* No check if None *) + | Lcase of annot_sw * lambda * lambda * lam_branches + (* annotations, term being matched, accu, branches *) + | Lif of lambda * lambda * lambda + | Lfix of (int array * (string * inductive) array * int) * fix_decl + | Lcofix of int * fix_decl + | Lmakeblock of prefix * pconstructor * int * lambda array + (* prefix, constructor Name.t, constructor tag, arguments *) + (* A fully applied constructor *) + | Lconstruct of prefix * pconstructor (* prefix, constructor Name.t *) + (* A partially applied constructor *) + | Luint of Uint63.t + | Lval of Nativevalues.t + | Lsort of Sorts.t + | Lind of prefix * pinductive + | Llazy + | Lforce + +and lam_branches = (constructor * Name.t array * lambda) array + +and fix_decl = Name.t array * lambda array * lambda array type evars = { evars_val : existential -> constr option; @@ -84,9 +118,9 @@ let get_const_prefix env c = (* A generic map function *) -let rec map_lam_with_binders g f n lam = +let map_lam_with_binders g f n lam = match lam with - | Lrel _ | Lvar _ | Lconst _ | Lproj _ | Lval _ | Lsort _ | Lind _ + | Lrel _ | Lvar _ | Lconst _ | Lproj _ | Lval _ | Lsort _ | Lind _ | Luint _ | Lconstruct _ | Llazy | Lforce | Lmeta _ -> lam | Lprod(dom,codom) -> let dom' = f n dom in @@ -95,6 +129,9 @@ let rec map_lam_with_binders g f n lam = | Llam(ids,body) -> let body' = f (g (Array.length ids) n) body in if body == body' then lam else mkLlam ids body' + | Lrec(id,body) -> + let body' = f (g 1 n) body in + if body == body' then lam else Lrec(id,body') | Llet(id,def,body) -> let def' = f n def in let body' = f (g 1 n) body in @@ -135,23 +172,10 @@ let rec map_lam_with_binders g f n lam = | Lmakeblock(prefix,cn,tag,args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Lmakeblock(prefix,cn,tag,args') - | Luint u -> - let u' = map_uint g f n u in - if u == u' then lam else Luint u' | Levar (evk, args) -> let args' = Array.Smart.map (f n) args in if args == args' then lam else Levar (evk, args') -and map_uint _g f n u = - match u with - | UintVal _ -> u - | UintDigits(prefix,c,args) -> - let args' = Array.Smart.map (f n) args in - if args == args' then u else UintDigits(prefix,c,args') - | UintDecomp(prefix,c,a) -> - let a' = f n a in - if a == a' then u else UintDecomp(prefix,c,a') - (*s Lift and substitution *) let rec lam_exlift el lam = @@ -186,7 +210,7 @@ let lam_subst_args subst args = (* [simplify subst lam] simplify the expression [lam_subst subst lam] *) (* that is : *) (* - Reduce [let] is the definition can be substituted i.e: *) -(* - a variable (rel or identifier) *) +(* - a variable (rel or Id.t) *) (* - a constant *) (* - a structured constant *) (* - a function *) @@ -298,7 +322,7 @@ let is_value lc = match lc with | Lval _ -> true | Lmakeblock(_,_,_,args) when Array.is_empty args -> true - | Luint (UintVal _) -> true + | Luint _ -> true | _ -> false let get_value lc = @@ -306,7 +330,7 @@ let get_value lc = | Lval v -> v | Lmakeblock(_,_,tag,args) when Array.is_empty args -> Nativevalues.mk_int tag - | Luint (UintVal i) -> Nativevalues.mk_uint i + | Luint i -> Nativevalues.mk_uint i | _ -> raise Not_found let make_args start _end = @@ -333,6 +357,20 @@ let rec get_alias env (kn, u as p) = | Cemitcodes.BCalias kn' -> get_alias env (kn', u) | _ -> p +let prim env kn p args = + let prefix = get_const_prefix env (fst kn) in + Lprim(prefix, kn, p, args) + +let expand_prim env kn op arity = + let ids = Array.make arity Anonymous in + let args = make_args arity 1 in + Llam(ids, prim env kn op args) + +let lambda_of_prim env kn op args = + let arity = CPrimitives.arity op in + if Array.length args >= arity then prim env kn op args + else mkLapp (expand_prim env kn op arity) args + (*i Global environment *) let get_names decl = @@ -368,22 +406,9 @@ module Cache = r end -let is_lazy env prefix t = - match kind t with - | App (f,_args) -> - begin match kind f with - | Construct (c,_) -> - let gr = GlobRef.IndRef (fst c) in - (try - let _ = - Retroknowledge.get_native_before_match_info env.retroknowledge - gr prefix c Llazy; - in - false - with Not_found -> true) - | _ -> true - end - | LetIn _ | Case _ | Proj _ -> true +let is_lazy t = + match Constr.kind t with + | App _ | LetIn _ | Case _ | Proj _ -> true | _ -> false let evar_value sigma ev = sigma.evars_val ev @@ -482,13 +507,6 @@ let rec lambda_of_constr cache env sigma c = in (* translation of the argument *) let la = lambda_of_constr cache env sigma a in - let gr = GlobRef.IndRef ind in - let la = - try - Retroknowledge.get_native_before_match_info (env).retroknowledge - gr prefix (ind,1) la - with Not_found -> la - in (* translation of the type *) let lt = lambda_of_constr cache env sigma t in (* translation of branches *) @@ -519,7 +537,7 @@ let rec lambda_of_constr cache env sigma c = let env = Environ.push_rec_types (names, type_bodies, rec_bodies) env in let lbodies = lambda_of_args cache env sigma 0 rec_bodies in Lfix((pos, inds, i), (names, ltypes, lbodies)) - + | CoFix(init,(names,type_bodies,rec_bodies)) -> let rec_bodies = Array.map2 (Reduction.eta_expand env) rec_bodies type_bodies in let ltypes = lambda_of_args cache env sigma 0 type_bodies in @@ -527,27 +545,22 @@ let rec lambda_of_constr cache env sigma c = let lbodies = lambda_of_args cache env sigma 0 rec_bodies in Lcofix(init, (names, ltypes, lbodies)) + | Int i -> Luint i + and lambda_of_app cache env sigma f args = match kind f with | Const (_kn,_u as c) -> let kn,u = get_alias env c in let cb = lookup_constant kn env in - (try - let prefix = get_const_prefix env kn in - (* We delay the compilation of arguments to avoid an exponential behavior *) - let f = Retroknowledge.get_native_compiling_info - (env).retroknowledge (GlobRef.ConstRef kn) prefix in - let args = lambda_of_args cache env sigma 0 args in - f args - with Not_found -> begin match cb.const_body with + | Primitive op -> lambda_of_prim env c op (lambda_of_args cache env sigma 0 args) | Def csubst -> (* TODO optimize if f is a proj and argument is known *) if cb.const_inline_code then lambda_of_app cache env sigma (Mod_subst.force_constr csubst) args else let prefix = get_const_prefix env kn in let t = - if is_lazy env prefix (Mod_subst.force_constr csubst) then + if is_lazy (Mod_subst.force_constr csubst) then mkLapp Lforce [|Lconst (prefix, (kn,u))|] else Lconst (prefix, (kn,u)) in @@ -555,34 +568,18 @@ and lambda_of_app cache env sigma f args = | OpaqueDef _ | Undef _ -> let prefix = get_const_prefix env kn in mkLapp (Lconst (prefix, (kn,u))) (lambda_of_args cache env sigma 0 args) - end) + end | Construct (c,u) -> let tag, nparams, arity = Cache.get_construct_info cache env c in let expected = nparams + arity in let nargs = Array.length args in let prefix = get_mind_prefix env (fst (fst c)) in - let gr = GlobRef.ConstructRef c in if Int.equal nargs expected then - try - try - Retroknowledge.get_native_constant_static_info - (env).retroknowledge - gr args - with NotClosed -> - assert (Int.equal nparams 0); (* should be fine for int31 *) - let args = lambda_of_args cache env sigma nparams args in - Retroknowledge.get_native_constant_dynamic_info - (env).retroknowledge gr prefix c args - with Not_found -> let args = lambda_of_args cache env sigma nparams args in makeblock env c u tag args else let args = lambda_of_args cache env sigma 0 args in - (try - Retroknowledge.get_native_constant_dynamic_info - (env).retroknowledge gr prefix c args - with Not_found -> - mkLapp (Lconstruct (prefix, (c,u))) args) + mkLapp (Lconstruct (prefix, (c,u))) args | _ -> let f = lambda_of_constr cache env sigma f in let args = lambda_of_args cache env sigma 0 args in @@ -615,45 +612,3 @@ let lambda_of_constr env sigma c = let mk_lazy c = mkLapp Llazy [|c|] - -(** Retroknowledge, to be removed once we move to primitive machine integers *) -let compile_static_int31 fc args = - if not fc then raise Not_found else - Luint (UintVal - (Uint31.of_int (Array.fold_left - (fun temp_i -> fun t -> match kind t with - | Construct ((_,d),_) -> 2*temp_i+d-1 - | _ -> raise NotClosed) - 0 args))) - -let compile_dynamic_int31 fc prefix c args = - if not fc then raise Not_found else - Luint (UintDigits (prefix,c,args)) - -(* We are relying here on the order of digits constructors *) -let digits_from_uint digits_ind prefix i = - let d0 = Lconstruct (prefix, ((digits_ind, 1), Univ.Instance.empty)) in - let d1 = Lconstruct (prefix, ((digits_ind, 2), Univ.Instance.empty)) in - let digits = Array.make 31 d0 in - for k = 0 to 30 do - if Int.equal ((Uint31.to_int i lsr k) land 1) 1 then - digits.(30-k) <- d1 - done; - digits - -let before_match_int31 digits_ind fc prefix c t = - if not fc then - raise Not_found - else - match t with - | Luint (UintVal i) -> - let digits = digits_from_uint digits_ind prefix i in - mkLapp (Lconstruct (prefix,(c, Univ.Instance.empty))) digits - | Luint (UintDigits (prefix,c,args)) -> - mkLapp (Lconstruct (prefix,(c, Univ.Instance.empty))) args - | _ -> Luint (UintDecomp (prefix,c,t)) - -let compile_prim prim kn fc prefix args = - if not fc then raise Not_found - else - Lprim(prefix, kn, prim, args) diff --git a/kernel/nativelambda.mli b/kernel/nativelambda.mli index 7b20258929..eb06522a33 100644 --- a/kernel/nativelambda.mli +++ b/kernel/nativelambda.mli @@ -10,9 +10,45 @@ open Names open Constr open Environ -open Nativeinstr +open Nativevalues (** This file defines the lambda code generation phase of the native compiler *) +type prefix = string + +type lambda = + | Lrel of Name.t * int + | Lvar of Id.t + | Lmeta of metavariable * lambda (* type *) + | Levar of Evar.t * lambda array (* arguments *) + | Lprod of lambda * lambda + | Llam of Name.t array * lambda + | Lrec of Name.t * lambda + | Llet of Name.t * lambda * lambda + | Lapp of lambda * lambda array + | Lconst of prefix * pconstant + | Lproj of prefix * inductive * int (* prefix, inductive, index starting from 0 *) + | Lprim of prefix * pconstant * CPrimitives.t * lambda array + | Lcase of annot_sw * lambda * lambda * lam_branches + (* annotations, term being matched, accu, branches *) + | Lif of lambda * lambda * lambda + | Lfix of (int array * (string * inductive) array * int) * fix_decl + | Lcofix of int * fix_decl + | Lmakeblock of prefix * pconstructor * int * lambda array + (* prefix, constructor Name.t, constructor tag, arguments *) + (* A fully applied constructor *) + | Lconstruct of prefix * pconstructor (* prefix, constructor Name.t *) + (* A partially applied constructor *) + | Luint of Uint63.t + | Lval of Nativevalues.t + | Lsort of Sorts.t + | Lind of prefix * pinductive + | Llazy + | Lforce + +and lam_branches = (constructor * Name.t array * lambda) array + +and fix_decl = Name.t array * lambda array * lambda array + type evars = { evars_val : existential -> constr option; evars_metas : metavariable -> types } @@ -22,7 +58,7 @@ val empty_evars : evars val decompose_Llam : lambda -> Name.t array * lambda val decompose_Llam_Llet : lambda -> (Name.t * lambda option) array * lambda -val is_lazy : env -> prefix -> constr -> bool +val is_lazy : constr -> bool val mk_lazy : lambda -> lambda val get_mind_prefix : env -> MutInd.t -> string @@ -30,14 +66,3 @@ val get_mind_prefix : env -> MutInd.t -> string val get_alias : env -> pconstant -> pconstant val lambda_of_constr : env -> evars -> Constr.constr -> lambda - -val compile_static_int31 : bool -> Constr.constr array -> lambda - -val compile_dynamic_int31 : bool -> prefix -> constructor -> lambda array -> - lambda - -val before_match_int31 : inductive -> bool -> prefix -> constructor -> lambda -> - lambda - -val compile_prim : CPrimitives.t -> Constant.t -> bool -> prefix -> lambda array -> - lambda diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml index 93e74af845..a6b48cd7e3 100644 --- a/kernel/nativevalues.ml +++ b/kernel/nativevalues.ml @@ -196,11 +196,17 @@ let dummy_value : unit -> t = fun () _ -> anomaly ~label:"native" (Pp.str "Evaluation failed.") let cast_accu v = (Obj.magic v:accumulator) +[@@ocaml.inline always] let mk_int (x : int) = (Obj.magic x : t) +[@@ocaml.inline always] + (* Coq's booleans are reversed... *) let mk_bool (b : bool) = (Obj.magic (not b) : t) -let mk_uint (x : Uint31.t) = (Obj.magic x : t) +[@@ocaml.inline always] + +let mk_uint (x : Uint63.t) = (Obj.magic x : t) +[@@ocaml.inline always] type block @@ -216,8 +222,9 @@ type kind_of_value = | Vaccu of accumulator | Vfun of (t -> t) | Vconst of int + | Vint64 of int64 | 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) @@ -225,8 +232,8 @@ let kind_of_value (v:t) = let tag = Obj.tag o in if Int.equal tag accumulate_tag then Vaccu (Obj.magic v) - else - if (tag < Obj.lazy_tag) then Vblock (Obj.magic v) + else if Int.equal tag Obj.custom_tag then Vint64 (Obj.magic v) + else if (tag < Obj.lazy_tag) then Vblock (Obj.magic v) else (* assert (tag = Obj.closure_tag || tag = Obj.infix_tag); or ??? what is 1002*) @@ -236,92 +243,105 @@ let kind_of_value (v:t) = let is_int (x:t) = let o = Obj.repr x in - Obj.is_int o + Obj.is_int o || Int.equal (Obj.tag o) Obj.custom_tag let val_to_int (x:t) = (Obj.magic x : int) +[@@ocaml.inline always] -let to_uint (x:t) = (Obj.magic x : Uint31.t) -let of_uint (x: Uint31.t) = (Obj.magic x : t) +let to_uint (x:t) = (Obj.magic x : Uint63.t) +[@@ocaml.inline always] let no_check_head0 x = - of_uint (Uint31.head0 (to_uint x)) + mk_uint (Uint63.head0 (to_uint x)) +[@@ocaml.inline always] let head0 accu x = if is_int x then no_check_head0 x else accu x let no_check_tail0 x = - of_uint (Uint31.tail0 (to_uint x)) + mk_uint (Uint63.tail0 (to_uint x)) +[@@ocaml.inline always] let tail0 accu x = if is_int x then no_check_tail0 x else accu x let no_check_add x y = - of_uint (Uint31.add (to_uint x) (to_uint y)) + mk_uint (Uint63.add (to_uint x) (to_uint y)) +[@@ocaml.inline always] let add accu x y = if is_int x && is_int y then no_check_add x y else accu x y let no_check_sub x y = - of_uint (Uint31.sub (to_uint x) (to_uint y)) + mk_uint (Uint63.sub (to_uint x) (to_uint y)) +[@@ocaml.inline always] let sub accu x y = if is_int x && is_int y then no_check_sub x y else accu x y let no_check_mul x y = - of_uint (Uint31.mul (to_uint x) (to_uint y)) + mk_uint (Uint63.mul (to_uint x) (to_uint y)) +[@@ocaml.inline always] let mul accu x y = if is_int x && is_int y then no_check_mul x y else accu x y let no_check_div x y = - of_uint (Uint31.div (to_uint x) (to_uint y)) + mk_uint (Uint63.div (to_uint x) (to_uint y)) +[@@ocaml.inline always] let div accu x y = if is_int x && is_int y then no_check_div x y else accu x y let no_check_rem x y = - of_uint (Uint31.rem (to_uint x) (to_uint y)) + mk_uint (Uint63.rem (to_uint x) (to_uint y)) +[@@ocaml.inline always] let rem accu x y = if is_int x && is_int y then no_check_rem x y else accu x y let no_check_l_sr x y = - of_uint (Uint31.l_sr (to_uint x) (to_uint y)) + mk_uint (Uint63.l_sr (to_uint x) (to_uint y)) +[@@ocaml.inline always] let l_sr accu x y = if is_int x && is_int y then no_check_l_sr x y else accu x y let no_check_l_sl x y = - of_uint (Uint31.l_sl (to_uint x) (to_uint y)) + mk_uint (Uint63.l_sl (to_uint x) (to_uint y)) +[@@ocaml.inline always] let l_sl accu x y = if is_int x && is_int y then no_check_l_sl x y else accu x y let no_check_l_and x y = - of_uint (Uint31.l_and (to_uint x) (to_uint y)) + mk_uint (Uint63.l_and (to_uint x) (to_uint y)) +[@@ocaml.inline always] let l_and accu x y = if is_int x && is_int y then no_check_l_and x y else accu x y let no_check_l_xor x y = - of_uint (Uint31.l_xor (to_uint x) (to_uint y)) + mk_uint (Uint63.l_xor (to_uint x) (to_uint y)) +[@@ocaml.inline always] let l_xor accu x y = if is_int x && is_int y then no_check_l_xor x y else accu x y let no_check_l_or x y = - of_uint (Uint31.l_or (to_uint x) (to_uint y)) + mk_uint (Uint63.l_or (to_uint x) (to_uint y)) +[@@ocaml.inline always] let l_or accu x y = if is_int x && is_int y then no_check_l_or x y @@ -337,61 +357,57 @@ type coq_pair = | Paccu of t | PPair of t * t -type coq_zn2z = - | Zaccu of t - | ZW0 - | ZWW of t * t - let mkCarry b i = - if b then (Obj.magic (C1(of_uint i)):t) - else (Obj.magic (C0(of_uint i)):t) + if b then (Obj.magic (C1(mk_uint i)):t) + else (Obj.magic (C0(mk_uint i)):t) let no_check_addc x y = - let s = Uint31.add (to_uint x) (to_uint y) in - mkCarry (Uint31.lt s (to_uint x)) s + let s = Uint63.add (to_uint x) (to_uint y) in + mkCarry (Uint63.lt s (to_uint x)) s +[@@ocaml.inline always] let addc accu x y = if is_int x && is_int y then no_check_addc x y else accu x y let no_check_subc x y = - let s = Uint31.sub (to_uint x) (to_uint y) in - mkCarry (Uint31.lt (to_uint x) (to_uint y)) s + let s = Uint63.sub (to_uint x) (to_uint y) in + mkCarry (Uint63.lt (to_uint x) (to_uint y)) s +[@@ocaml.inline always] let subc accu x y = if is_int x && is_int y then no_check_subc x y else accu x y -let no_check_addcarryc x y = +let no_check_addCarryC x y = let s = - Uint31.add (Uint31.add (to_uint x) (to_uint y)) - (Uint31.of_int 1) in - mkCarry (Uint31.le s (to_uint x)) s + Uint63.add (Uint63.add (to_uint x) (to_uint y)) + (Uint63.of_int 1) in + mkCarry (Uint63.le s (to_uint x)) s +[@@ocaml.inline always] -let addcarryc accu x y = - if is_int x && is_int y then no_check_addcarryc x y +let addCarryC accu x y = + if is_int x && is_int y then no_check_addCarryC x y else accu x y -let no_check_subcarryc x y = +let no_check_subCarryC x y = let s = - Uint31.sub (Uint31.sub (to_uint x) (to_uint y)) - (Uint31.of_int 1) in - mkCarry (Uint31.le (to_uint x) (to_uint y)) s + Uint63.sub (Uint63.sub (to_uint x) (to_uint y)) + (Uint63.of_int 1) in + mkCarry (Uint63.le (to_uint x) (to_uint y)) s +[@@ocaml.inline always] -let subcarryc accu x y = - if is_int x && is_int y then no_check_subcarryc x y +let subCarryC accu x y = + if is_int x && is_int y then no_check_subCarryC x y else accu x y let of_pair (x, y) = - (Obj.magic (PPair(of_uint x, of_uint y)):t) - -let zn2z_of_pair (x,y) = - if Uint31.equal x (Uint31.of_uint 0) && - Uint31.equal y (Uint31.of_uint 0) then Obj.magic ZW0 - else (Obj.magic (ZWW(of_uint x, of_uint y)) : t) + (Obj.magic (PPair(mk_uint x, mk_uint y)):t) +[@@ocaml.inline always] let no_check_mulc x y = - zn2z_of_pair(Uint31.mulc (to_uint x) (to_uint y)) + of_pair (Uint63.mulc (to_uint x) (to_uint y)) +[@@ocaml.inline always] let mulc accu x y = if is_int x && is_int y then no_check_mulc x y @@ -399,7 +415,8 @@ let mulc accu x y = let no_check_diveucl x y = let i1, i2 = to_uint x, to_uint y in - of_pair(Uint31.div i1 i2, Uint31.rem i1 i2) + of_pair(Uint63.div i1 i2, Uint63.rem i1 i2) +[@@ocaml.inline always] let diveucl accu x y = if is_int x && is_int y then no_check_diveucl x y @@ -407,21 +424,20 @@ let diveucl accu x y = let no_check_div21 x y z = let i1, i2, i3 = to_uint x, to_uint y, to_uint z in - of_pair (Uint31.div21 i1 i2 i3) + of_pair (Uint63.div21 i1 i2 i3) +[@@ocaml.inline always] let div21 accu x y z = if is_int x && is_int y && is_int z then no_check_div21 x y z else accu x y z -let no_check_addmuldiv x y z = +let no_check_addMulDiv x y z = let p, i, j = to_uint x, to_uint y, to_uint z in - let p' = Uint31.to_int p in - of_uint (Uint31.l_or - (Uint31.l_sl i p) - (Uint31.l_sr j (Uint31.of_int (31 - p')))) + mk_uint (Uint63.addmuldiv p i j) +[@@ocaml.inline always] -let addmuldiv accu x y z = - if is_int x && is_int y && is_int z then no_check_addmuldiv x y z +let addMulDiv accu x y z = + if is_int x && is_int y && is_int z then no_check_addMulDiv x y z else accu x y z [@@@ocaml.warning "-34"] @@ -436,29 +452,32 @@ type coq_cmp = | CmpLt | CmpGt -let no_check_eq x y = - mk_bool (Uint31.equal (to_uint x) (to_uint y)) +let no_check_eq x y = + mk_bool (Uint63.equal (to_uint x) (to_uint y)) +[@@ocaml.inline always] let eq accu x y = if is_int x && is_int y then no_check_eq x y else accu x y let no_check_lt x y = - mk_bool (Uint31.lt (to_uint x) (to_uint y)) + mk_bool (Uint63.lt (to_uint x) (to_uint y)) +[@@ocaml.inline always] let lt accu x y = if is_int x && is_int y then no_check_lt x y else accu x y let no_check_le x y = - mk_bool (Uint31.le (to_uint x) (to_uint y)) + mk_bool (Uint63.le (to_uint x) (to_uint y)) +[@@ocaml.inline always] let le accu x y = if is_int x && is_int y then no_check_le x y else accu x y let no_check_compare x y = - match Uint31.compare (to_uint x) (to_uint y) with + match Uint63.compare (to_uint x) (to_uint y) with | x when x < 0 -> (Obj.magic CmpLt:t) | 0 -> (Obj.magic CmpEq:t) | _ -> (Obj.magic CmpGt:t) @@ -467,6 +486,11 @@ let compare accu x y = if is_int x && is_int y then no_check_compare x y else accu x y +let print x = + Printf.fprintf stderr "%s" (Uint63.to_string (to_uint x)); + flush stderr; + 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) - @@ -491,63 +515,3 @@ let str_decode s = Buffer.add_char mshl_expr (bin_of_hex (Bytes.to_string buf)) done; Marshal.from_bytes (Buffer.to_bytes mshl_expr) 0 - -(** Retroknowledge, to be removed when we switch to primitive integers *) - -(* This will be unsafe with 63-bits integers *) -let digit_to_uint d = (Obj.magic d : Uint31.t) - -let mk_I31_accu c x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 x15 x16 x17 - x18 x19 x20 x21 x22 x23 x24 x25 x26 x27 x28 x29 x30 = - if is_int x0 && is_int x1 && is_int x2 && is_int x3 && is_int x4 && is_int x5 - && is_int x6 && is_int x7 && is_int x8 && is_int x9 && is_int x10 - && is_int x11 && is_int x12 && is_int x13 && is_int x14 && is_int x15 - && is_int x16 && is_int x17 && is_int x18 && is_int x19 && is_int x20 - && is_int x21 && is_int x22 && is_int x23 && is_int x24 && is_int x25 - && is_int x26 && is_int x27 && is_int x28 && is_int x29 && is_int x30 - then - let r = digit_to_uint x0 in - let r = Uint31.add_digit r (digit_to_uint x1) in - let r = Uint31.add_digit r (digit_to_uint x2) in - let r = Uint31.add_digit r (digit_to_uint x3) in - let r = Uint31.add_digit r (digit_to_uint x4) in - let r = Uint31.add_digit r (digit_to_uint x5) in - let r = Uint31.add_digit r (digit_to_uint x6) in - let r = Uint31.add_digit r (digit_to_uint x7) in - let r = Uint31.add_digit r (digit_to_uint x8) in - let r = Uint31.add_digit r (digit_to_uint x9) in - let r = Uint31.add_digit r (digit_to_uint x10) in - let r = Uint31.add_digit r (digit_to_uint x11) in - let r = Uint31.add_digit r (digit_to_uint x12) in - let r = Uint31.add_digit r (digit_to_uint x13) in - let r = Uint31.add_digit r (digit_to_uint x14) in - let r = Uint31.add_digit r (digit_to_uint x15) in - let r = Uint31.add_digit r (digit_to_uint x16) in - let r = Uint31.add_digit r (digit_to_uint x17) in - let r = Uint31.add_digit r (digit_to_uint x18) in - let r = Uint31.add_digit r (digit_to_uint x19) in - let r = Uint31.add_digit r (digit_to_uint x20) in - let r = Uint31.add_digit r (digit_to_uint x21) in - let r = Uint31.add_digit r (digit_to_uint x22) in - let r = Uint31.add_digit r (digit_to_uint x23) in - let r = Uint31.add_digit r (digit_to_uint x24) in - let r = Uint31.add_digit r (digit_to_uint x25) in - let r = Uint31.add_digit r (digit_to_uint x26) in - let r = Uint31.add_digit r (digit_to_uint x27) in - let r = Uint31.add_digit r (digit_to_uint x28) in - let r = Uint31.add_digit r (digit_to_uint x29) in - let r = Uint31.add_digit r (digit_to_uint x30) in - mk_uint r - else - c x0 x1 x2 x3 x4 x5 x6 x7 x8 x9 x10 x11 x12 x13 x14 x15 x16 x17 x18 x19 x20 - x21 x22 x23 x24 x25 x26 x27 x28 x29 x30 - -let decomp_uint c v = - if is_int v then - let r = ref c in - let v = val_to_int v in - for i = 30 downto 0 do - r := (!r) (mk_int ((v lsr i) land 1)); - done; - !r - else v diff --git a/kernel/nativevalues.mli b/kernel/nativevalues.mli index 10689941e8..58cb6e2c30 100644 --- a/kernel/nativevalues.mli +++ b/kernel/nativevalues.mli @@ -78,8 +78,13 @@ val mk_const : tag -> t val mk_block : tag -> t array -> t val mk_bool : bool -> t +[@@ocaml.inline always] + val mk_int : int -> t -val mk_uint : Uint31.t -> t +[@@ocaml.inline always] + +val mk_uint : Uint63.t -> t +[@@ocaml.inline always] val napply : t -> t array -> t (* Functions over accumulators *) @@ -90,6 +95,8 @@ val args_of_accu : accumulator -> t array val accu_nargs : accumulator -> int val cast_accu : t -> accumulator +[@@ocaml.inline always] + (* Functions over block: i.e constructors *) type block @@ -106,6 +113,7 @@ type kind_of_value = | Vaccu of accumulator | Vfun of (t -> t) | Vconst of int + | Vint64 of int64 | Vblock of block val kind_of_value : t -> kind_of_value @@ -136,51 +144,90 @@ val l_or : t -> t -> t -> t val addc : t -> t -> t -> t val subc : t -> t -> t -> t -val addcarryc : t -> t -> t -> t -val subcarryc : t -> t -> t -> t +val addCarryC : t -> t -> t -> t +val subCarryC : t -> t -> t -> t val mulc : t -> t -> t -> t val diveucl : t -> t -> t -> t val div21 : t -> t -> t -> t -> t -val addmuldiv : t -> t -> t -> t -> t +val addMulDiv : t -> t -> t -> t -> t val eq : t -> t -> t -> t val lt : t -> t -> t -> t val le : t -> t -> t -> t val compare : t -> t -> t -> t +val print : t -> t + (* Function without check *) val no_check_head0 : t -> t +[@@ocaml.inline always] + val no_check_tail0 : t -> t +[@@ocaml.inline always] val no_check_add : t -> t -> t +[@@ocaml.inline always] + val no_check_sub : t -> t -> t +[@@ocaml.inline always] + val no_check_mul : t -> t -> t +[@@ocaml.inline always] + val no_check_div : t -> t -> t +[@@ocaml.inline always] + val no_check_rem : t -> t -> t +[@@ocaml.inline always] val no_check_l_sr : t -> t -> t +[@@ocaml.inline always] + val no_check_l_sl : t -> t -> t +[@@ocaml.inline always] + val no_check_l_and : t -> t -> t +[@@ocaml.inline always] + val no_check_l_xor : t -> t -> t +[@@ocaml.inline always] + val no_check_l_or : t -> t -> t +[@@ocaml.inline always] val no_check_addc : t -> t -> t +[@@ocaml.inline always] + val no_check_subc : t -> t -> t -val no_check_addcarryc : t -> t -> t -val no_check_subcarryc : t -> t -> t +[@@ocaml.inline always] + +val no_check_addCarryC : t -> t -> t +[@@ocaml.inline always] + +val no_check_subCarryC : t -> t -> t +[@@ocaml.inline always] val no_check_mulc : t -> t -> t +[@@ocaml.inline always] + val no_check_diveucl : t -> t -> t +[@@ocaml.inline always] val no_check_div21 : t -> t -> t -> t -val no_check_addmuldiv : t -> t -> t -> t +[@@ocaml.inline always] + +val no_check_addMulDiv : t -> t -> t -> t +[@@ocaml.inline always] val no_check_eq : t -> t -> t +[@@ocaml.inline always] + val no_check_lt : t -> t -> t +[@@ocaml.inline always] + val no_check_le : t -> t -> t -val no_check_compare : t -> t -> t +[@@ocaml.inline always] -val mk_I31_accu : t -val decomp_uint : t -> t -> t +val no_check_compare : t -> t -> t diff --git a/kernel/primred.ml b/kernel/primred.ml new file mode 100644 index 0000000000..d95d7de7aa --- /dev/null +++ b/kernel/primred.ml @@ -0,0 +1,208 @@ +(* Reduction of native operators *) +open Names +open CPrimitives +open Retroknowledge +open Environ +open CErrors + +let add_retroknowledge env action = + match action with + | Register_type(PT_int63,c) -> + let retro = env.retroknowledge in + let retro = + match retro.retro_int63 with + | None -> { retro with retro_int63 = 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 = + match pit with + | PIT_bool -> + let r = + match retro.retro_bool with + | None -> ((ind,1), (ind,2)) + | Some (((ind',_),_) as t) -> assert (eq_ind ind ind'); t in + { retro with retro_bool = Some r } + | PIT_carry -> + let r = + match retro.retro_carry with + | None -> ((ind,1), (ind,2)) + | Some (((ind',_),_) as t) -> assert (eq_ind ind ind'); t in + { retro with retro_carry = Some r } + | PIT_pair -> + let r = + match retro.retro_pair with + | None -> (ind,1) + | Some ((ind',_) as t) -> assert (eq_ind ind ind'); t in + { retro with retro_pair = Some r } + | PIT_cmp -> + let r = + match retro.retro_cmp with + | None -> ((ind,1), (ind,2), (ind,3)) + | Some (((ind',_),_,_) as t) -> assert (eq_ind ind ind'); t in + { retro with retro_cmp = Some r } + in + set_retroknowledge env retro + | Register_inline(c) -> + let (cb,r) = lookup_constant_key c env in + let cb = {cb with Declarations.const_inline_code = true} in + add_constant_key c cb !(fst r) env + +let get_int_type env = + match env.retroknowledge.retro_int63 with + | Some c -> c + | None -> anomaly Pp.(str"Reduction of primitive: int63 not registered") + +let get_bool_constructors env = + match env.retroknowledge.retro_bool with + | Some r -> r + | None -> anomaly Pp.(str"Reduction of primitive: bool not registered") + +let get_carry_constructors env = + match env.retroknowledge.retro_carry with + | Some r -> r + | None -> anomaly Pp.(str"Reduction of primitive: carry not registered") + +let get_pair_constructor env = + match env.retroknowledge.retro_pair with + | Some c -> c + | None -> anomaly Pp.(str"Reduction of primitive: pair not registered") + +let get_cmp_constructors env = + match env.retroknowledge.retro_cmp with + | Some r -> r + | None -> anomaly Pp.(str"Reduction of primitive: cmp not registered") + +exception NativeDestKO + +module type RedNativeEntries = + sig + type elem + type args + type evd (* will be unit in kernel, evar_map outside *) + + val get : args -> int -> elem + val get_int : evd -> elem -> Uint63.t + val mkInt : env -> Uint63.t -> elem + val mkBool : env -> bool -> elem + val mkCarry : env -> bool -> elem -> elem (* true if carry *) + val mkIntPair : env -> elem -> elem -> elem + val mkLt : env -> elem + val mkEq : env -> elem + val mkGt : env -> elem + + end + +module type RedNative = + sig + type elem + type args + type evd + val red_prim : env -> evd -> CPrimitives.t -> args -> elem option + end + +module RedNative (E:RedNativeEntries) : + RedNative with type elem = E.elem + with type args = E.args + with type evd = E.evd = +struct + type elem = E.elem + type args = E.args + type evd = E.evd + + let get_int evd args i = E.get_int evd (E.get args i) + + let get_int1 evd args = get_int evd args 0 + + let get_int2 evd args = get_int evd args 0, get_int evd args 1 + + let get_int3 evd args = + get_int evd args 0, get_int evd args 1, get_int evd args 2 + + let red_prim_aux env evd op args = + let open CPrimitives in + match op with + | Int63head0 -> + let i = get_int1 evd args in E.mkInt env (Uint63.head0 i) + | Int63tail0 -> + let i = get_int1 evd args in E.mkInt env (Uint63.tail0 i) + | Int63add -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.add i1 i2) + | Int63sub -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.sub i1 i2) + | Int63mul -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.mul i1 i2) + | Int63div -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.div i1 i2) + | Int63mod -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.rem i1 i2) + | Int63lsr -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_sr i1 i2) + | Int63lsl -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_sl i1 i2) + | Int63land -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_and i1 i2) + | Int63lor -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_or i1 i2) + | Int63lxor -> + let i1, i2 = get_int2 evd args in E.mkInt env (Uint63.l_xor i1 i2) + | Int63addc -> + let i1, i2 = get_int2 evd args in + let s = Uint63.add i1 i2 in + E.mkCarry env (Uint63.lt s i1) (E.mkInt env s) + | Int63subc -> + let i1, i2 = get_int2 evd args in + let s = Uint63.sub i1 i2 in + E.mkCarry env (Uint63.lt i1 i2) (E.mkInt env s) + | Int63addCarryC -> + let i1, i2 = get_int2 evd args in + let s = Uint63.add (Uint63.add i1 i2) (Uint63.of_int 1) in + E.mkCarry env (Uint63.le s i1) (E.mkInt env s) + | Int63subCarryC -> + let i1, i2 = get_int2 evd args in + let s = Uint63.sub (Uint63.sub i1 i2) (Uint63.of_int 1) in + E.mkCarry env (Uint63.le i1 i2) (E.mkInt env s) + | Int63mulc -> + let i1, i2 = get_int2 evd args in + let (h, l) = Uint63.mulc i1 i2 in + E.mkIntPair env (E.mkInt env h) (E.mkInt env l) + | Int63diveucl -> + let i1, i2 = get_int2 evd args in + let q,r = Uint63.div i1 i2, Uint63.rem i1 i2 in + E.mkIntPair env (E.mkInt env q) (E.mkInt env r) + | Int63div21 -> + let i1, i2, i3 = get_int3 evd args in + let q,r = Uint63.div21 i1 i2 i3 in + E.mkIntPair env (E.mkInt env q) (E.mkInt env r) + | Int63addMulDiv -> + let p, i, j = get_int3 evd args in + E.mkInt env + (Uint63.l_or + (Uint63.l_sl i p) + (Uint63.l_sr j (Uint63.sub (Uint63.of_int Uint63.uint_size) p))) + | Int63eq -> + let i1, i2 = get_int2 evd args in + E.mkBool env (Uint63.equal i1 i2) + | Int63lt -> + let i1, i2 = get_int2 evd args in + E.mkBool env (Uint63.lt i1 i2) + | Int63le -> + let i1, i2 = get_int2 evd args in + E.mkBool env (Uint63.le i1 i2) + | Int63compare -> + let i1, i2 = get_int2 evd args in + begin match Uint63.compare i1 i2 with + | x when x < 0 -> E.mkLt env + | 0 -> E.mkEq env + | _ -> E.mkGt env + end + + let red_prim env evd p args = + try + let r = + red_prim_aux env evd p args + in Some r + with NativeDestKO -> None + +end diff --git a/kernel/primred.mli b/kernel/primred.mli new file mode 100644 index 0000000000..f5998982d7 --- /dev/null +++ b/kernel/primred.mli @@ -0,0 +1,44 @@ +open Names +open Environ + +(** {5 Reduction of primitives} *) +val add_retroknowledge : env -> Retroknowledge.action -> env + +val get_int_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 + +exception NativeDestKO (* Should be raised by get_* functions on failure *) + +module type RedNativeEntries = + sig + type elem + type args + type evd (* will be unit in kernel, evar_map outside *) + + val get : args -> int -> elem + val get_int : evd -> elem -> Uint63.t + val mkInt : env -> Uint63.t -> elem + val mkBool : env -> bool -> elem + val mkCarry : env -> bool -> elem -> elem (* true if carry *) + val mkIntPair : env -> elem -> elem -> elem + val mkLt : env -> elem + val mkEq : env -> elem + val mkGt : env -> elem + end + +module type RedNative = + sig + type elem + type args + type evd + val red_prim : env -> evd -> CPrimitives.t -> args -> elem option + end + +module RedNative : + functor (E:RedNativeEntries) -> + RedNative with type elem = E.elem + with type args = E.args + with type evd = E.evd diff --git a/kernel/reduction.ml b/kernel/reduction.ml index 97cd4c00d7..61051c001d 100644 --- a/kernel/reduction.ml +++ b/kernel/reduction.ml @@ -21,6 +21,7 @@ open CErrors open Util open Names open Constr +open Declarations open Vars open Environ open CClosure @@ -59,16 +60,23 @@ let compare_stack_shape stk1 stk2 = Int.equal bal 0 (* && c1.ci_ind = c2.ci_ind *) && compare_rec 0 s1 s2 | (Zfix(_,a1)::s1, Zfix(_,a2)::s2) -> Int.equal bal 0 && compare_rec 0 a1 a2 && compare_rec 0 s1 s2 + | Zprimitive(op1,_,rargs1, _kargs1)::s1, Zprimitive(op2,_,rargs2, _kargs2)::s2 -> + bal=0 && op1=op2 && List.length rargs1=List.length rargs2 && + compare_rec 0 s1 s2 | [], _ :: _ - | (Zproj _ | ZcaseT _ | Zfix _) :: _, _ -> false + | (Zproj _ | ZcaseT _ | Zfix _ | Zprimitive _) :: _, _ -> false in compare_rec 0 stk1 stk2 +type lft_fconstr = lift * fconstr + type lft_constr_stack_elt = Zlapp of (lift * fconstr) array | Zlproj of Projection.Repr.t * lift | Zlfix of (lift * fconstr) * lft_constr_stack | Zlcase of case_info * lift * constr * constr array * fconstr subs + | Zlprimitive of + CPrimitives.t * pconstant * lft_fconstr list * lft_fconstr next_native_args and lft_constr_stack = lft_constr_stack_elt list let rec zlapp v = function @@ -102,7 +110,10 @@ let pure_stack lfts stk = let (lfx,pa) = pure_rec l a in (l, Zlfix((lfx,fx),pa)::pstk) | (ZcaseT(ci,p,br,e),(l,pstk)) -> - (l,Zlcase(ci,l,p,br,e)::pstk)) + (l,Zlcase(ci,l,p,br,e)::pstk) + | (Zprimitive(op,c,rargs,kargs),(l,pstk)) -> + (l,Zlprimitive(op,c,List.map (fun t -> (l,t)) rargs, + List.map (fun (k,t) -> (k,(l,t))) kargs)::pstk)) in snd (pure_rec lfts stk) @@ -127,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 _) -> x + Prod _|Lambda _|Fix _|CoFix _|Int _) -> x | App (c, _) -> begin match kind c with - | Ind _ | Construct _ | Evar _ | Meta _ | Const _ -> x + | Ind _ | Construct _ | Evar _ | Meta _ | Const _ | Int _ -> x | Sort _ | Rel _ | Var _ | Cast _ | Prod _ | Lambda _ | LetIn _ | App _ | Case _ | Fix _ | CoFix _ | Proj _ -> whd_val (create_clos_infos betaiotazeta env) (create_tab ()) (inject x) @@ -141,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 _) -> t + Prod _|Lambda _|Fix _|CoFix _|Int _) -> t | App (c, _) -> begin match kind c with - | Ind _ | Construct _ | Evar _ | Meta _ -> t + | Ind _ | Construct _ | Evar _ | Meta _ | Int _ -> 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) @@ -155,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 _) -> t + Prod _|Lambda _|Fix _|CoFix _|LetIn _|Int _) -> t | App (c, _) -> begin match kind c with - | Ind _ | Construct _ | Evar _ | Meta _ | LetIn _ -> t + | Ind _ | Construct _ | Evar _ | Meta _ | LetIn _ | Int _ -> t | Sort _ | Rel _ | Var _ | Cast _ | Prod _ | Lambda _ | App _ | Const _ | Case _ | Fix _ | CoFix _ | Proj _ -> whd_val (create_clos_infos allnolet env) (create_tab ()) (inject t) @@ -296,6 +307,15 @@ let rec skip_pattern n c = | Lambda (_, _, c) -> skip_pattern (pred n) c | _ -> raise IrregularPatternShape +let unfold_ref_with_args infos tab fl v = + match unfold_reference infos tab fl with + | Def def -> Some (def, v) + | Primitive op when check_native_args op v -> + let c = match fl with ConstKey c -> c | _ -> assert false in + let rargs, a, nargs, v = get_native_args1 op c v in + Some (whd_stack infos tab a (Zupdate a::(Zprimitive(op,c,rargs,nargs)::v))) + | Undef _ | OpaqueDef _ | Primitive _ -> None + type conv_tab = { cnv_inf : clos_infos; lft_tab : clos_tab; @@ -355,28 +375,26 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = (* 2 constants, 2 local defined vars or 2 defined rels *) | (FFlex fl1, FFlex fl2) -> (try - let cuniv = conv_table_key infos.cnv_inf fl1 fl2 cuniv in - convert_stacks l2r infos lft1 lft2 v1 v2 cuniv + let cuniv = conv_table_key infos.cnv_inf fl1 fl2 cuniv in + convert_stacks l2r infos lft1 lft2 v1 v2 cuniv with NotConvertible | Univ.UniverseInconsistency _ -> - (* else the oracle tells which constant is to be expanded *) + (* else the oracle tells which constant is to be expanded *) let oracle = CClosure.oracle_of_infos infos.cnv_inf in let (app1,app2) = + let aux appr1 lft1 fl1 tab1 v1 appr2 lft2 fl2 tab2 v2 = + match unfold_ref_with_args infos.cnv_inf tab1 fl1 v1 with + | Some t1 -> ((lft1, t1), appr2) + | None -> match unfold_ref_with_args infos.cnv_inf tab2 fl2 v2 with + | Some t2 -> (appr1, (lft2, t2)) + | None -> raise NotConvertible + in if Conv_oracle.oracle_order Univ.out_punivs oracle l2r fl1 fl2 then - match unfold_reference infos.cnv_inf infos.lft_tab fl1 with - | Some def1 -> ((lft1, (def1, v1)), appr2) - | None -> - (match unfold_reference infos.cnv_inf infos.rgt_tab fl2 with - | Some def2 -> (appr1, (lft2, (def2, v2))) - | None -> raise NotConvertible) + aux appr1 lft1 fl1 infos.lft_tab v1 appr2 lft2 fl2 infos.rgt_tab v2 else - match unfold_reference infos.cnv_inf infos.rgt_tab fl2 with - | Some def2 -> (appr1, (lft2, (def2, v2))) - | None -> - (match unfold_reference infos.cnv_inf infos.lft_tab fl1 with - | Some def1 -> ((lft1, (def1, v1)), appr2) - | None -> raise NotConvertible) - in - eqappr cv_pb l2r infos app1 app2 cuniv) + let (app2,app1) = aux appr2 lft2 fl2 infos.rgt_tab v2 appr1 lft1 fl1 infos.lft_tab v1 in + (app1,app2) + in + eqappr cv_pb l2r infos app1 app2 cuniv) | (FProj (p1,c1), FProj (p2, c2)) -> (* Projections: prefer unfolding to first-order unification, @@ -400,31 +418,37 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = raise NotConvertible) | (FProj (p1,c1), t2) -> - (match unfold_projection infos.cnv_inf p1 with - | Some s1 -> + begin match unfold_projection infos.cnv_inf p1 with + | Some s1 -> eqappr cv_pb l2r infos (lft1, (c1, (s1 :: v1))) appr2 cuniv - | None -> - (match t2 with - | FFlex fl2 -> - (match unfold_reference infos.cnv_inf infos.rgt_tab fl2 with - | Some def2 -> - eqappr cv_pb l2r infos appr1 (lft2, (def2, v2)) cuniv - | None -> raise NotConvertible) - | _ -> raise NotConvertible)) - + | None -> + begin match t2 with + | FFlex fl2 -> + begin match unfold_ref_with_args infos.cnv_inf infos.rgt_tab fl2 v2 with + | Some t2 -> + eqappr cv_pb l2r infos appr1 (lft2, t2) cuniv + | None -> raise NotConvertible + end + | _ -> raise NotConvertible + end + end + | (t1, FProj (p2,c2)) -> - (match unfold_projection infos.cnv_inf p2 with - | Some s2 -> + begin match unfold_projection infos.cnv_inf p2 with + | Some s2 -> eqappr cv_pb l2r infos appr1 (lft2, (c2, (s2 :: v2))) cuniv - | None -> - (match t1 with - | FFlex fl1 -> - (match unfold_reference infos.cnv_inf infos.lft_tab fl1 with - | Some def1 -> - eqappr cv_pb l2r infos (lft1, (def1, v1)) appr2 cuniv - | None -> raise NotConvertible) - | _ -> raise NotConvertible)) - + | None -> + begin match t1 with + | FFlex fl1 -> + begin match unfold_ref_with_args infos.cnv_inf infos.lft_tab fl1 v1 with + | Some t1 -> + eqappr cv_pb l2r infos (lft1, t1) appr2 cuniv + | None -> raise NotConvertible + end + | _ -> raise NotConvertible + end + end + (* other constructors *) | (FLambda _, FLambda _) -> (* Inconsistency: we tolerate that v1, v2 contain shift and update but @@ -469,8 +493,8 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = (* only one constant, defined var or defined rel *) | (FFlex fl1, c2) -> - (match unfold_reference infos.cnv_inf infos.lft_tab fl1 with - | Some def1 -> + begin match unfold_ref_with_args infos.cnv_inf infos.lft_tab fl1 v1 with + | Some (def1,v1) -> (** By virtue of the previous case analyses, we know [c2] is rigid. Conversion check to rigid terms eventually implies full weak-head reduction, so instead of repeatedly performing small-step @@ -478,32 +502,34 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = let all = RedFlags.red_add_transparent all (RedFlags.red_transparent (info_flags infos.cnv_inf)) in let r1 = whd_stack (infos_with_reds infos.cnv_inf all) infos.lft_tab def1 v1 in eqappr cv_pb l2r infos (lft1, r1) appr2 cuniv - | None -> - match c2 with + | None -> + (match c2 with | FConstruct ((ind2,_j2),_u2) -> - (try - let v2, v1 = - eta_expand_ind_stack (info_env infos.cnv_inf) ind2 hd2 v2 (snd appr1) - in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv - with Not_found -> raise NotConvertible) - | _ -> raise NotConvertible) - + (try + let v2, v1 = + eta_expand_ind_stack (info_env infos.cnv_inf) ind2 hd2 v2 (snd appr1) + in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv + with Not_found -> raise NotConvertible) + | _ -> raise NotConvertible) + end + | (c1, FFlex fl2) -> - (match unfold_reference infos.cnv_inf infos.rgt_tab fl2 with - | Some def2 -> + begin match unfold_ref_with_args infos.cnv_inf infos.rgt_tab fl2 v2 with + | Some (def2, v2) -> (** Symmetrical case of above. *) let all = RedFlags.red_add_transparent all (RedFlags.red_transparent (info_flags infos.cnv_inf)) in let r2 = whd_stack (infos_with_reds infos.cnv_inf all) infos.rgt_tab def2 v2 in eqappr cv_pb l2r infos appr1 (lft2, r2) cuniv - | None -> - match c1 with - | FConstruct ((ind1,_j1),_u1) -> - (try let v1, v2 = - eta_expand_ind_stack (info_env infos.cnv_inf) ind1 hd1 v1 (snd appr2) - in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv - with Not_found -> raise NotConvertible) - | _ -> raise NotConvertible) - + | None -> + match c1 with + | FConstruct ((ind1,_j1),_u1) -> + (try let v1, v2 = + eta_expand_ind_stack (info_env infos.cnv_inf) ind1 hd1 v1 (snd appr2) + in convert_stacks l2r infos lft1 lft2 v1 v2 cuniv + with Not_found -> raise NotConvertible) + | _ -> raise NotConvertible + end + (* Inductive types: MutInd MutConstruct Fix Cofix *) | (FInd (ind1,u1), FInd (ind2,u2)) -> if eq_ind ind1 ind2 then @@ -584,13 +610,17 @@ and eqappr cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv = convert_stacks l2r infos lft1 lft2 v1 v2 cuniv else raise NotConvertible + | FInt i1, FInt i2 -> + if Uint63.equal i1 i2 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 _), _ -> raise NotConvertible + | FProd _ | FEvar _ | FInt _), _ -> 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 @@ -613,7 +643,12 @@ and convert_stacks l2r infos lft1 lft2 stk1 stk2 cuniv = raise NotConvertible; let cu2 = f (l1, mk_clos e1 p1) (l2, mk_clos e2 p2) cu1 in convert_branches l2r infos ci1 e1 e2 l1 l2 br1 br2 cu2 - | _ -> assert false) + | (Zlprimitive(op1,_,rargs1,kargs1),Zlprimitive(op2,_,rargs2,kargs2)) -> + if not (CPrimitives.equal op1 op2) then raise NotConvertible else + let cu2 = List.fold_right2 f rargs1 rargs2 cu1 in + let fk (_,a1) (_,a2) cu = f a1 a2 cu in + List.fold_right2 fk kargs1 kargs2 cu2 + | ((Zlapp _ | Zlproj _ | Zlfix _| Zlcase _| Zlprimitive _), _) -> assert false) | _ -> cuniv in if compare_stack_shape stk1 stk2 then cmp_rec (pure_stack lft1 stk1) (pure_stack lft2 stk2) cuniv diff --git a/kernel/retroknowledge.ml b/kernel/retroknowledge.ml index e51c25c06b..18fafdb6d3 100644 --- a/kernel/retroknowledge.ml +++ b/kernel/retroknowledge.ml @@ -12,249 +12,30 @@ (* Addition of native Head (nb of heading 0) and Tail (nb of trailing 0) by Benjamin Grégoire, Jun 2007 *) -(* This file defines the knowledge that the kernel is able to optimize - for evaluation in the bytecode virtual machine *) +(* This file defines the knowledge that the kernel is able to optimize. *) open Names -open Constr - -(* The retroknowledge defines a bijective correspondance between some - [entry]-s (which are, in fact, merely names) and [field]-s which - are roles assigned to these entries. *) - -type int31_field = - | Int31Bits - | Int31Type - | Int31Constructor - | Int31Twice - | Int31TwicePlusOne - | Int31Phi - | Int31PhiInv - | Int31Plus - | Int31PlusC - | Int31PlusCarryC - | Int31Minus - | Int31MinusC - | Int31MinusCarryC - | Int31Times - | Int31TimesC - | Int31Div21 - | Int31Div - | Int31Diveucl - | Int31AddMulDiv - | Int31Compare - | Int31Head0 - | Int31Tail0 - | Int31Lor - | Int31Land - | Int31Lxor - -type field = - | KInt31 of int31_field - -let int31_field_of_string = - function - | "bits" -> Int31Bits - | "type" -> Int31Type - | "twice" -> Int31Twice - | "twice_plus_one" -> Int31TwicePlusOne - | "phi" -> Int31Phi - | "phi_inv" -> Int31PhiInv - | "plus" -> Int31Plus - | "plusc" -> Int31PlusC - | "pluscarryc" -> Int31PlusCarryC - | "minus" -> Int31Minus - | "minusc" -> Int31MinusC - | "minuscarryc" -> Int31MinusCarryC - | "times" -> Int31Times - | "timesc" -> Int31TimesC - | "div21" -> Int31Div21 - | "div" -> Int31Div - | "diveucl" -> Int31Diveucl - | "addmuldiv" -> Int31AddMulDiv - | "compare" -> Int31Compare - | "head0" -> Int31Head0 - | "tail0" -> Int31Tail0 - | "lor" -> Int31Lor - | "land" -> Int31Land - | "lxor" -> Int31Lxor - | s -> CErrors.user_err Pp.(str "Registering unknown int31 operator " ++ str s) - -let int31_path = DirPath.make [ Id.of_string "int31" ] - -(* record representing all the flags of the internal state of the kernel *) -type flags = {fastcomputation : bool} - - - - - -(* The [proactive] knowledge contains the mapping [field->entry]. *) - -module Proactive = - Map.Make (struct type t = field let compare = Pervasives.compare end) - -type proactive = GlobRef.t Proactive.t - -(* The [reactive] knowledge contains the mapping - [entry->field]. Fields are later to be interpreted as a - [reactive_info]. *) - -module Reactive = GlobRef.Map - -type reactive_info = {(*information required by the compiler of the VM *) - vm_compiling : - (*fastcomputation flag -> continuation -> result *) - (bool -> Cinstr.lambda array -> Cinstr.lambda) - option; - vm_constant_static : - (*fastcomputation flag -> constructor -> args -> result*) - (bool -> constr array -> Cinstr.lambda) - option; - vm_constant_dynamic : - (*fastcomputation flag -> constructor -> reloc -> args -> sz -> cont -> result *) - (bool -> Cinstr.lambda array -> Cinstr.lambda) - option; - (* fastcomputation flag -> cont -> result *) - vm_before_match : (bool -> Cinstr.lambda -> Cinstr.lambda) option; - (* tag (= compiled int for instance) -> result *) - vm_decompile_const : (int -> constr) option; - - native_compiling : - (bool -> Nativeinstr.prefix -> Nativeinstr.lambda array -> - Nativeinstr.lambda) option; - - native_constant_static : - (bool -> constr array -> Nativeinstr.lambda) option; - - native_constant_dynamic : - (bool -> Nativeinstr.prefix -> constructor -> - Nativeinstr.lambda array -> Nativeinstr.lambda) option; - - native_before_match : (bool -> Nativeinstr.prefix -> constructor -> - Nativeinstr.lambda -> Nativeinstr.lambda) option +type retroknowledge = { + retro_int63 : 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; } +let empty = { + retro_int63 = None; + retro_bool = None; + retro_carry = None; + retro_pair = None; + retro_cmp = None; + retro_refl = None; +} - -and reactive = field Reactive.t - -and retroknowledge = {flags : flags; proactive : proactive; reactive : reactive} - -(* This type represent an atomic action of the retroknowledge. It - is stored in the compiled libraries *) -(* As per now, there is only the possibility of registering things - the possibility of unregistering or changing the flag is under study *) type action = - | RKRegister of field * GlobRef.t - - -(*initialisation*) -let initial_flags = - {fastcomputation = true;} - -let initial_proactive = - (Proactive.empty:proactive) - -let initial_reactive = - (Reactive.empty:reactive) - -let initial_retroknowledge = - {flags = initial_flags; - proactive = initial_proactive; - reactive = initial_reactive } - -let empty_reactive_info = - { vm_compiling = None ; - vm_constant_static = None; - vm_constant_dynamic = None; - vm_before_match = None; - vm_decompile_const = None; - native_compiling = None; - native_constant_static = None; - native_constant_dynamic = None; - native_before_match = None; - } - - - -(* adds a binding [entry<->field]. *) -let add_field knowledge field entry = - {knowledge with - proactive = Proactive.add field entry knowledge.proactive; - reactive = Reactive.add entry field knowledge.reactive} - -(* acces functions for proactive retroknowledge *) -let mem knowledge field = - Proactive.mem field knowledge.proactive - -let find knowledge field = - Proactive.find field knowledge.proactive - - -let (dispatch,dispatch_hook) = Hook.make () - -let dispatch_reactive entry retroknowledge = - Hook.get dispatch retroknowledge entry (Reactive.find entry retroknowledge.reactive) - -(*access functions for reactive retroknowledge*) - -(* used for compiling of functions (add, mult, etc..) *) -let get_vm_compiling_info knowledge key = - match (dispatch_reactive key knowledge).vm_compiling - with - | None -> raise Not_found - | Some f -> f knowledge.flags.fastcomputation - -(* used for compilation of fully applied constructors *) -let get_vm_constant_static_info knowledge key = - match (dispatch_reactive key knowledge).vm_constant_static - with - | None -> raise Not_found - | Some f -> f knowledge.flags.fastcomputation - -(* used for compilation of partially applied constructors *) -let get_vm_constant_dynamic_info knowledge key = - match (dispatch_reactive key knowledge).vm_constant_dynamic - with - | None -> raise Not_found - | Some f -> f knowledge.flags.fastcomputation - -let get_vm_before_match_info knowledge key = - match (dispatch_reactive key knowledge).vm_before_match - with - | None -> raise Not_found - | Some f -> f knowledge.flags.fastcomputation - -let get_vm_decompile_constant_info knowledge key = - match (dispatch_reactive key knowledge).vm_decompile_const - with - | None -> raise Not_found - | Some f -> f - -let get_native_compiling_info knowledge key = - match (dispatch_reactive key knowledge).native_compiling - with - | None -> raise Not_found - | Some f -> f knowledge.flags.fastcomputation - -(* used for compilation of fully applied constructors *) -let get_native_constant_static_info knowledge key = - match (dispatch_reactive key knowledge).native_constant_static - with - | None -> raise Not_found - | Some f -> f knowledge.flags.fastcomputation - -(* used for compilation of partially applied constructors *) -let get_native_constant_dynamic_info knowledge key = - match (dispatch_reactive key knowledge).native_constant_dynamic - with - | None -> raise Not_found - | Some f -> f knowledge.flags.fastcomputation - -let get_native_before_match_info knowledge key = - match (dispatch_reactive key knowledge).native_before_match - with - | None -> raise Not_found - | Some f -> f knowledge.flags.fastcomputation + | Register_ind of CPrimitives.prim_ind * inductive + | Register_type of CPrimitives.prim_type * Constant.t + | Register_inline of Constant.t diff --git a/kernel/retroknowledge.mli b/kernel/retroknowledge.mli index 0a2ef5300e..1554fe88da 100644 --- a/kernel/retroknowledge.mli +++ b/kernel/retroknowledge.mli @@ -9,157 +9,20 @@ (************************************************************************) open Names -open Constr - -type retroknowledge - -(** the following types correspond to the different "things" - the kernel can learn about.*) -type int31_field = - | Int31Bits - | Int31Type - | Int31Constructor - | Int31Twice - | Int31TwicePlusOne - | Int31Phi - | Int31PhiInv - | Int31Plus - | Int31PlusC - | Int31PlusCarryC - | Int31Minus - | Int31MinusC - | Int31MinusCarryC - | Int31Times - | Int31TimesC - | Int31Div21 - | Int31Div - | Int31Diveucl - | Int31AddMulDiv - | Int31Compare - | Int31Head0 - | Int31Tail0 - | Int31Lor - | Int31Land - | Int31Lxor - -type field = - | KInt31 of int31_field - -val int31_field_of_string : string -> int31_field - -val int31_path : DirPath.t - -(** This type represent an atomic action of the retroknowledge. It - is stored in the compiled libraries - As per now, there is only the possibility of registering things - the possibility of unregistering or changing the flag is under study *) -type action = - | RKRegister of field * GlobRef.t - - -(** initial value for retroknowledge *) -val initial_retroknowledge : retroknowledge - - -(** Given an identifier id (usually Const _) - and the continuation cont of the bytecode compilation - returns the compilation of id in cont if it has a specific treatment - or raises Not_found if id should be compiled as usual *) -val get_vm_compiling_info : retroknowledge -> GlobRef.t -> - Cinstr.lambda array -> Cinstr.lambda -(*Given an identifier id (usually Construct _) - and its argument array, returns a function that tries an ad-hoc optimisated - compilation (in the case of the 31-bit integers it means compiling them - directly into an integer) - raises Not_found if id should be compiled as usual, and expectingly - CBytecodes.NotClosed if the term is not a closed constructor pattern - (a constant for the compiler) *) -val get_vm_constant_static_info : retroknowledge -> GlobRef.t -> - constr array -> Cinstr.lambda - -(*Given an identifier id (usually Construct _ ) - its argument array and a continuation, returns the compiled version - of id+args+cont when id has a specific treatment (in the case of - 31-bit integers, that would be the dynamic compilation into integers) - or raises Not_found if id should be compiled as usual *) -val get_vm_constant_dynamic_info : retroknowledge -> GlobRef.t -> - Cinstr.lambda array -> Cinstr.lambda - -(** Given a type identifier, this function is used before compiling a match - over this type. In the case of 31-bit integers for instance, it is used - to add the instruction sequence which would perform a dynamic decompilation - in case the argument of the match is not in coq representation *) -val get_vm_before_match_info : retroknowledge -> GlobRef.t -> Cinstr.lambda - -> Cinstr.lambda - -(** Given a type identifier, this function is used by pretyping/vnorm.ml to - recover the elements of that type from their compiled form if it's non - standard (it is used (and can be used) only when the compiled form - is not a block *) -val get_vm_decompile_constant_info : retroknowledge -> GlobRef.t -> int -> constr - - -val get_native_compiling_info : retroknowledge -> GlobRef.t -> Nativeinstr.prefix -> - Nativeinstr.lambda array -> Nativeinstr.lambda - -val get_native_constant_static_info : retroknowledge -> GlobRef.t -> - constr array -> Nativeinstr.lambda - -val get_native_constant_dynamic_info : retroknowledge -> GlobRef.t -> - Nativeinstr.prefix -> constructor -> - Nativeinstr.lambda array -> - Nativeinstr.lambda - -val get_native_before_match_info : retroknowledge -> GlobRef.t -> - Nativeinstr.prefix -> constructor -> - Nativeinstr.lambda -> Nativeinstr.lambda - - -(** the following functions are solely used in Environ and Safe_typing to implement - the functions register and unregister (and mem) of Environ *) -val add_field : retroknowledge -> field -> GlobRef.t -> retroknowledge -val mem : retroknowledge -> field -> bool -(* val remove : retroknowledge -> field -> retroknowledge *) -val find : retroknowledge -> field -> GlobRef.t - - -(** Dispatching type for the above [get_*] functions. *) -type reactive_info = {(*information required by the compiler of the VM *) - vm_compiling : - (*fastcomputation flag -> continuation -> result *) - (bool -> Cinstr.lambda array -> Cinstr.lambda) - option; - vm_constant_static : - (*fastcomputation flag -> constructor -> args -> result*) - (bool -> constr array -> Cinstr.lambda) - option; - vm_constant_dynamic : - (*fastcomputation flag -> constructor -> reloc -> args -> sz -> cont -> result *) - (bool -> Cinstr.lambda array -> Cinstr.lambda) - option; - (* fastcomputation flag -> cont -> result *) - vm_before_match : (bool -> Cinstr.lambda -> Cinstr.lambda) option; - (* tag (= compiled int for instance) -> result *) - vm_decompile_const : (int -> constr) option; - - native_compiling : - (bool -> Nativeinstr.prefix -> Nativeinstr.lambda array -> - Nativeinstr.lambda) option; - - native_constant_static : - (bool -> constr array -> Nativeinstr.lambda) option; - - native_constant_dynamic : - (bool -> Nativeinstr.prefix -> constructor -> - Nativeinstr.lambda array -> Nativeinstr.lambda) option; - - native_before_match : (bool -> Nativeinstr.prefix -> constructor -> - Nativeinstr.lambda -> Nativeinstr.lambda) option +type retroknowledge = { + retro_int63 : 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; } -val empty_reactive_info : reactive_info +val empty : retroknowledge -(** Hook to be set after the compiler are installed to dispatch fields - into the above [get_*] functions. *) -val dispatch_hook : (retroknowledge -> GlobRef.t -> field -> reactive_info) Hook.t +type action = + | Register_ind of CPrimitives.prim_ind * inductive + | Register_type of CPrimitives.prim_type * Constant.t + | Register_inline of Constant.t diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml index df9e253135..474ce3e871 100644 --- a/kernel/safe_typing.ml +++ b/kernel/safe_typing.ml @@ -59,10 +59,10 @@ etc. *) -open CErrors open Util open Names open Declarations +open Constr open Context.Named.Declaration module NamedDecl = Context.Named.Declaration @@ -227,6 +227,7 @@ let check_engagement env expected_impredicative_set = let get_opaque_body env cbo = match cbo.const_body with | Undef _ -> assert false + | Primitive _ -> assert false | Def _ -> `Nothing | OpaqueDef opaque -> `Opaque @@ -467,7 +468,7 @@ let globalize_constant_universes env cb = | Monomorphic_const cstrs -> Now (false, cstrs) :: (match cb.const_body with - | (Undef _ | Def _) -> [] + | (Undef _ | Def _ | Primitive _) -> [] | OpaqueDef lc -> match Opaqueproof.get_constraints (Environ.opaque_tables env) lc with | None -> [] @@ -492,6 +493,11 @@ let constraints_of_sfb env sfb = | SFBmodtype mtb -> [Now (false, mtb.mod_constraints)] | SFBmodule mb -> [Now (false, mb.mod_constraints)] +let add_retroknowledge pttc senv = + { senv with + env = Primred.add_retroknowledge senv.env pttc; + local_retroknowledge = pttc::senv.local_retroknowledge } + (** A generic function for adding a new field in a same environment. It also performs the corresponding [add_constraints]. *) @@ -809,6 +815,13 @@ let add_constant ~in_section l decl senv = let cb = Term_typing.translate_recipe ~hcons:(not in_section) senv.env kn r in if in_section then cb else Declareops.hcons_const_body cb in add_constant_aux ~in_section senv (kn, cb) in + let senv = + match decl with + | ConstantEntry (_,(Entries.PrimitiveEntry { Entries.prim_entry_content = CPrimitives.OT_type t; _ })) -> + if in_section then CErrors.anomaly (Pp.str "Primitive type not allowed in sections"); + add_retroknowledge (Retroknowledge.Register_type(t,kn)) senv + | _ -> senv + in kn, senv (** Insertion of inductive types *) @@ -1173,18 +1186,6 @@ let typing senv = Typeops.infer (env_of_senv senv) (** {6 Retroknowledge / native compiler } *) -let register field value senv = - (* todo : value closed *) - (* spiwack : updates the safe_env with the information that the register - action has to be performed (again) when the environment is imported *) - { senv with - env = Environ.register senv.env field value; - local_retroknowledge = - Retroknowledge.RKRegister (field,value)::senv.local_retroknowledge - } - -(* This function serves only for inlining constants in native compiler for now, -but it is meant to become a replacement for environ.register *) let register_inline kn senv = let open Environ in if not (evaluable_constant kn senv.env) then @@ -1194,6 +1195,80 @@ 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 (mind,i) r env = + let mb = Environ.lookup_mind mind env in + let check_if b s = + if not b then + CErrors.user_err ~hdr:"check_register_ind" (Pp.str s) in + check_if (Int.equal (Array.length mb.mind_packets) 1) "A non mutual inductive is expected"; + let ob = mb.mind_packets.(i) in + let check_nconstr n = + check_if (Int.equal (Array.length ob.mind_consnames) n) + ("an inductive type with "^(string_of_int n)^" constructors is expected") + in + let check_name pos s = + check_if (Id.equal ob.mind_consnames.(pos) (Id.of_string s)) + ("the "^(string_of_int (pos + 1))^ + "th constructor does not have the expected name: " ^ s) in + let check_type pos t = + check_if (Constr.equal t ob.mind_user_lc.(pos)) + ("the "^(string_of_int (pos + 1))^ + "th constructor does not have the expected type") in + let check_type_cte pos = check_type pos (Constr.mkRel 1) in + match r with + | CPrimitives.PIT_bool -> + check_nconstr 2; + check_name 0 "true"; + check_type_cte 0; + check_name 1 "false"; + check_type_cte 1 + | CPrimitives.PIT_carry -> + check_nconstr 2; + let test_type pos = + let c = ob.mind_user_lc.(pos) in + let s = "the "^(string_of_int (pos + 1))^ + "th constructor does not have the expected type" in + check_if (Constr.isProd c) s; + let (_,d,cd) = Constr.destProd c in + check_if (Constr.is_Type d) s; + check_if + (Constr.equal + (mkProd (Anonymous,mkRel 1, mkApp (mkRel 3,[|mkRel 2|]))) + cd) + s in + check_name 0 "C0"; + test_type 0; + check_name 1 "C1"; + test_type 1; + | CPrimitives.PIT_pair -> + check_nconstr 1; + check_name 0 "pair"; + let c = ob.mind_user_lc.(0) in + let s = "the "^(string_of_int 1)^ + "th constructor does not have the expected type" in + begin match Term.decompose_prod c with + | ([_,b;_,a;_,_B;_,_A], codom) -> + check_if (is_Type _A) s; + check_if (is_Type _B) s; + check_if (Constr.equal a (mkRel 2)) s; + check_if (Constr.equal b (mkRel 2)) s; + check_if (Constr.equal codom (mkApp (mkRel 5,[|mkRel 4; mkRel 3|]))) s + | _ -> check_if false s + end + | CPrimitives.PIT_cmp -> + check_nconstr 3; + check_name 0 "Eq"; + check_type_cte 0; + check_name 1 "Lt"; + check_type_cte 1; + check_name 2 "Gt"; + check_type_cte 2 + +let register_inductive ind prim senv = + check_register_ind ind prim senv.env; + let action = Retroknowledge.Register_ind(prim,ind) in + add_retroknowledge action senv + let add_constraints c = add_constraints (Now (false, Univ.ContextSet.add_constraints c Univ.ContextSet.empty)) @@ -1222,125 +1297,3 @@ Would this be correct with respect to undo's and stuff ? let set_strategy k l e = { e with env = (Environ.set_oracle e.env (Conv_oracle.set_strategy (Environ.oracle e.env) k l)) } - -(** Register retroknowledge hooks *) - -open Retroknowledge - -(* the Environ.register function synchronizes the proactive and reactive - retroknowledge. *) -let dispatch = - - (* subfunction used for static decompilation of int31 (after a vm_compute, - see pretyping/vnorm.ml for more information) *) - let constr_of_int31 = - let nth_digit_plus_one i n = (* calculates the nth (starting with 0) - digit of i and adds 1 to it - (nth_digit_plus_one 1 3 = 2) *) - if Int.equal (i land (1 lsl n)) 0 then - 1 - else - 2 - in - fun ind -> fun digit_ind -> fun tag -> - let array_of_int i = - Array.init 31 (fun n -> Constr.mkConstruct - (digit_ind, nth_digit_plus_one i (30-n))) - in - (* We check that no bit above 31 is set to one. This assertion used to - fail in the VM, and led to conversion tests failing at Qed. *) - assert (Int.equal (tag lsr 31) 0); - Constr.mkApp(Constr.mkConstruct(ind, 1), array_of_int tag) - in - - (* subfunction which dispatches the compiling information of an - int31 operation which has a specific vm instruction (associates - it to the name of the coq definition in the reactive retroknowledge) *) - let int31_op n op prim kn = - { empty_reactive_info with - vm_compiling = Some (Clambda.compile_prim n op (kn, Univ.Instance.empty)); (*XXX: FIXME universes? *) - native_compiling = Some (Nativelambda.compile_prim prim kn); - } - in - -fun rk value field -> - (* subfunction which shortens the (very common) dispatch of operations *) - let int31_op_from_const n op prim = - match value with - | GlobRef.ConstRef kn -> int31_op n op prim kn - | _ -> anomaly ~label:"Environ.register" (Pp.str "should be a constant.") - in - let int31_binop_from_const op prim = int31_op_from_const 2 op prim in - let int31_unop_from_const op prim = int31_op_from_const 1 op prim in - match field with - | KInt31 Int31Type -> - let int31bit = - (* invariant : the type of bits is registered, otherwise the function - would raise Not_found. The invariant is enforced in safe_typing.ml *) - match field with - | KInt31 Int31Type -> Retroknowledge.find rk (KInt31 Int31Bits) - | _ -> anomaly ~label:"Environ.register" - (Pp.str "add_int31_decompilation_from_type called with an abnormal field.") - in - let i31bit_type = - match int31bit with - | GlobRef.IndRef i31bit_type -> i31bit_type - | _ -> anomaly ~label:"Environ.register" - (Pp.str "Int31Bits should be an inductive type.") - in - let int31_decompilation = - match value with - | GlobRef.IndRef i31t -> - constr_of_int31 i31t i31bit_type - | _ -> anomaly ~label:"Environ.register" - (Pp.str "should be an inductive type.") - in - { empty_reactive_info with - vm_decompile_const = Some int31_decompilation; - vm_before_match = Some Clambda.int31_escape_before_match; - native_before_match = Some (Nativelambda.before_match_int31 i31bit_type); - } - | KInt31 Int31Constructor -> - { empty_reactive_info with - vm_constant_static = Some Clambda.compile_structured_int31; - vm_constant_dynamic = Some Clambda.dynamic_int31_compilation; - native_constant_static = Some Nativelambda.compile_static_int31; - native_constant_dynamic = Some Nativelambda.compile_dynamic_int31; - } - | KInt31 Int31Plus -> int31_binop_from_const Cbytecodes.Kaddint31 - CPrimitives.Int31add - | KInt31 Int31PlusC -> int31_binop_from_const Cbytecodes.Kaddcint31 - CPrimitives.Int31addc - | KInt31 Int31PlusCarryC -> int31_binop_from_const Cbytecodes.Kaddcarrycint31 - CPrimitives.Int31addcarryc - | KInt31 Int31Minus -> int31_binop_from_const Cbytecodes.Ksubint31 - CPrimitives.Int31sub - | KInt31 Int31MinusC -> int31_binop_from_const Cbytecodes.Ksubcint31 - CPrimitives.Int31subc - | KInt31 Int31MinusCarryC -> int31_binop_from_const - Cbytecodes.Ksubcarrycint31 CPrimitives.Int31subcarryc - | KInt31 Int31Times -> int31_binop_from_const Cbytecodes.Kmulint31 - CPrimitives.Int31mul - | KInt31 Int31TimesC -> int31_binop_from_const Cbytecodes.Kmulcint31 - CPrimitives.Int31mulc - | KInt31 Int31Div21 -> int31_op_from_const 3 Cbytecodes.Kdiv21int31 - CPrimitives.Int31div21 - | KInt31 Int31Diveucl -> int31_binop_from_const Cbytecodes.Kdivint31 - CPrimitives.Int31diveucl - | KInt31 Int31AddMulDiv -> int31_op_from_const 3 Cbytecodes.Kaddmuldivint31 - CPrimitives.Int31addmuldiv - | KInt31 Int31Compare -> int31_binop_from_const Cbytecodes.Kcompareint31 - CPrimitives.Int31compare - | KInt31 Int31Head0 -> int31_unop_from_const Cbytecodes.Khead0int31 - CPrimitives.Int31head0 - | KInt31 Int31Tail0 -> int31_unop_from_const Cbytecodes.Ktail0int31 - CPrimitives.Int31tail0 - | KInt31 Int31Lor -> int31_binop_from_const Cbytecodes.Klorint31 - CPrimitives.Int31lor - | KInt31 Int31Land -> int31_binop_from_const Cbytecodes.Klandint31 - CPrimitives.Int31land - | KInt31 Int31Lxor -> int31_binop_from_const Cbytecodes.Klxorint31 - CPrimitives.Int31lxor - | _ -> empty_reactive_info - -let _ = Hook.set Retroknowledge.dispatch_hook dispatch diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli index 57b01f15e3..1b7239da23 100644 --- a/kernel/safe_typing.mli +++ b/kernel/safe_typing.mli @@ -217,12 +217,8 @@ val mind_of_delta_kn_senv : safe_environment -> KerName.t -> MutInd.t (** {6 Retroknowledge / Native compiler } *) -open Retroknowledge - -val register : - field -> GlobRef.t -> safe_transformer0 - val register_inline : Constant.t -> safe_transformer0 +val register_inductive : inductive -> CPrimitives.prim_ind -> safe_transformer0 val set_strategy : Names.Constant.t Names.tableKey -> Conv_oracle.level -> safe_transformer0 diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml index 347c30dd64..2fc3aa99b5 100644 --- a/kernel/subtyping.ml +++ b/kernel/subtyping.ml @@ -257,10 +257,10 @@ let check_constant cst env l info1 cb2 spec2 subst1 subst2 = anything of the right type can implement it, even if bodies differ. *) (match cb2.const_body with - | Undef _ | OpaqueDef _ -> cst + | Primitive _ | Undef _ | OpaqueDef _ -> cst | Def lc2 -> (match cb1.const_body with - | Undef _ | OpaqueDef _ -> error NotConvertibleBodyField + | Primitive _ | Undef _ | OpaqueDef _ -> error NotConvertibleBodyField | Def lc1 -> (* NB: cb1 might have been strengthened and appear as transparent. Anyway [check_conv] will handle that afterwards. *) diff --git a/kernel/term.ml b/kernel/term.ml index 795cdeb040..58b289eaa5 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -379,4 +379,4 @@ let kind_of_type t = match kind t with | (Rel _ | Meta _ | Var _ | Evar _ | Const _ | Proj _ | Case _ | Fix _ | CoFix _ | Ind _) -> AtomicType (t,[||]) - | (Lambda _ | Construct _) -> failwith "Not a type" + | (Lambda _ | Construct _ | Int _) -> failwith "Not a type" diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml index f9fdbdd68e..3cb5d17d34 100644 --- a/kernel/term_typing.ml +++ b/kernel/term_typing.ml @@ -93,10 +93,45 @@ let infer_declaration (type a) ~(trust : a trust) env (dcl : a constant_entry) = cook_context = ctx; } + (** Primitives cannot be universe polymorphic *) + | PrimitiveEntry ({ prim_entry_type = otyp; + prim_entry_univs = uctxt; + prim_entry_content = op_t; + }) -> + let env = push_context_set ~strict:true uctxt env in + let ty = match otyp with + | Some typ -> + let tyj = infer_type env typ in + check_primitive_type env op_t tyj.utj_val; + Constr.hcons tyj.utj_val + | None -> + match op_t with + | CPrimitives.OT_op op -> type_of_prim env op + | CPrimitives.OT_type _ -> mkSet + in + let cd = + match op_t with + | CPrimitives.OT_op op -> Declarations.Primitive op + | CPrimitives.OT_type _ -> Undef None in + { Cooking.cook_body = cd; + cook_type = ty; + cook_universes = Monomorphic_const uctxt; + cook_private_univs = None; + cook_inline = false; + cook_context = None + } + (** Definition [c] is opaque (Qed), non polymorphic and with a specified type, so we delay the typing and hash consing of its body. Remark: when the universe quantification is given explicitly, we could delay even in the polymorphic case. *) + +(** Definition is opaque (Qed) and non polymorphic with known type, so we delay +the typing and hash consing of its body. + +TODO: if the universe quantification is given explicitly, we could delay even in +the polymorphic case + *) | DefinitionEntry ({ const_entry_type = Some typ; const_entry_opaque = true; const_entry_universes = Monomorphic_const_entry univs; _ } as c) -> @@ -238,7 +273,7 @@ let build_constant_declaration _kn env result = we must look at the body NOW, if any *) let ids_typ = global_vars_set env typ in let ids_def = match def with - | Undef _ -> Id.Set.empty + | Undef _ | Primitive _ -> Id.Set.empty | Def cs -> global_vars_set env (Mod_subst.force_constr cs) | OpaqueDef lc -> let vars = @@ -258,7 +293,7 @@ let build_constant_declaration _kn env result = (* We use the declared set and chain a check of correctness *) sort declared, match def with - | Undef _ as x -> x (* nothing to check *) + | Undef _ | Primitive _ as x -> x (* nothing to check *) | Def cs as x -> let ids_typ = global_vars_set env typ in let ids_def = global_vars_set env (Mod_subst.force_constr cs) in @@ -316,6 +351,7 @@ let translate_local_def env _id centry = if Option.is_empty decl.cook_context && !Flags.record_aux_file then begin match decl.cook_body with | Undef _ -> () + | Primitive _ -> () | Def _ -> () | OpaqueDef lc -> let ids_typ = global_vars_set env typ in @@ -336,7 +372,7 @@ let translate_local_def env _id centry = the body by virtue of the typing of [Entries.section_def_entry]. *) let () = assert (Univ.ContextSet.is_empty cst) in p - | Undef _ -> assert false + | Undef _ | Primitive _ -> assert false in c, typ diff --git a/kernel/typeops.ml b/kernel/typeops.ml index c9acd168e8..7eb8e803b3 100644 --- a/kernel/typeops.ml +++ b/kernel/typeops.ml @@ -12,7 +12,7 @@ open CErrors open Util open Names open Univ -open Sorts +open Term open Constr open Vars open Declarations @@ -187,6 +187,74 @@ let type_of_apply env func funt argsv argstv = in apply_rec 0 (inject funt) +(* Type of primitive constructs *) +let type_of_prim_type _env = function + | CPrimitives.PT_int63 -> 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_prim env t = + let int_ty = type_of_int env in + let bool_ty () = + match env.retroknowledge.Retroknowledge.retro_bool with + | Some ((ind,_),_) -> Constr.mkInd ind + | None -> CErrors.user_err Pp.(str"The type bool must be registered before this primitive.") + in + let compare_ty () = + match env.retroknowledge.Retroknowledge.retro_cmp with + | Some ((ind,_),_,_) -> Constr.mkInd ind + | None -> CErrors.user_err Pp.(str"The type compare 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|]) + | None -> CErrors.user_err Pp.(str"The type pair must be registered before this primitive.") + in + let carry_ty int_ty = + match env.retroknowledge.Retroknowledge.retro_carry with + | 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(Name (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 judge_of_int env i = + make_judge (Constr.mkInt i) (type_of_int env) + (* Type of product *) let sort_of_product env domsort rangsort = @@ -468,6 +536,9 @@ let rec execute env cstr = let (fix_ty,recdef') = execute_recdef env recdef i in let cofix = (i,recdef') in check_cofix env cofix; fix_ty + + (* Primitive types *) + | Int _ -> type_of_int env (* Partial proofs: unsupported by the kernel *) | Meta _ -> @@ -590,3 +661,20 @@ let judge_of_case env ci pj cj lfj = let lf, lft = dest_judgev lfj in make_judge (mkCase (ci, (*nf_betaiota*) pj.uj_val, cj.uj_val, lft)) (type_of_case env ci pj.uj_val pj.uj_type cj.uj_val cj.uj_type lf lft) + +(* Building type of primitive operators and type *) + +open CPrimitives + +let check_primitive_type env op_t t = + match op_t with + | OT_type PT_int63 -> + (try + default_conv ~l2r:false CUMUL env mkSet t + with NotConvertible -> + CErrors.user_err Pp.(str"Was expecting the sort of this primitive type to be Set")) + | OT_op p -> + (try + default_conv ~l2r:false CUMUL env (type_of_prim env p) t + with NotConvertible -> + CErrors.user_err Pp.(str"Not the expected type for this primitive")) diff --git a/kernel/typeops.mli b/kernel/typeops.mli index 4193324136..52c261c5e8 100644 --- a/kernel/typeops.mli +++ b/kernel/typeops.mli @@ -118,3 +118,13 @@ val constr_of_global_in_context : env -> GlobRef.t -> types * Univ.AUContext.t (** Check that hyps are included in env and fails with error otherwise *) val check_hyps_inclusion : env -> ?evars:((existential->constr option) * UGraph.t) -> ('a -> constr) -> 'a -> Constr.named_context -> unit + +val check_primitive_type : env -> CPrimitives.op_or_type -> types -> unit + +(** Types for primitives *) + +val type_of_int : env -> types +val judge_of_int : env -> Uint63.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/uint31.ml b/kernel/uint31.ml deleted file mode 100644 index d9c723c243..0000000000 --- a/kernel/uint31.ml +++ /dev/null @@ -1,153 +0,0 @@ - (* Invariant: For arch64 all extra bytes are set to 0 *) -type t = int - - (* to be used only on 32 bits architectures *) -let maxuint31 = Int32.of_string "0x7FFFFFFF" -let uint_32 i = Int32.logand (Int32.of_int i) maxuint31 - -let select f32 f64 = if Sys.word_size = 64 then f64 else f32 - - (* conversion to an int *) -let to_int i = i - -let of_int_32 i = i -let of_int_64 i = i land 0x7FFFFFFF - -let of_int = select of_int_32 of_int_64 -let of_uint i = i - - (* conversion of an uint31 to a string *) -let to_string_32 i = Int32.to_string (uint_32 i) -let to_string_64 = string_of_int - -let to_string = select to_string_32 to_string_64 -let of_string s = - let i32 = Int32.of_string s in - if Int32.compare Int32.zero i32 <= 0 - && Int32.compare i32 maxuint31 <= 0 - then Int32.to_int i32 - else raise (Failure "int_of_string") - - - - (* logical shift *) -let l_sl x y = - of_int (if 0 <= y && y < 31 then x lsl y else 0) - -let l_sr x y = - if 0 <= y && y < 31 then x lsr y else 0 - -let l_and x y = x land y -let l_or x y = x lor y -let l_xor x y = x lxor y - - (* addition of int31 *) -let add x y = of_int (x + y) - - (* subtraction *) -let sub x y = of_int (x - y) - - (* multiplication *) -let mul x y = of_int (x * y) - - (* exact multiplication *) -let mulc_32 x y = - let x = Int64.of_int32 (uint_32 x) in - let y = Int64.of_int32 (uint_32 y) in - let m = Int64.mul x y in - let l = Int64.to_int m in - let h = Int64.to_int (Int64.shift_right_logical m 31) in - h,l - -let mulc_64 x y = - let m = x * y in - let l = of_int_64 m in - let h = of_int_64 (m lsr 31) in - h, l -let mulc = select mulc_32 mulc_64 - - (* division *) -let div_32 x y = - if y = 0 then 0 else - Int32.to_int (Int32.div (uint_32 x) (uint_32 y)) -let div_64 x y = if y = 0 then 0 else x / y -let div = select div_32 div_64 - - (* modulo *) -let rem_32 x y = - if y = 0 then 0 - else Int32.to_int (Int32.rem (uint_32 x) (uint_32 y)) -let rem_64 x y = if y = 0 then 0 else x mod y -let rem = select rem_32 rem_64 - - (* division of two numbers by one *) -let div21_32 xh xl y = - if y = 0 then (0,0) - else - let x = - Int64.logor - (Int64.shift_left (Int64.of_int32 (uint_32 xh)) 31) - (Int64.of_int32 (uint_32 xl)) in - let y = Int64.of_int32 (uint_32 y) in - let q = Int64.div x y in - let r = Int64.rem x y in - Int64.to_int q, Int64.to_int r -let div21_64 xh xl y = - if y = 0 then (0,0) - else - let x = (xh lsl 31) lor xl in - let q = x / y in - let r = x mod y in - q, r -let div21 = select div21_32 div21_64 - - (* comparison *) -let lt_32 x y = (x lxor 0x40000000) < (y lxor 0x40000000) - -(* Do not remove the type information it is really important for - efficiency *) -let lt_64 (x:int) (y:int) = x < y -let lt = select lt_32 lt_64 - -let le_32 x y = - (x lxor 0x40000000) <= (y lxor 0x40000000) - -(* Do not remove the type information it is really important for - efficiency *) -let le_64 (x:int) (y:int) = x <= y -let le = select le_32 le_64 - -let equal (x:int) (y:int) = x == y - -let cmp_32 x y = Int32.compare (uint_32 x) (uint_32 y) -(* Do not remove the type information it is really important for - efficiency *) -let cmp_64 (x:int) (y:int) = compare x y -let compare = select cmp_32 cmp_64 - - (* head tail *) - -let head0 x = - let r = ref 0 in - let x = ref x in - if !x land 0x7FFF0000 = 0 then r := !r + 15 - else x := !x lsr 15; - if !x land 0xFF00 = 0 then (x := !x lsl 8; r := !r + 8); - if !x land 0xF000 = 0 then (x := !x lsl 4; r := !r + 4); - if !x land 0xC000 = 0 then (x := !x lsl 2; r := !r + 2); - if !x land 0x8000 = 0 then (x := !x lsl 1; r := !r + 1); - if !x land 0x8000 = 0 then ( r := !r + 1); - !r;; - -let tail0 x = - let r = ref 0 in - let x = ref x in - if !x land 0xFFFF = 0 then (x := !x lsr 16; r := !r + 16); - if !x land 0xFF = 0 then (x := !x lsr 8; r := !r + 8); - if !x land 0xF = 0 then (x := !x lsr 4; r := !r + 4); - if !x land 0x3 = 0 then (x := !x lsr 2; r := !r + 2); - if !x land 0x1 = 0 then ( r := !r + 1); - !r - -let add_digit x d = - (x lsl 1) lor d diff --git a/kernel/uint31.mli b/kernel/uint63.mli index d1f933cc4e..b5f40ca804 100644 --- a/kernel/uint31.mli +++ b/kernel/uint63.mli @@ -1,14 +1,28 @@ type t +val uint_size : int +val maxuint31 : t + (* conversion to int *) -val to_int : t -> int val of_int : int -> t +val to_int2 : t -> int * int (* msb, lsb *) +val of_int64 : Int64.t -> t +(* val of_uint : int -> t +*) + +val hash : t -> int - (* conversion to a string *) + (* convertion to a string *) val to_string : t -> string val of_string : string -> t +val compile : t -> string + +(* constants *) +val zero : t +val one : t + (* logical operations *) val l_sl : t -> t -> t val l_sr : t -> t -> t @@ -16,20 +30,21 @@ val l_and : t -> t -> t val l_xor : t -> t -> t val l_or : t -> t -> t - (* Arithmetic operations *) + (* Arithmetic operations *) val add : t -> t -> t val sub : t -> t -> t val mul : t -> t -> t val div : t -> t -> t val rem : t -> t -> t - + (* Specific arithmetic operations *) val mulc : t -> t -> t * t +val addmuldiv : t -> t -> t -> t val div21 : t -> t -> t -> t * t - + (* comparison *) val lt : t -> t -> bool -val equal : t -> t -> bool +val equal : t -> t -> bool val le : t -> t -> bool val compare : t -> t -> int @@ -37,5 +52,4 @@ val compare : t -> t -> int val head0 : t -> t val tail0 : t -> t -(** Used by retroknowledge *) -val add_digit : t -> t -> t +val is_uint63 : Obj.t -> bool diff --git a/kernel/uint63_amd64.ml b/kernel/uint63_amd64.ml new file mode 100644 index 0000000000..010b594de8 --- /dev/null +++ b/kernel/uint63_amd64.ml @@ -0,0 +1,215 @@ +(************************************************************************) +(* * 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) *) +(************************************************************************) + +type t = int + +let _ = assert (Sys.word_size = 64) + +let uint_size = 63 + +let maxuint63 = Int64.of_string "0x7FFFFFFFFFFFFFFF" +let maxuint31 = 0x7FFFFFFF + + (* conversion from an int *) +let to_uint64 i = Int64.logand (Int64.of_int i) maxuint63 + +let of_int i = i +[@@ocaml.inline always] + +let to_int2 i = (0,i) + +let of_int64 _i = assert false + +let hash i = i +[@@ocaml.inline always] + + (* conversion of an uint63 to a string *) +let to_string i = Int64.to_string (to_uint64 i) + +let of_string s = + let i64 = Int64.of_string s in + if Int64.compare Int64.zero i64 <= 0 + && Int64.compare i64 maxuint63 <= 0 + then Int64.to_int i64 + else raise (Failure "Int64.of_string") + +(* Compiles an unsigned int to OCaml code *) +let compile i = Printf.sprintf "Uint63.of_int (%i)" i + +let zero = 0 +let one = 1 + + (* logical shift *) +let l_sl x y = + if 0 <= y && y < 63 then x lsl y else 0 + +let l_sr x y = + if 0 <= y && y < 63 then x lsr y else 0 + +let l_and x y = x land y +[@@ocaml.inline always] + +let l_or x y = x lor y +[@@ocaml.inline always] + +let l_xor x y = x lxor y +[@@ocaml.inline always] + + (* addition of int63 *) +let add x y = x + y +[@@ocaml.inline always] + + (* subtraction *) +let sub x y = x - y +[@@ocaml.inline always] + + (* multiplication *) +let mul x y = x * y +[@@ocaml.inline always] + + (* division *) +let div (x : int) (y : int) = + if y = 0 then 0 else Int64.to_int (Int64.div (to_uint64 x) (to_uint64 y)) + + (* modulo *) +let rem (x : int) (y : int) = + if y = 0 then 0 else Int64.to_int (Int64.rem (to_uint64 x) (to_uint64 y)) + +let addmuldiv p x y = + l_or (l_sl x p) (l_sr y (uint_size - p)) + + (* comparison *) +let lt (x : int) (y : int) = + (x lxor 0x4000000000000000) < (y lxor 0x4000000000000000) +[@@ocaml.inline always] + +let le (x : int) (y : int) = + (x lxor 0x4000000000000000) <= (y lxor 0x4000000000000000) +[@@ocaml.inline always] + +(* A few helper functions on 128 bits *) +let lt128 xh xl yh yl = + lt xh yh || (xh = yh && lt xl yl) + +let le128 xh xl yh yl = + lt xh yh || (xh = yh && le xl yl) + + (* division of two numbers by one *) +let div21 xh xl y = + let maskh = ref 0 in + let maskl = ref 1 in + let dh = ref 0 in + let dl = ref y in + let cmp = ref true in + while !dh >= 0 && !cmp do + cmp := lt128 !dh !dl xh xl; + (* We don't use addmuldiv below to avoid checks on 1 *) + dh := (!dh lsl 1) lor (!dl lsr (uint_size - 1)); + dl := !dl lsl 1; + maskh := (!maskh lsl 1) lor (!maskl lsr (uint_size - 1)); + maskl := !maskl lsl 1 + done; (* mask = 2^N, d = 2^N * d, d >= x *) + let remh = ref xh in + let reml = ref xl in + let quotient = ref 0 in + while !maskh lor !maskl <> 0 do + if le128 !dh !dl !remh !reml then begin (* if rem >= d, add one bit and subtract d *) + quotient := !quotient lor !maskl; + remh := if lt !reml !dl then !remh - !dh - 1 else !remh - !dh; + reml := !reml - !dl; + end; + maskl := (!maskl lsr 1) lor (!maskh lsl (uint_size - 1)); + maskh := !maskh lsr 1; + dl := (!dl lsr 1) lor (!dh lsl (uint_size - 1)); + dh := !dh lsr 1; + done; + !quotient, !reml + + (* exact multiplication *) +(* TODO: check that none of these additions could be a logical or *) + + +(* size = 32 + 31 + (hx << 31 + lx) * (hy << 31 + ly) = + hxhy << 62 + (hxly + lxhy) << 31 + lxly +*) + +(* precondition : (x lsr 62 = 0 || y lsr 62 = 0) *) +let mulc_aux x y = + let lx = x land maxuint31 in + let ly = y land maxuint31 in + let hx = x lsr 31 in + let hy = y lsr 31 in + (* hx and hy are 32 bits value but at most one is 32 *) + let hxy = hx * hy in (* 63 bits *) + let hxly = hx * ly in (* 63 bits *) + let lxhy = lx * hy in (* 63 bits *) + let lxy = lx * ly in (* 62 bits *) + let l = lxy lor (hxy lsl 62) (* 63 bits *) in + let h = hxy lsr 1 in (* 62 bits *) + let hl = hxly + lxhy in (* We can have a carry *) + let h = if lt hl hxly then h + (1 lsl 31) else h in + let hl'= hl lsl 31 in + let l = l + hl' in + let h = if lt l hl' then h + 1 else h in + let h = h + (hl lsr 32) in + (h,l) + +let mulc x y = + if (x lsr 62 = 0 || y lsr 62 = 0) then mulc_aux x y + else + let yl = y lxor (1 lsl 62) in + let (h,l) = mulc_aux x yl in + (* h << 63 + l = x * yl + x * y = x * (1 << 62 + yl) = + x << 62 + x*yl = x << 62 + h << 63 + l *) + let l' = l + (x lsl 62) in + let h = if lt l' l then h + 1 else h in + (h + (x lsr 1), l') + +let equal (x : int) (y : int) = x = y +[@@ocaml.inline always] + +let compare (x:int) (y:int) = + let x = x lxor 0x4000000000000000 in + let y = y lxor 0x4000000000000000 in + if x > y then 1 + else if y > x then -1 + else 0 + + (* head tail *) + +let head0 x = + let r = ref 0 in + let x = ref x in + if !x land 0x7FFFFFFF00000000 = 0 then r := !r + 31 + else x := !x lsr 31; + if !x land 0xFFFF0000 = 0 then (x := !x lsl 16; r := !r + 16); + if !x land 0xFF000000 = 0 then (x := !x lsl 8; r := !r + 8); + if !x land 0xF0000000 = 0 then (x := !x lsl 4; r := !r + 4); + if !x land 0xC0000000 = 0 then (x := !x lsl 2; r := !r + 2); + if !x land 0x80000000 = 0 then (x := !x lsl 1; r := !r + 1); + if !x land 0x80000000 = 0 then ( r := !r + 1); + !r;; + +let tail0 x = + let r = ref 0 in + let x = ref x in + if !x land 0xFFFFFFFF = 0 then (x := !x lsr 32; r := !r + 32); + if !x land 0xFFFF = 0 then (x := !x lsr 16; r := !r + 16); + if !x land 0xFF = 0 then (x := !x lsr 8; r := !r + 8); + if !x land 0xF = 0 then (x := !x lsr 4; r := !r + 4); + if !x land 0x3 = 0 then (x := !x lsr 2; r := !r + 2); + if !x land 0x1 = 0 then ( r := !r + 1); + !r + +let is_uint63 t = + Obj.is_int t +[@@ocaml.inline always] diff --git a/kernel/uint63_x86.ml b/kernel/uint63_x86.ml new file mode 100644 index 0000000000..461184c432 --- /dev/null +++ b/kernel/uint63_x86.ml @@ -0,0 +1,209 @@ +(************************************************************************) +(* * 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) *) +(************************************************************************) + +(* Invariant: the msb should be 0 *) +type t = Int64.t + +let _ = assert (Sys.word_size = 32) + +let uint_size = 63 + +let maxuint63 = Int64.of_string "0x7FFFFFFFFFFFFFFF" +let maxuint31 = Int64.of_string "0x7FFFFFFF" + +let zero = Int64.zero +let one = Int64.one + + (* conversion from an int *) +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 hash i = + let (h,l) = to_int2 i in + (*Hashset.combine h l*) + h * 65599 + l + + (* conversion of an uint63 to a string *) +let to_string i = Int64.to_string i + +let of_string s = + let i64 = Int64.of_string s in + if Int64.compare Int64.zero i64 <= 0 + && Int64.compare i64 maxuint63 <= 0 + then i64 + else raise (Failure "Int63.of_string") + +(* Compiles an unsigned int to OCaml code *) +let compile i = Printf.sprintf "Uint63.of_int64 (%LiL)" i + + (* comparison *) +let lt x y = + Int64.compare x y < 0 + +let le x y = + Int64.compare x y <= 0 + + (* logical shift *) +let l_sl x y = + if le 0L y && lt y 63L then mask63 (Int64.shift_left x (Int64.to_int y)) else 0L + +let l_sr x y = + if le 0L y && lt y 63L then Int64.shift_right x (Int64.to_int y) else 0L + +let l_and x y = Int64.logand x y +let l_or x y = Int64.logor x y +let l_xor x y = Int64.logxor x y + + (* addition of int63 *) +let add x y = mask63 (Int64.add x y) + +let addcarry x y = add (add x y) Int64.one + + (* subtraction *) +let sub x y = mask63 (Int64.sub x y) + +let subcarry x y = sub (sub x y) Int64.one + + (* multiplication *) +let mul x y = mask63 (Int64.mul x y) + + (* division *) +let div x y = + if y = 0L then 0L else Int64.div x y + + (* modulo *) +let rem x y = + if y = 0L then 0L else Int64.rem x y + +let addmuldiv p x y = + l_or (l_sl x p) (l_sr y Int64.(sub (of_int uint_size) p)) + +(* A few helper functions on 128 bits *) +let lt128 xh xl yh yl = + lt xh yh || (xh = yh && lt xl yl) + +let le128 xh xl yh yl = + lt xh yh || (xh = yh && le xl yl) + + (* division of two numbers by one *) +let div21 xh xl y = + let maskh = ref zero in + let maskl = ref one in + let dh = ref zero in + let dl = ref y in + let cmp = ref true in + while le zero !dh && !cmp do + cmp := lt128 !dh !dl xh xl; + (* We don't use addmuldiv below to avoid checks on 1 *) + dh := l_or (l_sl !dh one) (l_sr !dl (of_int (uint_size - 1))); + dl := l_sl !dl one; + maskh := l_or (l_sl !maskh one) (l_sr !maskl (of_int (uint_size - 1))); + maskl := l_sl !maskl one + done; (* mask = 2^N, d = 2^N * d, d >= x *) + let remh = ref xh in + let reml = ref xl in + let quotient = ref zero in + while not (Int64.equal (l_or !maskh !maskl) zero) do + if le128 !dh !dl !remh !reml then begin (* if rem >= d, add one bit and subtract d *) + quotient := l_or !quotient !maskl; + remh := if lt !reml !dl then sub (sub !remh !dh) one else sub !remh !dh; + reml := sub !reml !dl + end; + maskl := l_or (l_sr !maskl one) (l_sl !maskh (of_int (uint_size - 1))); + maskh := l_sr !maskh one; + dl := l_or (l_sr !dl one) (l_sl !dh (of_int (uint_size - 1))); + dh := l_sr !dh one + done; + !quotient, !reml + + + (* exact multiplication *) +let mulc x y = + let lx = ref (Int64.logand x maxuint31) in + let ly = ref (Int64.logand y maxuint31) in + let hx = Int64.shift_right x 31 in + let hy = Int64.shift_right y 31 in + let hr = ref (Int64.mul hx hy) in + let lr = ref (Int64.logor (Int64.mul !lx !ly) (Int64.shift_left !hr 62)) in + hr := (Int64.shift_right_logical !hr 1); + lx := Int64.mul !lx hy; + ly := Int64.mul hx !ly; + hr := Int64.logor !hr (Int64.add (Int64.shift_right !lx 32) (Int64.shift_right !ly 32)); + lr := Int64.add !lr (Int64.shift_left !lx 31); + hr := Int64.add !hr (Int64.shift_right_logical !lr 63); + lr := Int64.add (Int64.shift_left !ly 31) (mask63 !lr); + hr := Int64.add !hr (Int64.shift_right_logical !lr 63); + if Int64.logand !lr Int64.min_int <> 0L + then Int64.(sub !hr one, mask63 !lr) + else (!hr, !lr) + +let equal x y = mask63 x = mask63 y + +let compare x y = Int64.compare x y + +(* Number of leading zeroes *) +let head0 x = + let r = ref 0 in + let x = ref x in + if Int64.logand !x 0x7FFFFFFF00000000L = 0L then r := !r + 31 + else x := Int64.shift_right !x 31; + if Int64.logand !x 0xFFFF0000L = 0L then (x := Int64.shift_left !x 16; r := !r + 16); + if Int64.logand !x 0xFF000000L = 0L then (x := Int64.shift_left !x 8; r := !r + 8); + if Int64.logand !x 0xF0000000L = 0L then (x := Int64.shift_left !x 4; r := !r + 4); + if Int64.logand !x 0xC0000000L = 0L then (x := Int64.shift_left !x 2; r := !r + 2); + if Int64.logand !x 0x80000000L = 0L then (x := Int64.shift_left !x 1; r := !r + 1); + if Int64.logand !x 0x80000000L = 0L then (r := !r + 1); + Int64.of_int !r + +(* Number of trailing zeroes *) +let tail0 x = + let r = ref 0 in + let x = ref x in + if Int64.logand !x 0xFFFFFFFFL = 0L then (x := Int64.shift_right !x 32; r := !r + 32); + if Int64.logand !x 0xFFFFL = 0L then (x := Int64.shift_right !x 16; r := !r + 16); + if Int64.logand !x 0xFFL = 0L then (x := Int64.shift_right !x 8; r := !r + 8); + if Int64.logand !x 0xFL = 0L then (x := Int64.shift_right !x 4; r := !r + 4); + if Int64.logand !x 0x3L = 0L then (x := Int64.shift_right !x 2; r := !r + 2); + if Int64.logand !x 0x1L = 0L then (r := !r + 1); + Int64.of_int !r + +(* May an object be safely cast into an Uint63.t ? *) +let is_uint63 t = + Obj.is_block t && Int.equal (Obj.tag t) Obj.custom_tag + && le (Obj.magic t) maxuint63 + +(* Register all exported functions so that they can be called from C code *) + +let () = + Callback.register "uint63 add" add; + Callback.register "uint63 addcarry" addcarry; + Callback.register "uint63 addmuldiv" addmuldiv; + Callback.register "uint63 div" div; + Callback.register "uint63 div21_ml" div21; + Callback.register "uint63 eq" equal; + Callback.register "uint63 eq0" (equal Int64.zero); + Callback.register "uint63 head0" head0; + Callback.register "uint63 land" l_and; + Callback.register "uint63 leq" le; + Callback.register "uint63 lor" l_or; + Callback.register "uint63 lsl" l_sl; + Callback.register "uint63 lsl1" (fun x -> l_sl x Int64.one); + Callback.register "uint63 lsr" l_sr; + Callback.register "uint63 lsr1" (fun x -> l_sr x Int64.one); + Callback.register "uint63 lt" lt; + Callback.register "uint63 lxor" l_xor; + Callback.register "uint63 mod" rem; + Callback.register "uint63 mul" mul; + Callback.register "uint63 mulc_ml" mulc; + Callback.register "uint63 one" one; + Callback.register "uint63 sub" sub; + Callback.register "uint63 subcarry" subcarry; + Callback.register "uint63 tail0" tail0 diff --git a/kernel/vconv.ml b/kernel/vconv.ml index 246c90c09d..04a17f7b08 100644 --- a/kernel/vconv.ml +++ b/kernel/vconv.ml @@ -71,13 +71,15 @@ and conv_whd env pb k whd1 whd2 cu = done; !rcu else raise NotConvertible + | Vint64 i1, Vint64 i2 -> + if Int64.equal i1 i2 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 _ -> (* on the fly eta expansion *) conv_val env CONV (k+1) (apply_whd k whd1) (apply_whd k whd2) cu - | Vprod _, _ | Vfix _, _ | Vcofix _, _ | Vconstr_const _, _ + | Vprod _, _ | Vfix _, _ | Vcofix _, _ | Vconstr_const _, _ | Vint64 _, _ | Vconstr_block _, _ | Vatom_stk _, _ -> raise NotConvertible diff --git a/kernel/vm.ml b/kernel/vm.ml index eaf64ba4af..83312a8530 100644 --- a/kernel/vm.ml +++ b/kernel/vm.ml @@ -7,7 +7,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (* * (see LICENSE file for the text of the license) *) (************************************************************************) - open Vmvalues external set_drawinstr : unit -> unit = "coq_set_drawinstr" @@ -170,7 +169,7 @@ 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 _ -> assert false + | Vprod _ | Vconstr_const _ | Vconstr_block _ | Vint64 _ -> 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 217ef4b8e5..9a3eadf747 100644 --- a/kernel/vmvalues.ml +++ b/kernel/vmvalues.ml @@ -57,6 +57,7 @@ type structured_constant = | Const_b0 of tag | Const_univ_level of Univ.Level.t | Const_val of structured_values + | Const_uint of Uint63.t type reloc_table = (tag * int) array @@ -73,7 +74,9 @@ let rec eq_structured_values v1 v2 = let t2 = Obj.tag o2 in if Int.equal t1 t2 && Int.equal (Obj.size o1) (Obj.size o2) - then begin + then if Int.equal t1 Obj.custom_tag + then Int64.equal (Obj.magic v1 : int64) (Obj.magic v2 : int64) + else begin assert (t1 <= Obj.last_non_constant_constructor_tag && t2 <= Obj.last_non_constant_constructor_tag); let i = ref 0 in @@ -100,7 +103,9 @@ let eq_structured_constant c1 c2 = match c1, c2 with | Const_univ_level l1 , Const_univ_level l2 -> Univ.Level.equal l1 l2 | Const_univ_level _ , _ -> false | Const_val v1, Const_val v2 -> eq_structured_values v1 v2 -| Const_val _v1, _ -> false +| Const_val _, _ -> false +| Const_uint i1, Const_uint i2 -> Uint63.equal i1 i2 +| Const_uint _, _ -> false let hash_structured_constant c = let open Hashset.Combine in @@ -110,6 +115,7 @@ let hash_structured_constant c = | Const_b0 t -> combinesmall 3 (Int.hash t) | 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) let eq_annot_switch asw1 asw2 = let eq_ci ci1 ci2 = @@ -142,6 +148,7 @@ let pp_struct_const = function | Const_b0 i -> Pp.int i | Const_univ_level l -> Univ.Level.pr l | Const_val _ -> Pp.str "(value)" + | Const_uint i -> Pp.str (Uint63.to_string i) (* Abstract data *) type vprod @@ -276,6 +283,7 @@ type whd = | Vcofix of vcofix * to_update * arguments option | Vconstr_const of int | Vconstr_block of vblock + | Vint64 of int64 | Vatom_stk of atom * stack | Vuniv_level of Univ.Level.t @@ -306,8 +314,9 @@ let uni_lvl_val (v : values) : Univ.Level.t = | Vcofix _ -> str "Vcofix" | Vconstr_const _i -> str "Vconstr_const" | Vconstr_block _b -> str "Vconstr_block" + | Vint64 _ -> str "Vint64" | Vatom_stk (_a,_stk) -> str "Vatom_stk" - | _ -> assert false + | Vuniv_level _ -> assert false in CErrors.anomaly Pp.( strbrk "Parsing virtual machine value expected universe level, got " @@ -363,6 +372,8 @@ let rec whd_accu a stk = | [Zapp args] -> Vcofix(vcofix, res, Some args) | _ -> assert false end + | i when Int.equal i Obj.custom_tag -> + Vint64 (Obj.magic i) | tg -> CErrors.anomaly Pp.(strbrk "Failed to parse VM value. Tag = " ++ int tg ++ str ".") @@ -391,6 +402,7 @@ let whd_val : values -> whd = | 2 -> Vfix(Obj.obj (Obj.field o 1), Some (Obj.obj o)) | 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 Vconstr_block(Obj.obj o) @@ -413,6 +425,7 @@ let obj_of_str_const str = | Const_b0 tag -> Obj.repr tag | Const_univ_level l -> Obj.repr (Vuniv_level l) | Const_val v -> Obj.repr v + | Const_uint i -> Obj.repr i let val_of_block tag (args : structured_values array) = let nargs = Array.length args in @@ -430,6 +443,8 @@ let val_of_atom a = val_of_obj (obj_of_atom a) let val_of_int i = (Obj.magic i : values) +let val_of_uint i = (Obj.magic i : values) + let atom_of_proj kn v = let r = Obj.new_block proj_tag 2 in Obj.set_field r 0 (Obj.repr kn); @@ -659,6 +674,7 @@ and pr_whd w = | Vcofix _ -> str "Vcofix" | Vconstr_const i -> str "Vconstr_const(" ++ int i ++ str ")" | Vconstr_block _b -> str "Vconstr_block" + | Vint64 i -> i |> Format.sprintf "Vint64(%LiL)" |> 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 ae1d416ed5..6d984d5f49 100644 --- a/kernel/vmvalues.mli +++ b/kernel/vmvalues.mli @@ -44,6 +44,7 @@ type structured_constant = | Const_b0 of tag | Const_univ_level of Univ.Level.t | Const_val of structured_values + | Const_uint of Uint63.t val pp_struct_const : structured_constant -> Pp.t @@ -125,6 +126,7 @@ type whd = | Vcofix of vcofix * to_update * arguments option | Vconstr_const of int | Vconstr_block of vblock + | Vint64 of int64 | Vatom_stk of atom * stack | Vuniv_level of Univ.Level.t @@ -145,6 +147,7 @@ val val_of_proj : Projection.Repr.t -> values -> values val val_of_atom : atom -> values val val_of_int : int -> structured_values val val_of_block : tag -> structured_values array -> structured_values +val val_of_uint : Uint63.t -> structured_values external val_of_annot_switch : annot_switch -> values = "%identity" external val_of_proj_name : Projection.Repr.t -> values = "%identity" diff --git a/kernel/write_uint63.ml b/kernel/write_uint63.ml new file mode 100644 index 0000000000..0fcaf4f10a --- /dev/null +++ b/kernel/write_uint63.ml @@ -0,0 +1,30 @@ +(** Equivalent of rm -f *) + +let safe_remove f = + try Unix.chmod f 0o644; Sys.remove f with _ -> () + +(** * Generate an implementation of 63-bit arithmetic *) + +let ml_file_copy input output = + safe_remove output; + let i = open_in input in + let o = open_out output in + let pr s = Printf.fprintf o s in + pr "(* DO NOT EDIT THIS FILE: automatically generated by ./write_uint63.ml *)\n"; + pr "(* see uint63_amd64.ml and uint63_x86.ml *)\n"; + try + while true do + output_string o (input_line i); output_char o '\n' + done + with End_of_file -> + close_in i; + close_out o; + Unix.chmod output 0o444 + +let write_uint63 () = + ml_file_copy + (if max_int = 1073741823 (* 32-bits *) then "uint63_x86.ml" + else (* 64 bits *) "uint63_amd64.ml") + "uint63.ml" + +let () = write_uint63 () diff --git a/library/decl_kinds.ml b/library/decl_kinds.ml index 171d51800e..8d5c2fb687 100644 --- a/library/decl_kinds.ml +++ b/library/decl_kinds.ml @@ -70,6 +70,7 @@ type goal_kind = locality * polymorphic * goal_object_kind (** Kinds used in library *) type logical_kind = + | IsPrimitive | IsAssumption of assumption_object_kind | IsDefinition of definition_object_kind | IsProof of theorem_kind diff --git a/library/global.ml b/library/global.ml index 84d2a37170..784a02449c 100644 --- a/library/global.ml +++ b/library/global.ml @@ -175,11 +175,8 @@ let with_global f = let (a, ctx) = f (env ()) (current_dirpath ()) in push_context_set false ctx; a -(* spiwack: register/unregister functions for retroknowledge *) -let register field value = - globalize0 (Safe_typing.register field value) - let register_inline c = globalize0 (Safe_typing.register_inline c) +let register_inductive c r = globalize0 (Safe_typing.register_inductive c r) let set_strategy k l = globalize0 (Safe_typing.set_strategy k l) diff --git a/library/global.mli b/library/global.mli index 40962e21af..df18625a5f 100644 --- a/library/global.mli +++ b/library/global.mli @@ -142,10 +142,8 @@ val universes_of_global : GlobRef.t -> Univ.AUContext.t (** {6 Retroknowledge } *) -val register : - Retroknowledge.field -> GlobRef.t -> unit - val register_inline : Constant.t -> unit +val register_inductive : inductive -> CPrimitives.prim_ind -> unit (** {6 Oracle } *) diff --git a/library/keys.ml b/library/keys.ml index 58883ccc88..45b6fae2f0 100644 --- a/library/keys.ml +++ b/library/keys.ml @@ -25,13 +25,14 @@ type key = | KFix | KCoFix | KRel + | KInt module KeyOrdered = struct type t = key let hash gr = match gr with - | KGlob gr -> 8 + GlobRef.Ordered.hash gr + | KGlob gr -> 9 + GlobRef.Ordered.hash gr | KLam -> 0 | KLet -> 1 | KProd -> 2 @@ -40,6 +41,7 @@ module KeyOrdered = struct | KFix -> 5 | KCoFix -> 6 | KRel -> 7 + | KInt -> 8 let compare gr1 gr2 = match gr1, gr2 with @@ -133,6 +135,7 @@ let constr_key kind c = | Evar _ -> raise Not_found | Sort _ -> KSort | LetIn _ -> KLet + | Int _ -> KInt in Some (aux c) with Not_found -> None @@ -148,6 +151,7 @@ let pr_key pr_global = function | KFix -> str"Fix" | KCoFix -> str"CoFix" | KRel -> str"Rel" + | KInt -> str"Int" let pr_keyset pr_global v = prlist_with_sep spc (pr_key pr_global) (Keyset.elements v) diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml index b0f6301192..b59e3b608c 100644 --- a/plugins/extraction/extract_env.ml +++ b/plugins/extraction/extract_env.ml @@ -147,7 +147,7 @@ let check_fix env sg cb i = | Fix ((_,j),recd) when Int.equal i j -> check_arity env cb; (true,recd) | CoFix (j,recd) when Int.equal i j -> check_arity env cb; (false,recd) | _ -> raise Impossible) - | Undef _ | OpaqueDef _ -> raise Impossible + | Undef _ | OpaqueDef _ | Primitive _ -> raise Impossible let prec_declaration_equal sg (na1, ca1, ta1) (na2, ca2, ta2) = Array.equal Name.equal na1 na2 && diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 67c605ea1d..c15486ea10 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -304,9 +304,9 @@ let rec extract_type env sg db j c args = | (Info, TypeScheme) -> let mlt = extract_type_app env sg db (r, type_sign env sg typ) args in (match (lookup_constant kn env).const_body with - | Undef _ | OpaqueDef _ -> mlt - | Def _ when is_custom (ConstRef kn) -> mlt - | Def lbody -> + | Undef _ | OpaqueDef _ | Primitive _ -> mlt + | Def _ when is_custom (ConstRef kn) -> mlt + | Def lbody -> let newc = applistc (get_body lbody) args in let mlt' = extract_type env sg db j newc [] in (* ML type abbreviations interact badly with Coq *) @@ -318,7 +318,7 @@ let rec extract_type env sg db j c args = | (Info, Default) -> (* Not an ML type, for example [(c:forall X, X->X) Type nat] *) (match (lookup_constant kn env).const_body with - | Undef _ | OpaqueDef _ -> Tunknown (* Brutal approx ... *) + | Undef _ | OpaqueDef _ | Primitive _ -> Tunknown (* Brutal approx ... *) | Def lbody -> (* We try to reduce. *) let newc = applistc (get_body lbody) args in @@ -346,7 +346,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 _ -> assert false + | Cast _ | LetIn _ | Construct _ | Int _ -> assert false (*s Auxiliary function dealing with type application. Precondition: [r] is a type scheme represented by the signature [s], @@ -564,7 +564,7 @@ and mlt_env env r = match r with | ConstRef kn -> let cb = Environ.lookup_constant kn env in match cb.const_body with - | Undef _ | OpaqueDef _ -> None + | Undef _ | OpaqueDef _ | Primitive _ -> None | Def l_body -> match lookup_typedef kn cb with | Some _ as o -> o @@ -683,6 +683,7 @@ let rec extract_term env sg mle mlt c args = let vty = extract_type env sg [] 0 ty [] in let extract_var mlt = put_magic (mlt,vty) (MLglob (VarRef v)) in extract_app env sg mle mlt extract_var args + | Int i -> assert (args = []); MLuint i | Ind _ | Prod _ | Sort _ -> assert false (*s [extract_maybe_term] is [extract_term] for usual terms, else [MLdummy] *) @@ -1063,7 +1064,7 @@ let extract_constant env kn cb = | (Logic,Default) -> warn_log (); Dterm (r, MLdummy Kprop, Tdummy Kprop) | (Info,TypeScheme) -> (match cb.const_body with - | Undef _ -> warn_info (); mk_typ_ax () + | Primitive _ | Undef _ -> warn_info (); mk_typ_ax () | Def c -> (match Recordops.find_primitive_projection kn with | None -> mk_typ (get_body c) @@ -1079,7 +1080,7 @@ let extract_constant env kn cb = else mk_typ_ax ()) | (Info,Default) -> (match cb.const_body with - | Undef _ -> warn_info (); mk_ax () + | Primitive _ | Undef _ -> warn_info (); mk_ax () | Def c -> (match Recordops.find_primitive_projection kn with | None -> mk_def (get_body c) @@ -1107,7 +1108,7 @@ let extract_constant_spec env kn cb = | (Info, TypeScheme) -> let s,vl = type_sign_vl env sg typ in (match cb.const_body with - | Undef _ | OpaqueDef _ -> Stype (r, vl, None) + | Undef _ | OpaqueDef _ | Primitive _ -> Stype (r, vl, None) | Def body -> let db = db_from_sign s in let body = get_body body in diff --git a/plugins/extraction/haskell.ml b/plugins/extraction/haskell.ml index 97fe9f24d5..a3cd92d556 100644 --- a/plugins/extraction/haskell.ml +++ b/plugins/extraction/haskell.ml @@ -214,6 +214,8 @@ let rec pp_expr par env args = | MLmagic a -> pp_apply (str "unsafeCoerce") par (pp_expr true env [] a :: args) | MLaxiom -> pp_par par (str "Prelude.error \"AXIOM TO BE REALIZED\"") + | MLuint _ -> + pp_par par (str "Prelude.error \"EXTRACTION OF UINT 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 e43c47d050..f88d29e9ed 100644 --- a/plugins/extraction/json.ml +++ b/plugins/extraction/json.ml @@ -155,6 +155,10 @@ let rec json_expr env = function ("value", json_expr env a) ] | MLaxiom -> json_dict [("what", json_str "expr:axiom")] + | MLuint i -> json_dict [ + ("what", json_str "expr:int"); + ("int", json_str (Uint63.to_string i)) + ] 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 ce920ad6a0..b7f80d543b 100644 --- a/plugins/extraction/miniml.ml +++ b/plugins/extraction/miniml.ml @@ -126,6 +126,7 @@ and ml_ast = | MLdummy of kill_reason | MLaxiom | MLmagic of ml_ast + | MLuint of Uint63.t and ml_pattern = | Pcons of GlobRef.t * ml_pattern list diff --git a/plugins/extraction/miniml.mli b/plugins/extraction/miniml.mli index ce920ad6a0..9df0f4964e 100644 --- a/plugins/extraction/miniml.mli +++ b/plugins/extraction/miniml.mli @@ -126,6 +126,7 @@ and ml_ast = | MLdummy of kill_reason | MLaxiom | MLmagic of ml_ast + | MLuint of Uint63.t and ml_pattern = | Pcons of GlobRef.t * ml_pattern list diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml index 9f5c1f1a17..2432887673 100644 --- a/plugins/extraction/mlutil.ml +++ b/plugins/extraction/mlutil.ml @@ -66,7 +66,8 @@ let rec eq_ml_type t1 t2 = match t1, t2 with | Tdummy k1, Tdummy k2 -> k1 == k2 | Tunknown, Tunknown -> true | Taxiom, Taxiom -> true -| _ -> false +| (Tarr _ | Tglob _ | Tvar _ | Tvar' _ | Tmeta _ | Tdummy _ | Tunknown | Taxiom), _ + -> false and eq_ml_meta m1 m2 = Int.equal m1.id m2.id && Option.equal eq_ml_type m1.contents m2.contents @@ -107,7 +108,7 @@ let rec type_occurs alpha t = | Tmeta {contents=Some u} -> type_occurs alpha u | Tarr (t1, t2) -> type_occurs alpha t1 || type_occurs alpha t2 | Tglob (r,l) -> List.exists (type_occurs alpha) l - | _ -> false + | (Tdummy _ | Tvar _ | Tvar' _ | Taxiom | Tunknown) -> false (*s Most General Unificator *) @@ -310,7 +311,7 @@ let isMLdummy = function MLdummy _ -> true | _ -> false let sign_of_id = function | Dummy -> Kill Kprop - | _ -> Keep + | (Id _ | Tmp _) -> Keep (* Classification of signatures *) @@ -370,7 +371,10 @@ let eq_ml_ident i1 i2 = match i1, i2 with | Dummy, Dummy -> true | Id id1, Id id2 -> Id.equal id1 id2 | Tmp id1, Tmp id2 -> Id.equal id1 id2 -| _ -> false +| Dummy, (Id _ | Tmp _) +| Id _, (Dummy | Tmp _) +| Tmp _, (Dummy | Id _) + -> false let rec eq_ml_ast t1 t2 = match t1, t2 with | MLrel i1, MLrel i2 -> @@ -394,7 +398,8 @@ let rec eq_ml_ast t1 t2 = match t1, t2 with | MLdummy k1, MLdummy k2 -> k1 == k2 | MLaxiom, MLaxiom -> true | MLmagic t1, MLmagic t2 -> eq_ml_ast t1 t2 -| _ -> false +| MLuint i1, MLuint i2 -> Uint63.equal i1 i2 +| _, _ -> false and eq_ml_pattern p1 p2 = match p1, p2 with | Pcons (gr1, p1), Pcons (gr2, p2) -> @@ -426,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 -> () + | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> () in iter 0 (*s Map over asts. *) @@ -445,7 +450,7 @@ 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 as a -> a + | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ as a -> a (*s Map over asts, with binding depth as parameter. *) @@ -463,7 +468,7 @@ 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 as a -> a + | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ as a -> a (*s Iter over asts. *) @@ -477,7 +482,7 @@ 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 -> () + | MLrel _ | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> () (*S Operations concerning De Bruijn indices. *) @@ -513,7 +518,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 -> 0 + | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> 0 in nb 1 (* Replace unused variables by _ *) @@ -565,7 +570,7 @@ let dump_unused_vars a = let b' = ren env b in if b' == b then a else MLmagic b' - | MLglob _ | MLexn _ | MLdummy _ | MLaxiom -> a + | MLglob _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> a and ren_branch env ((ids,p,b) as tr) = let occs = List.map (fun _ -> ref false) ids in @@ -1398,7 +1403,7 @@ 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 -> 0 + | MLglob _ | MLrel _ | MLexn _ | MLdummy _ | MLaxiom | MLuint _ -> 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 b398bc07a0..654695c232 100644 --- a/plugins/extraction/modutil.ml +++ b/plugins/extraction/modutil.ml @@ -108,7 +108,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 _ -> () + | MLdummy _ | MLaxiom | MLmagic _ | MLuint _ -> () 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 96d8760404..8940aedd6d 100644 --- a/plugins/extraction/ocaml.ml +++ b/plugins/extraction/ocaml.ml @@ -310,6 +310,10 @@ let rec pp_expr par env args = apply2 (v 0 (str "match " ++ head ++ str " with" ++ fnl () ++ pp_pat env pv)))) + | MLuint i -> + assert (args=[]); + str "(" ++ str (Uint63.compile i) ++ str ")" + and pp_record_proj par env typ t pv args = (* Can a match be printed as a mere record projection ? *) diff --git a/plugins/extraction/scheme.ml b/plugins/extraction/scheme.ml index 76a0c74068..6aa3a6220e 100644 --- a/plugins/extraction/scheme.ml +++ b/plugins/extraction/scheme.ml @@ -129,6 +129,8 @@ let rec pp_expr env args = | MLmagic a -> pp_expr env args a | MLaxiom -> paren (str "error \"AXIOM TO BE REALIZED\"") + | MLuint _ -> + paren (str "Prelude.error \"EXTRACTION OF UINT 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 3b95423067..8da30bd9c9 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -763,12 +763,13 @@ let build_proof end | Cast(t,_,_) -> build_proof do_finalize {dyn_infos with info = t} g - | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ -> + | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ -> 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") | 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/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 4b6caea70d..02964d7ba5 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 funnames avoid rt : glob_constr build_entry_return = 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 _ -> + | GRef _ | GVar _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ -> (* do nothing (except changing type of course) *) mk_result [] rt avoid | GApp(_,_) -> @@ -588,6 +588,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return = build_entry_lc env funnames avoid (mkGApp(b,args)) | 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") end (* end of the application treatement *) | GLambda(n,_,t,b) -> @@ -1221,7 +1222,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 _ -> params + | GRef _ | GVar _ | GEvar _ | GPatVar _ | GInt _ -> 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 5b45a8dbed..13ff19a46b 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -106,6 +106,7 @@ let change_vars = | GRec _ -> user_err ?loc Pp.(str "Local (co)fixes are not supported") | GSort _ as x -> x | GHole _ as x -> x + | GInt _ as x -> x | GCast(b,c) -> GCast(change_vars mapping b, Glob_ops.map_cast_type (change_vars mapping) c) @@ -285,6 +286,7 @@ let rec alpha_rt excluded rt = ) | GRec _ -> user_err Pp.(str "Not handled GRec") | GSort _ + | GInt _ | GHole _ as rt -> rt | GCast (b,c) -> GCast(alpha_rt excluded b, @@ -344,6 +346,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 ) x and is_free_in_br {CAst.v=(ids,_,rt)} = (not (Id.List.mem id ids)) && is_free_in rt @@ -434,6 +437,7 @@ let replace_var_by_term x_id term = | GRec _ -> raise (UserError(None,str "Not handled GRec")) | GSort _ | GHole _ as rt -> rt + | GInt _ as rt -> rt | GCast(b,c) -> GCast(replace_var_by_pattern b, Glob_ops.map_cast_type replace_var_by_pattern c) @@ -516,7 +520,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 _ as rt -> rt + | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ as rt -> rt | GVar id as rt -> begin try diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 3a04c753ea..d9b0330e2b 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -201,7 +201,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 _ -> false + | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ -> false | GCast(b,_) -> lookup names b | GRec _ -> error "GRec not handled" | GIf(b,_,lhs,rhs) -> diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 38f27f760b..1b5286dce4 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -306,6 +306,7 @@ let check_not_nested sigma forbidden e = let rec check_not_nested e = match EConstr.kind sigma e with | Rel _ -> () + | Int _ -> () | Var x -> if Id.List.mem x forbidden then user_err ~hdr:"Recdef.check_not_nested" @@ -487,7 +488,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g = | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_leconstr_env (pf_env g) 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 _ -> + | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ | Int _ -> let new_continuation_tac = jinfo.otherS () expr_info continuation_tac in new_continuation_tac expr_info g @@ -1294,6 +1295,7 @@ let is_opaque_constant c = | Declarations.OpaqueDef _ -> Proof_global.Opaque | Declarations.Undef _ -> Proof_global.Opaque | Declarations.Def _ -> Proof_global.Transparent + | Declarations.Primitive _ -> Proof_global.Opaque let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decompose_and_tac,nb_goal) = (* Pp.msgnl (str "gls_type := " ++ Printer.pr_lconstr gls_type); *) diff --git a/plugins/ssrmatching/ssrmatching.ml b/plugins/ssrmatching/ssrmatching.ml index efd65ade15..552a4048b1 100644 --- a/plugins/ssrmatching/ssrmatching.ml +++ b/plugins/ssrmatching/ssrmatching.ml @@ -312,7 +312,8 @@ let iter_constr_LR f c = match kind c with | Fix (_, (_, t, b)) | CoFix (_, (_, t, b)) -> 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 _) -> () + | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ + | Int _) -> () (* 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/int31_syntax.ml b/plugins/syntax/int31_syntax.ml deleted file mode 100644 index e34a401c2c..0000000000 --- a/plugins/syntax/int31_syntax.ml +++ /dev/null @@ -1,114 +0,0 @@ -(************************************************************************) -(* * 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) *) -(************************************************************************) - - -(* Poor's man DECLARE PLUGIN *) -let __coq_plugin_name = "int31_syntax_plugin" -let () = Mltop.add_known_module __coq_plugin_name - -(* digit-based syntax for int31 *) - -open Bigint -open Names -open Globnames -open Glob_term - -(*** Constants for locating int31 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) - -let is_gr c gr = match DAst.get c with -| GRef (r, _) -> GlobRef.equal r gr -| _ -> false - -let make_mind mp id = Names.MutInd.make2 mp (Label.make id) -let make_mind_mpfile dir id = make_mind (ModPath.MPfile (make_dir dir)) id -let make_mind_mpdot dir modname id = - let mp = ModPath.MPdot (ModPath.MPfile (make_dir dir), Label.make modname) - in make_mind mp id - - -(* int31 stuff *) -let int31_module = ["Coq"; "Numbers"; "Cyclic"; "Int31"; "Int31"] -let int31_path = make_path int31_module "int31" -let int31_id = make_mind_mpfile int31_module -let int31_scope = "int31_scope" - -let int31_construct = ConstructRef ((int31_id "int31",0),1) - -let int31_0 = ConstructRef ((int31_id "digits",0),1) -let int31_1 = ConstructRef ((int31_id "digits",0),2) - -(*** Definition of the Non_closed exception, used in the pretty printing ***) -exception Non_closed - -(*** Parsing for int31 in digital notation ***) - -(* parses a *non-negative* integer (from bigint.ml) into an int31 - wraps modulo 2^31 *) -let int31_of_pos_bigint ?loc n = - let ref_construct = DAst.make ?loc (GRef (int31_construct, None)) in - let ref_0 = DAst.make ?loc (GRef (int31_0, None)) in - let ref_1 = DAst.make ?loc (GRef (int31_1, None)) in - let rec args counter n = - if counter <= 0 then - [] - else - let (q,r) = div2_with_rest n in - (if r then ref_1 else ref_0)::(args (counter-1) q) - in - DAst.make ?loc (GApp (ref_construct, List.rev (args 31 n))) - -let error_negative ?loc = - CErrors.user_err ?loc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.") - -let interp_int31 ?loc n = - if is_pos_or_zero n then - int31_of_pos_bigint ?loc n - else - error_negative ?loc - -(* Pretty prints an int31 *) - -let bigint_of_int31 = - let rec args_parsing args cur = - match args with - | [] -> cur - | r::l when is_gr r int31_0 -> args_parsing l (mult_2 cur) - | r::l when is_gr r int31_1 -> args_parsing l (add_1 (mult_2 cur)) - | _ -> raise Non_closed - in - fun c -> match DAst.get c with - | GApp (r, args) when is_gr r int31_construct -> args_parsing args zero - | _ -> raise Non_closed - -let uninterp_int31 (AnyGlobConstr i) = - try - Some (bigint_of_int31 i) - with Non_closed -> - None - -open Notation - -let at_declare_ml_module f x = - Mltop.declare_cache_obj (fun () -> f x) __coq_plugin_name - -(* Actually declares the interpreter for int31 *) - -let _ = - register_bignumeral_interpretation int31_scope (interp_int31,uninterp_int31); - at_declare_ml_module enable_prim_token_interpretation - { pt_local = false; - pt_scope = int31_scope; - pt_interp_info = Uid int31_scope; - pt_required = (int31_path,int31_module); - pt_refs = [int31_construct]; - pt_in_match = true } diff --git a/plugins/syntax/int31_syntax_plugin.mlpack b/plugins/syntax/int31_syntax_plugin.mlpack deleted file mode 100644 index 54a5bc0cd1..0000000000 --- a/plugins/syntax/int31_syntax_plugin.mlpack +++ /dev/null @@ -1 +0,0 @@ -Int31_syntax diff --git a/plugins/syntax/int63_syntax.ml b/plugins/syntax/int63_syntax.ml new file mode 100644 index 0000000000..992b7a986b --- /dev/null +++ b/plugins/syntax/int63_syntax.ml @@ -0,0 +1,55 @@ +(************************************************************************) +(* 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 *) +(************************************************************************) + + +(* Poor's man DECLARE PLUGIN *) +let __coq_plugin_name = "int63_syntax_plugin" +let () = Mltop.add_known_module __coq_plugin_name + +(* digit-based syntax for int63 *) + +open Names +open Libnames + +(*** Constants for locating int63 constructors ***) + +let q_int63 = qualid_of_string "Coq.Numbers.Cyclic.Int63.Int63.int" +let q_id_int63 = qualid_of_string "Coq.Numbers.Cyclic.Int63.Int63.id_int" + +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) + +(* int63 stuff *) +let int63_module = ["Coq"; "Numbers"; "Cyclic"; "Int63"; "Int63"] +let int63_path = make_path int63_module "int" +let int63_scope = "int63_scope" + +let at_declare_ml_module f x = + Mltop.declare_cache_obj (fun () -> f x) __coq_plugin_name + +(* Actually declares the interpreter for int63 *) + +let _ = + let open Notation in + at_declare_ml_module + (fun () -> + let id_int63 = Nametab.locate q_id_int63 in + let o = { to_kind = Int63, Direct; + to_ty = id_int63; + of_kind = Int63, Direct; + of_ty = id_int63; + ty_name = q_int63; + warning = Nop } in + enable_prim_token_interpretation + { pt_local = false; + pt_scope = int63_scope; + pt_interp_info = NumeralNotation o; + pt_required = (int63_path, int63_module); + pt_refs = []; + pt_in_match = false }) + () diff --git a/plugins/syntax/int63_syntax_plugin.mlpack b/plugins/syntax/int63_syntax_plugin.mlpack new file mode 100644 index 0000000000..d83d554fe6 --- /dev/null +++ b/plugins/syntax/int63_syntax_plugin.mlpack @@ -0,0 +1 @@ +Int63_syntax diff --git a/plugins/syntax/numeral.ml b/plugins/syntax/numeral.ml index ea564ae2ba..0c6d2ac0d1 100644 --- a/plugins/syntax/numeral.ml +++ b/plugins/syntax/numeral.ml @@ -69,6 +69,14 @@ let locate_int () = }, mkRefC q_int, mkRefC q_uint) else None +let locate_int63 () = + let int63n = "num.int63.type" in + if Coqlib.has_ref int63n + then + let q_int63 = qualid_of_ref int63n in + Some (mkRefC q_int63) + else None + let has_type f ty = let (sigma, env) = Pfedit.get_current_context () in let c = mkCastC (mkRefC f, Glob_term.CastConv ty) in @@ -79,17 +87,18 @@ let type_error_to f ty = CErrors.user_err (pr_qualid f ++ str " should go from Decimal.int to " ++ pr_qualid ty ++ str " or (option " ++ pr_qualid ty ++ str ")." ++ - fnl () ++ str "Instead of Decimal.int, the types Decimal.uint or Z could be used (you may need to require BinNums or Decimal first).") + fnl () ++ str "Instead of Decimal.int, the types Decimal.uint or Z or Int63.int could be used (you may need to require BinNums or Decimal or Int63 first).") let type_error_of g ty = CErrors.user_err (pr_qualid g ++ str " should go from " ++ pr_qualid ty ++ str " to Decimal.int or (option Decimal.int)." ++ fnl () ++ - str "Instead of Decimal.int, the types Decimal.uint or Z could be used (you may need to require BinNums or Decimal first).") + str "Instead of Decimal.int, the types Decimal.uint or Z or Int63.int could be used (you may need to require BinNums or Decimal or Int63 first).") let vernac_numeral_notation local ty f g scope opts = let int_ty = locate_int () in let z_pos_ty = locate_z () in + let int63_ty = locate_int63 () in let tyc = Smartlocate.global_inductive_with_alias ty in let to_ty = Smartlocate.global_with_alias f in let of_ty = Smartlocate.global_with_alias g in @@ -111,6 +120,10 @@ let vernac_numeral_notation local ty f g scope opts = match z_pos_ty with | Some (z_pos_ty, cZ) when has_type f (arrow cZ cty) -> Z z_pos_ty, Direct | Some (z_pos_ty, cZ) when has_type f (arrow cZ (opt cty)) -> Z z_pos_ty, Option + | _ -> + match int63_ty with + | Some cint63 when has_type f (arrow cint63 cty) -> Int63, Direct + | Some cint63 when has_type f (arrow cint63 (opt cty)) -> Int63, Option | _ -> type_error_to f ty in (* Check the type of g *) @@ -124,6 +137,10 @@ let vernac_numeral_notation local ty f g scope opts = match z_pos_ty with | Some (z_pos_ty, cZ) when has_type g (arrow cty cZ) -> Z z_pos_ty, Direct | Some (z_pos_ty, cZ) when has_type g (arrow cty (opt cZ)) -> Z z_pos_ty, Option + | _ -> + match int63_ty with + | Some cint63 when has_type g (arrow cty cint63) -> Int63, Direct + | Some cint63 when has_type g (arrow cty (opt cint63)) -> Int63, Option | _ -> type_error_of g ty in let o = { to_kind; to_ty; of_kind; of_ty; diff --git a/plugins/syntax/plugin_base.dune b/plugins/syntax/plugin_base.dune index 1ab16c700d..aac46338ea 100644 --- a/plugins/syntax/plugin_base.dune +++ b/plugins/syntax/plugin_base.dune @@ -20,8 +20,8 @@ (libraries coq.vernac)) (library - (name int31_syntax_plugin) - (public_name coq.plugins.int31_syntax) - (synopsis "Coq syntax plugin: int31") - (modules int31_syntax) + (name int63_syntax_plugin) + (public_name coq.plugins.int63_syntax) + (synopsis "Coq syntax plugin: int63") + (modules int63_syntax) (libraries coq.vernac)) diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml index f8289f558c..e27fc536eb 100644 --- a/pretyping/cbv.ml +++ b/pretyping/cbv.ml @@ -38,6 +38,10 @@ open Esubst * the bindings S, and then applied to args. Here again, * weak reduction. * CONSTR(c,args) is the constructor [c] applied to [args]. + * PRIMITIVE(cop,args) represent a particial application of + * a primitive, or a fully applied primitive + * which does not reduce. + * cop is the constr representing op. * *) type cbv_value = @@ -48,6 +52,7 @@ type cbv_value = | FIXP of fixpoint * cbv_value subs * cbv_value array | COFIXP of cofixpoint * cbv_value subs * cbv_value array | CONSTR of constructor Univ.puniverses * cbv_value array + | PRIMITIVE of CPrimitives.t * constr * cbv_value array (* type of terms with a hole. This hole can appear only under App or Case. * TOP means the term is considered without context @@ -90,6 +95,9 @@ let rec shift_value n = function COFIXP (cofix,subs_shft (n,s), Array.map (shift_value n) args) | CONSTR (c,args) -> CONSTR (c, Array.map (shift_value n) args) + | PRIMITIVE(op,c,args) -> + PRIMITIVE(op,c,Array.map (shift_value n) args) + let shift_value n v = if Int.equal n 0 then v else shift_value n v @@ -109,10 +117,11 @@ let contract_cofixp env (i,(_,_,bds as bodies)) = let n = Array.length bds in subs_cons(Array.init n make_body, env), bds.(i) -let make_constr_ref n = function +let make_constr_ref n k t = + match k with | RelKey p -> mkRel (n+p) - | VarKey id -> mkVar id - | ConstKey cst -> mkConstU cst + | VarKey id -> t + | ConstKey cst -> t (* Adds an application list. Collapse APPs! *) let stack_app appl stack = @@ -136,7 +145,7 @@ let mkSTACK = function type cbv_infos = { env : Environ.env; - tab : cbv_value KeyTable.t; + tab : cbv_value Declarations.constant_def KeyTable.t; reds : RedFlags.reds; sigma : Evd.evar_map } @@ -159,7 +168,8 @@ let strip_appl head stack = | FIXP (fix,env,app) -> (FIXP(fix,env,[||]), stack_app app stack) | COFIXP (cofix,env,app) -> (COFIXP(cofix,env,[||]), stack_app app stack) | CONSTR (c,app) -> (CONSTR(c,[||]), stack_app app stack) - | _ -> (head, stack) + | PRIMITIVE(op,c,app) -> (PRIMITIVE(op,c,[||]), stack_app app stack) + | VAL _ | STACK _ | CBN _ | LAM _ -> (head, stack) (* Tests if fixpoint reduction is possible. *) @@ -189,6 +199,60 @@ let get_debug_cbv = Goptions.declare_bool_option_and_ref ~name:"cbv visited constants display" ~key:["Debug";"Cbv"] +(* Reduction of primitives *) + +open Primred + +module VNativeEntries = + struct + + type elem = cbv_value + type args = cbv_value array + type evd = unit + + let get = Array.get + + let get_int () e = + match e with + | VAL(_, ci) -> + (match kind ci with + | Int i -> i + | _ -> raise Primred.NativeDestKO) + | _ -> raise Primred.NativeDestKO + + let mkInt env i = VAL(0, mkInt i) + + 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 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|]) + + let mkIntPair env e1 e2 = + let int_ty = int_ty env in + let c = get_pair_constructor env in + CONSTR(Univ.in_punivs c, [|int_ty;int_ty;e1;e2|]) + + let mkLt env = + let (_eq,lt,_gt) = get_cmp_constructors env in + CONSTR(Univ.in_punivs lt, [||]) + + let mkEq env = + let (eq,_lt,_gt) = get_cmp_constructors env in + CONSTR(Univ.in_punivs eq, [||]) + + let mkGt env = + let (_eq,_lt,gt) = get_cmp_constructors env in + CONSTR(Univ.in_punivs gt, [||]) + + end + +module VredNative = RedNative(VNativeEntries) + let debug_pr_key = function | ConstKey (sp,_) -> Names.Constant.print sp | VarKey id -> Names.Id.print id @@ -225,6 +289,8 @@ and reify_value = function (* reduction under binders *) mkApp (apply_env env cofix, Array.map reify_value args) | CONSTR (c,args) -> mkApp(mkConstructU c, Array.map reify_value args) + | PRIMITIVE(op,c,args) -> + mkApp(c, Array.map reify_value args) and apply_env env t = match kind t with @@ -277,14 +343,14 @@ let rec norm_head info env t stack = | Inl (0,v) -> strip_appl v stack | Inl (n,v) -> strip_appl (shift_value n v) stack | Inr (n,None) -> (VAL(0, mkRel n), stack) - | Inr (n,Some p) -> norm_head_ref (n-p) info env stack (RelKey p)) + | Inr (n,Some p) -> norm_head_ref (n-p) info env stack (RelKey p) t) - | Var id -> norm_head_ref 0 info env stack (VarKey id) + | Var id -> norm_head_ref 0 info env stack (VarKey id) t | Const sp -> Reductionops.reduction_effect_hook info.env info.sigma (fst sp) (lazy (reify_stack t stack)); - norm_head_ref 0 info env stack (ConstKey sp) + norm_head_ref 0 info env stack (ConstKey sp) t | LetIn (_, b, _, c) -> (* zeta means letin are contracted; delta without zeta means we *) @@ -315,22 +381,23 @@ let rec norm_head info env t stack = | Construct c -> (CONSTR(c, [||]), stack) (* neutral cases *) - | (Sort _ | Meta _ | Ind _) -> (VAL(0, t), stack) + | (Sort _ | Meta _ | Ind _ | Int _) -> (VAL(0, t), stack) | Prod _ -> (CBN(t,env), stack) -and norm_head_ref k info env stack normt = +and norm_head_ref k info env stack normt t = if red_set_ref info.reds normt then match cbv_value_cache info normt with - | Some body -> + | Declarations.Def body -> if get_debug_cbv () then Feedback.msg_debug Pp.(str "Unfolding " ++ debug_pr_key normt); strip_appl (shift_value k body) stack - | None -> + | Declarations.Primitive op -> (PRIMITIVE(op,t,[||]),stack) + | Declarations.OpaqueDef _ | Declarations.Undef _ -> if get_debug_cbv () then Feedback.msg_debug Pp.(str "Not unfolding " ++ debug_pr_key normt); - (VAL(0,make_constr_ref k normt),stack) + (VAL(0,make_constr_ref k normt t),stack) else begin if get_debug_cbv () then Feedback.msg_debug Pp.(str "Not unfolding " ++ debug_pr_key normt); - (VAL(0,make_constr_ref k normt),stack) + (VAL(0,make_constr_ref k normt t),stack) end (* cbv_stack_term performs weak reduction on constr t under the subs @@ -392,32 +459,58 @@ and cbv_stack_value info env = function | (COFIXP(cofix,env,[||]), APP(appl,TOP)) -> COFIXP(cofix,env,appl) | (CONSTR(c,[||]), APP(appl,TOP)) -> CONSTR(c,appl) + (* primitive apply to arguments *) + | (PRIMITIVE(op,c,[||]), APP(appl,stk)) -> + let nargs = CPrimitives.arity op in + let len = Array.length appl in + if nargs <= len then + let args = + if len = nargs then appl + else Array.sub appl 0 nargs in + let stk = + if nargs < len then + stack_app (Array.sub appl nargs (len - nargs)) stk + else stk in + match VredNative.red_prim info.env () op args with + | Some (CONSTR (c, args)) -> + (* args must be moved to the stack to allow future reductions *) + cbv_stack_value info env (CONSTR(c, [||]), stack_app args stk) + | Some v -> cbv_stack_value info env (v,stk) + | None -> mkSTACK(PRIMITIVE(op,c,args), stk) + else (* partical application *) + (assert (stk = TOP); + PRIMITIVE(op,c,appl)) + (* definitely a value *) | (head,stk) -> mkSTACK(head, stk) -and cbv_value_cache info ref = match KeyTable.find info.tab ref with -| v -> Some v -| exception Not_found -> - try - let body = match ref with - | RelKey n -> - let open Context.Rel.Declaration in - begin match Environ.lookup_rel n info.env with - | LocalDef (_, c, _) -> lift n c - | LocalAssum _ -> raise Not_found - end - | VarKey id -> - let open Context.Named.Declaration in - begin match Environ.lookup_named id info.env with - | LocalDef (_, c, _) -> c - | LocalAssum _ -> raise Not_found - end - | ConstKey cst -> Environ.constant_value_in info.env cst +and cbv_value_cache info ref = + try KeyTable.find info.tab ref with + Not_found -> + let v = + try + let body = match ref with + | RelKey n -> + let open Context.Rel.Declaration in + begin match Environ.lookup_rel n info.env with + | LocalDef (_, c, _) -> lift n c + | LocalAssum _ -> raise Not_found + end + | VarKey id -> + let open Context.Named.Declaration in + begin match Environ.lookup_named id info.env with + | LocalDef (_, c, _) -> c + | LocalAssum _ -> raise Not_found + end + | ConstKey cst -> Environ.constant_value_in info.env cst + in + let v = cbv_stack_term info TOP (subs_id 0) body in + Declarations.Def v + with + | Environ.NotEvaluableConst (Environ.IsPrimitive op) -> Declarations.Primitive op + | Not_found | Environ.NotEvaluableConst _ -> Declarations.Undef None in - let v = cbv_stack_term info TOP (subs_id 0) body in - let () = KeyTable.add info.tab ref v in - Some v - with Not_found | Environ.NotEvaluableConst _ -> None + KeyTable.add info.tab ref v; v (* When we are sure t will never produce a redex with its stack, we * normalize (even under binders) the applied terms and we build the @@ -471,6 +564,8 @@ and cbv_norm_value info = function (* reduction under binders *) Array.map (cbv_norm_value info) args) | CONSTR (c,args) -> mkApp(mkConstructU c, Array.map (cbv_norm_value info) args) + | PRIMITIVE(op,c,args) -> + mkApp(c,Array.map (cbv_norm_value info) args) (* with profiling *) let cbv_norm infos constr = diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli index 83844c95a7..0a1e771921 100644 --- a/pretyping/cbv.mli +++ b/pretyping/cbv.mli @@ -36,6 +36,7 @@ type cbv_value = | FIXP of fixpoint * cbv_value subs * cbv_value array | COFIXP of cofixpoint * cbv_value subs * cbv_value array | CONSTR of constructor Univ.puniverses * cbv_value array + | PRIMITIVE of CPrimitives.t * Constr.t * cbv_value array and cbv_stack = | TOP diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml index 032e4bbf85..94257fedd7 100644 --- a/pretyping/constr_matching.ml +++ b/pretyping/constr_matching.ml @@ -410,9 +410,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 | (PRef _ | PVar _ | PRel _ | PApp _ | PProj _ | PLambda _ | PProd _ | PLetIn _ | PSort _ | PIf _ | PCase _ - | PFix _ | PCoFix _| PEvar _), _ -> raise PatternMatchingFailure + | PFix _ | PCoFix _| PEvar _ | PInt _), _ -> raise PatternMatchingFailure in sorec [] env ((Id.Map.empty,Id.Set.empty), Id.Map.empty) pat c @@ -530,7 +531,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 _ -> + | Construct _| Ind _|Evar _|Const _ | Rel _|Meta _|Var _|Sort _ | Int _ -> next () in here next diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index c7cc2dbc8a..6746e4b902 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -766,6 +766,7 @@ and detype_r d flags avoid env sigma t = p c bl | 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 and detype_eqns d flags avoid env sigma ci computable constructs consnargsl bl = try @@ -959,6 +960,7 @@ let rec subst_glob_constr subst = DAst.map (function | GSort _ | GVar _ | GEvar _ + | GInt _ | GPatVar _ as raw -> raw | GApp (r,rl) as raw -> diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index ed28cc7725..aa30ed8d2e 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -107,7 +107,7 @@ let flex_kind_of_term ts env evd c sk = Option.cata (fun x -> MaybeFlexible x) Rigid (eval_flexible_term ts env evd c) | Lambda _ when not (Option.is_empty (Stack.decomp sk)) -> MaybeFlexible c | Evar ev -> Flexible ev - | Lambda _ | Prod _ | Sort _ | Ind _ | Construct _ | CoFix _ -> Rigid + | Lambda _ | Prod _ | Sort _ | Ind _ | Construct _ | CoFix _ | Int _ -> Rigid | Meta _ -> Rigid | Fix _ -> Rigid (* happens when the fixpoint is partially applied *) | Cast _ | App _ | Case _ -> assert false @@ -127,7 +127,7 @@ let occur_rigidly (evk,_ as ev) evd t = let rec aux t = match EConstr.kind evd t with | App (f, c) -> if aux f then Array.exists aux c else false - | Construct _ | Ind _ | Sort _ | Meta _ | Fix _ | CoFix _ -> true + | Construct _ | Ind _ | Sort _ | Meta _ | Fix _ | CoFix _ | Int _ -> true | Proj (p, c) -> not (aux c) | Evar (evk',_) -> if Evar.equal evk evk' then raise Occur else false | Cast (p, _, _) -> aux p @@ -769,7 +769,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts 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 _) -> + | (Var _|Construct _|Ind _|Const _|Prod _|Sort _|Int _) -> Stack.not_purely_applicative args | (CoFix _|Meta _|Rel _)-> true | Evar _ -> Stack.not_purely_applicative args @@ -887,8 +887,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty | Const _, Const _ | Ind _, Ind _ - | Construct _, Construct _ -> - rigids env evd sk1 term1 sk2 term2 + | Construct _, Construct _ + | Int _, Int _ -> + rigids env evd sk1 term1 sk2 term2 | Construct u, _ -> eta_constructor ts env evd sk1 u sk2 term2 @@ -923,9 +924,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty |Some (sk1',sk2'), Success i' -> evar_conv_x ts env i' CONV (Stack.zip i' (term1,sk1')) (Stack.zip i' (term2,sk2')) end - | (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _), _ -> - UnifFailure (evd,NotSameHead) - | _, (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _) -> + | (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _ | Int _), _ -> UnifFailure (evd,NotSameHead) | (App _ | Cast _ | Case _ | Proj _), _ -> assert false diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml index e14766f370..6b61b1452e 100644 --- a/pretyping/glob_ops.ml +++ b/pretyping/glob_ops.ml @@ -152,8 +152,10 @@ let mk_glob_constr_eq f c1 c2 = match DAst.get c1, DAst.get c2 with Namegen.intro_pattern_naming_eq nam1 nam2 | GCast (c1, t1), GCast (c2, t2) -> f c1 c2 && cast_type_eq f t1 t2 + | GInt i1, GInt i2 -> Uint63.equal i1 i2 | (GRef _ | GVar _ | GEvar _ | GPatVar _ | GApp _ | GLambda _ | GProd _ | GLetIn _ | - GCases _ | GLetTuple _ | GIf _ | GRec _ | GSort _ | GHole _ | GCast _ ), _ -> false + GCases _ | GLetTuple _ | GIf _ | GRec _ | GSort _ | GHole _ | GCast _ | + GInt _), _ -> false let rec glob_constr_eq c = mk_glob_constr_eq glob_constr_eq c @@ -214,7 +216,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 _) as x -> x + | (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GInt _) as x -> x ) let map_glob_constr = map_glob_constr_left_to_right @@ -246,9 +248,8 @@ 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 _) -> acc + | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GInt _) -> 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 @@ -289,7 +290,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 _) -> acc)) + | (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GInt _) -> acc)) let iter_glob_constr f = fold_glob_constr (fun () -> f) () diff --git a/pretyping/glob_term.ml b/pretyping/glob_term.ml index c405fcfc72..8670c1d964 100644 --- a/pretyping/glob_term.ml +++ b/pretyping/glob_term.ml @@ -82,6 +82,7 @@ type 'a glob_constr_r = | GSort of glob_sort | 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 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 ccbb2934bc..cdeec875a2 100644 --- a/pretyping/heads.ml +++ b/pretyping/heads.ml @@ -91,6 +91,7 @@ let kind_of_head env t = | Proj (p,c) -> RigidHead RigidOther | Case (_,_,c,_) -> aux k [] c true + | Int _ -> ConstructorHead | Fix ((i,j),_) -> let n = i.(j) in try aux k [] (List.nth l n) true diff --git a/pretyping/inferCumulativity.ml b/pretyping/inferCumulativity.ml index e46d03b743..b5a6ba6be5 100644 --- a/pretyping/inferCumulativity.ml +++ b/pretyping/inferCumulativity.ml @@ -100,6 +100,7 @@ let rec infer_fterm cv_pb infos variances hd stk = let variances = infer_stack infos variances stk in infer_vect infos variances (Array.map (mk_clos e) args) | FRel _ -> infer_stack infos variances stk + | FInt _ -> infer_stack infos variances stk | FFlex fl -> let variances = infer_table_key infos variances fl in infer_stack infos variances stk @@ -155,6 +156,10 @@ and infer_stack infos variances (stk:CClosure.stack) = infer_vect infos variances (Array.map (mk_clos e) br) | Zshift _ -> variances | Zupdate _ -> variances + | Zprimitive (_,_,rargs,kargs) -> + let variances = List.fold_left (fun variances c -> infer_fterm CONV infos variances c []) variances rargs in + let variances = List.fold_left (fun variances (_,c) -> infer_fterm CONV infos variances c []) variances kargs in + variances in infer_stack infos variances stk diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml index dc2663c1ca..b7090e69da 100644 --- a/pretyping/nativenorm.ml +++ b/pretyping/nativenorm.ml @@ -120,13 +120,6 @@ let construct_of_constr_notnative const env tag (mind, _ as ind) u allargs = let mib,mip = lookup_mind_specif env ind in let nparams = mib.mind_nparams in let params = Array.sub allargs 0 nparams in - try - if const then - let ctyp = type_constructor mind mib u (mip.mind_nf_lc.(0)) params in - Retroknowledge.get_vm_decompile_constant_info env.retroknowledge (GlobRef.IndRef ind) tag, ctyp - else - raise Not_found - with Not_found -> let i = invert_tag const tag mip.mind_reloc_tbl in let ctyp = type_constructor mind mib u (mip.mind_nf_lc.(i-1)) params in (mkApp(mkConstructU((ind,i),u), params), ctyp) @@ -137,7 +130,9 @@ let construct_of_constr const env sigma tag typ = match EConstr.kind_upto sigma t with | Ind (ind,u) -> construct_of_constr_notnative const env tag ind u l - | _ -> assert false + | _ -> + assert (Constr.equal t (Typeops.type_of_int env)); + (mkInt (Uint63.of_int tag), t) let construct_of_constr_const env sigma tag typ = fst (construct_of_constr true env sigma tag typ) @@ -208,6 +203,7 @@ let rec nf_val env sigma v typ = let body = nf_val env sigma (f (mk_rel_accu lvl)) codom in mkLambda(name,dom,body) | Vconst n -> construct_of_constr_const env sigma n typ + | Vint64 i -> i |> Uint63.of_int64 |> mkInt | 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 be7ebe49cf..2ca7f21e8d 100644 --- a/pretyping/pattern.ml +++ b/pretyping/pattern.ml @@ -39,6 +39,7 @@ type constr_pattern = (int * bool list * constr_pattern) list (** index of constructor, nb of args *) | 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 (** 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 248d5d0a0e..6803ea7d9b 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -61,9 +61,11 @@ let rec constr_pattern_eq p1 p2 = match p1, p2 with Int.equal i1 i2 && rec_declaration_eq f1 f2 | PProj (p1, t1), PProj (p2, t2) -> Projection.equal p1 p2 && constr_pattern_eq t1 t2 +| PInt i1, PInt i2 -> + Uint63.equal i1 i2 | (PRef _ | PVar _ | PEvar _ | PRel _ | PApp _ | PSoApp _ | PLambda _ | PProd _ | PLetIn _ | PSort _ | PMeta _ - | PIf _ | PCase _ | PFix _ | PCoFix _ | PProj _), _ -> false + | PIf _ | PCase _ | PFix _ | PCoFix _ | PProj _ | PInt _), _ -> false (** FIXME: fixpoint and cofixpoint should be relativized to pattern *) and pattern_eq (i1, j1, p1) (i2, j2, p2) = @@ -90,7 +92,8 @@ let rec occur_meta_pattern = function (occur_meta_pattern c) || (List.exists (fun (_,_,p) -> occur_meta_pattern p) br) | PMeta _ | PSoApp _ -> true - | PEvar _ | PVar _ | PRef _ | PRel _ | PSort _ | PFix _ | PCoFix _ -> false + | PEvar _ | PVar _ | PRef _ | PRel _ | PSort _ | PFix _ | PCoFix _ + | PInt _ -> false let rec occurn_pattern n = function | PRel p -> Int.equal n p @@ -111,7 +114,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 _ -> false + | PVar _ | PRef _ | PSort _ | PInt _ -> false | PFix (_,(_,tl,bl)) -> Array.exists (occurn_pattern n) tl || Array.exists (occurn_pattern (n+Array.length tl)) bl | PCoFix (_,(_,tl,bl)) -> @@ -134,7 +137,7 @@ let rec head_pattern_bound t = -> raise BoundPattern (* Perhaps they were arguments, but we don't beta-reduce *) | PLambda _ -> raise BoundPattern - | PCoFix _ -> anomaly ~label:"head_pattern_bound" (Pp.str "not a type.") + | PCoFix _ | PInt _ -> 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,_) -> ConstRef sp @@ -209,7 +212,8 @@ let pattern_of_constr env sigma t = let push env na2 c2 = push_rel (LocalAssum (na2,c2)) env in let env' = Array.fold_left2 push env lna tl in PCoFix (ln,(lna,Array.map (pattern_of_constr env) tl, - Array.map (pattern_of_constr env') bl)) in + Array.map (pattern_of_constr env') bl)) + | Int i -> PInt i in pattern_of_constr env t (* To process patterns, we need a translation without typing at all. *) @@ -231,7 +235,7 @@ 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 _ as x) -> x + | (PVar _ | PEvar _ | PRel _ | PRef _ | PSort _ | PMeta _ | PInt _ as x) -> x let error_instantiate_pattern id l = let is = match l with @@ -287,7 +291,8 @@ let rec subst_pattern subst pat = pattern_of_constr env evd t.Univ.univ_abstracted_value) | PVar _ | PEvar _ - | PRel _ -> pat + | PRel _ + | PInt _ -> pat | PProj (p,c) -> let p' = Projection.map (subst_mind subst) p in let c' = subst_pattern subst c in @@ -488,6 +493,7 @@ let rec pat_of_raw metas vars = DAst.with_loc_val (fun ?loc -> function let names = Array.map (fun id -> Name id) ids in PCoFix (n, (names, tl, cl)) + | GInt i -> PInt i | GPatVar _ | GIf _ | GLetTuple _ | GCases _ | GEvar _ -> err ?loc (Pp.str "Non supported pattern.")) diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index c23b20a5d3..57705fa209 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -953,6 +953,14 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : GlobEnv.t) (sigma sigma, { uj_val = v; uj_type = tval } in inh_conv_coerce_to_tycon ?loc env sigma cj tycon + | GInt i -> + let resj = + try Typing.judge_of_int !!env i + with Invalid_argument _ -> + user_err ?loc ~hdr:"pretype" (str "Type of int63 should be registered first.") + in + inh_conv_coerce_to_tycon ?loc env sigma resj tycon + and pretype_instance k0 resolve_tc env sigma loc hyps evk update = let f decl (subst,update,sigma) = let id = NamedDecl.get_id decl in diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml index 9c9877fd23..98ca329117 100644 --- a/pretyping/reductionops.ml +++ b/pretyping/reductionops.ml @@ -279,7 +279,9 @@ sig | Case of case_info * 'a * 'a array * Cst_stack.t | Proj of Projection.t * Cst_stack.t | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t + | Primitive of CPrimitives.t * (Constant.t * EInstance.t) * 'a t * CPrimitives.args_red * Cst_stack.t | Cst of cst_member * int * int list * 'a t * Cst_stack.t + and 'a t = 'a member list exception IncompatibleFold2 @@ -308,6 +310,8 @@ sig val nth : 'a t -> int -> 'a val best_state : evar_map -> constr * constr t -> Cst_stack.t -> constr * constr t val zip : ?refold:bool -> evar_map -> constr * constr t -> constr + val check_native_args : CPrimitives.t -> 'a t -> bool + val get_next_primitive_args : CPrimitives.args_red -> 'a t -> CPrimitives.args_red * ('a t * 'a * 'a t) option end = struct open EConstr @@ -336,7 +340,9 @@ struct | Case of case_info * 'a * 'a array * Cst_stack.t | Proj of Projection.t * Cst_stack.t | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t + | Primitive of CPrimitives.t * (Constant.t * EInstance.t) * 'a t * CPrimitives.args_red * Cst_stack.t | Cst of cst_member * int * int list * 'a t * Cst_stack.t + and 'a t = 'a member list (* Debugging printer *) @@ -354,6 +360,9 @@ struct | Fix (f,args,cst) -> str "ZFix(" ++ Constr.debug_print_fix pr_c f ++ pr_comma () ++ pr pr_c args ++ str ")" + | Primitive (p,c,args,kargs,cst_l) -> + str "ZPrimitive(" ++ str (CPrimitives.to_string p) + ++ pr_comma () ++ pr pr_c args ++ str ")" | Cst (mem,curr,remains,params,cst_l) -> str "ZCst(" ++ pr_cst_member pr_c mem ++ pr_comma () ++ int curr ++ pr_comma () ++ @@ -420,7 +429,7 @@ struct equal_cst_member c1 c2 && equal_rec (List.rev params1) (List.rev params2) && equal_rec s1' s2' - | ((App _|Case _|Proj _|Fix _|Cst _)::_|[]), _ -> false + | ((App _|Case _|Proj _|Fix _|Cst _|Primitive _)::_|[]), _ -> false in equal_rec (List.rev sk1) (List.rev sk2) let compare_shape stk1 stk2 = @@ -435,9 +444,11 @@ struct Int.equal bal 0 && compare_rec 0 s1 s2 | (Fix(_,a1,_)::s1, Fix(_,a2,_)::s2) -> Int.equal bal 0 && compare_rec 0 a1 a2 && compare_rec 0 s1 s2 + | (Primitive(_,_,a1,_,_)::s1, Primitive(_,_,a2,_,_)::s2) -> + Int.equal bal 0 && compare_rec 0 a1 a2 && compare_rec 0 s1 s2 | (Cst (_,_,_,p1,_)::s1, Cst (_,_,_,p2,_)::s2) -> - Int.equal bal 0 && compare_rec 0 p1 p2 && compare_rec 0 s1 s2 - | (_,_) -> false in + Int.equal bal 0 && compare_rec 0 p1 p2 && compare_rec 0 s1 s2 + | ((Case _|Proj _|Fix _|Cst _|Primitive _) :: _ | []) ,_ -> false in compare_rec 0 stk1 stk2 exception IncompatibleFold2 @@ -459,7 +470,7 @@ struct | Cst (cst1,_,_,params1,_) :: q1, Cst (cst2,_,_,params2,_) :: q2 -> let o' = aux o (List.rev params1) (List.rev params2) in aux o' q1 q2 - | (((App _|Case _|Proj _|Fix _| Cst _) :: _|[]), _) -> + | (((App _|Case _|Proj _|Fix _|Cst _|Primitive _) :: _|[]), _) -> raise IncompatibleFold2 in aux o (List.rev sk1) (List.rev sk2) @@ -473,7 +484,9 @@ struct Fix ((r,(na,Array.map f ty, Array.map f bo)),map f arg,alt) | Cst (cst,curr,remains,params,alt) -> Cst (cst,curr,remains,map f params,alt) - ) x + | Primitive (p,c,args,kargs,cst_l) -> + Primitive(p,c, map f args, kargs, cst_l) + ) x let append_app_list l s = let a = Array.of_list l in @@ -481,7 +494,7 @@ struct let rec args_size = function | App (i,_,j)::s -> j + 1 - i + args_size s - | (Case _|Fix _|Proj _|Cst _)::_ | [] -> 0 + | (Case _|Fix _|Proj _|Cst _|Primitive _)::_ | [] -> 0 let strip_app s = let rec aux out = function @@ -504,7 +517,8 @@ struct in aux n [] s let not_purely_applicative args = - List.exists (function (Fix _ | Case _ | Proj _ | Cst _) -> true | _ -> false) args + List.exists (function (Fix _ | Case _ | Proj _ | Cst _) -> true + | App _ | Primitive _ -> false) args let will_expose_iota args = List.exists (function (Fix (_,_,l) | Case (_,_,_,l) | @@ -588,9 +602,26 @@ struct | f, (Proj (p,cst_l)::s) when refold -> zip (best_state sigma (mkProj (p,f),s) cst_l) | f, (Proj (p,_)::s) -> zip (mkProj (p,f),s) + | f, (Primitive (p,c,args,kargs,cst_l)::s) -> + zip (mkConstU c, args @ append_app [|f|] s) in zip s + (* Check if there is enough arguments on [stk] w.r.t. arity of [op] *) + let check_native_args op stk = + let nargs = CPrimitives.arity op in + let rargs = args_size stk in + nargs <= rargs + + let get_next_primitive_args kargs stk = + let rec nargs = function + | [] -> 0 + | CPrimitives.Kwhnf :: _ -> 0 + | _ :: s -> 1 + nargs s + in + let n = nargs kargs in + (List.skipn (n+1) kargs, strip_n_app n stk) + end (** The type of (machine) states (= lambda-bar-calculus' cuts) *) @@ -815,6 +846,57 @@ let fix_recarg ((recindices,bodynum),_) stack = with Not_found -> None +open Primred + +module CNativeEntries = +struct + + type elem = EConstr.t + type args = EConstr.t array + type evd = evar_map + + let get = Array.get + + let get_int evd e = + match EConstr.kind evd e with + | Int i -> i + | _ -> raise Primred.NativeDestKO + + let mkInt env i = + mkInt i + + let mkBool env b = + let (ct,cf) = get_bool_constructors env in + mkConstruct (if b then ct else cf) + + let mkCarry env b e = + let int_ty = mkConst @@ get_int_type env in + let (c0,c1) = get_carry_constructors env in + mkApp (mkConstruct (if b then c1 else c0),[|int_ty;e|]) + + let mkIntPair env e1 e2 = + let int_ty = mkConst @@ get_int_type env in + let c = get_pair_constructor env in + mkApp(mkConstruct c, [|int_ty;int_ty;e1;e2|]) + + let mkLt env = + let (_eq, lt, _gt) = get_cmp_constructors env in + mkConstruct lt + + let mkEq env = + let (eq, _lt, _gt) = get_cmp_constructors env in + mkConstruct eq + + let mkGt env = + let (_eq, _lt, gt) = get_cmp_constructors env in + mkConstruct gt + +end + +module CredNative = RedNative(CNativeEntries) + + + (** Generic reduction function with environment Here is where unfolded constant are stored in order to be @@ -881,47 +963,55 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = (lazy (EConstr.to_constr sigma (Stack.zip sigma (x,stack)))); if CClosure.RedFlags.red_set flags (CClosure.RedFlags.fCONST c) then let u' = EInstance.kind sigma u in - (match constant_opt_value_in env (c, u') with - | None -> fold () - | Some body -> + match constant_value_in env (c, u') with + | body -> + begin let body = EConstr.of_constr body in - if not tactic_mode - then whrec (if refold then Cst_stack.add_cst (mkConstU const) cst_l else cst_l) - (body, stack) - else (* Looks for ReductionBehaviour *) - match ReductionBehaviour.get (Globnames.ConstRef c) with - | None -> whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, stack) - | Some (recargs, nargs, flags) -> - if (List.mem `ReductionNeverUnfold flags - || (nargs > 0 && Stack.args_size stack < nargs)) - then fold () - else (* maybe unfolds *) - if List.mem `ReductionDontExposeCase flags then - let app_sk,sk = Stack.strip_app stack in - let (tm',sk'),cst_l' = - whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, app_sk) - in - let rec is_case x = match EConstr.kind sigma x with - | Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x - | App (hd, _) -> is_case hd - | Case _ -> true - | _ -> false in - if equal_stacks sigma (x, app_sk) (tm', sk') - || Stack.will_expose_iota sk' - || is_case tm' - then fold () - else whrec cst_l' (tm', sk' @ sk) - else match recargs with - |[] -> (* if nargs has been specified *) - (* CAUTION : the constant is NEVER refold - (even when it hides a (co)fix) *) - whrec cst_l (body, stack) - |curr::remains -> match Stack.strip_n_app curr stack with - | None -> fold () - | Some (bef,arg,s') -> - whrec Cst_stack.empty - (arg,Stack.Cst(Stack.Cst_const (fst const, u'),curr,remains,bef,cst_l)::s') - ) else fold () + if not tactic_mode + then whrec (if refold then Cst_stack.add_cst (mkConstU const) cst_l else cst_l) + (body, stack) + else (* Looks for ReductionBehaviour *) + match ReductionBehaviour.get (Globnames.ConstRef c) with + | None -> whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, stack) + | Some (recargs, nargs, flags) -> + if (List.mem `ReductionNeverUnfold flags + || (nargs > 0 && Stack.args_size stack < nargs)) + then fold () + else (* maybe unfolds *) + if List.mem `ReductionDontExposeCase flags then + let app_sk,sk = Stack.strip_app stack in + let (tm',sk'),cst_l' = + whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, app_sk) + in + let rec is_case x = match EConstr.kind sigma x with + | Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x + | App (hd, _) -> is_case hd + | Case _ -> true + | _ -> false in + if equal_stacks sigma (x, app_sk) (tm', sk') + || Stack.will_expose_iota sk' + || is_case tm' + then fold () + else whrec cst_l' (tm', sk' @ sk) + else match recargs with + |[] -> (* if nargs has been specified *) + (* CAUTION : the constant is NEVER refold + (even when it hides a (co)fix) *) + whrec cst_l (body, stack) + |curr::remains -> match Stack.strip_n_app curr stack with + | None -> fold () + | Some (bef,arg,s') -> + whrec Cst_stack.empty + (arg,Stack.Cst(Stack.Cst_const (fst const, u'),curr,remains,bef,cst_l)::s') + end + | exception NotEvaluableConst (IsPrimitive p) when Stack.check_native_args p stack -> + let kargs = CPrimitives.kind p in + let (kargs,o) = Stack.get_next_primitive_args kargs stack in + (* Should not fail thanks to [check_native_args] *) + let (before,a,after) = Option.get o in + whrec Cst_stack.empty (a,Stack.Primitive(p,const,before,kargs,cst_l)::after) + | exception NotEvaluableConst _ -> fold () + else fold () | Proj (p, c) when CClosure.RedFlags.red_projection flags p -> (let npars = Projection.npars p in if not tactic_mode then @@ -1049,6 +1139,30 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma = |_ -> fold () else fold () + | Int i -> + 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 + if more_to_reduce then + let (kargs,o) = Stack.get_next_primitive_args kargs s in + (* Should not fail because Primitive is put on the stack only if fully applied *) + let (before,a,after) = Option.get o in + whrec Cst_stack.empty (a,Stack.Primitive(p,kn,rargs @ Stack.append_app [|x|] before,kargs,cst_l')::after) + else + let n = List.length kargs in + let (args,s) = Stack.strip_app s in + let (args,extra_args) = + try List.chop n args + with List.IndexOutOfRange -> (args,[]) (* FIXME probably useless *) + in + let args = Array.of_list (Option.get (Stack.list_of_app_stack (rargs @ Stack.append_app [|x|] args))) in + begin match CredNative.red_prim env sigma p args with + | Some t -> whrec cst_l' (t,s) + | None -> ((mkApp (mkConstU kn, args), s), cst_l) + end + | _ -> fold () + end + | Rel _ | Var _ | LetIn _ | Proj _ -> fold () | Sort _ | Ind _ | Prod _ -> fold () in @@ -1127,7 +1241,8 @@ let local_whd_state_gen flags sigma = |_ -> s else s - | Rel _ | Var _ | Sort _ | Prod _ | LetIn _ | Const _ | Ind _ | Proj _ -> s + | Rel _ | Var _ | Sort _ | Prod _ | LetIn _ | Const _ | Ind _ | Proj _ + | Int _ -> s in whrec @@ -1577,7 +1692,7 @@ let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s = if isConstruct sigma t_o then whrec Cst_stack.empty (Stack.nth stack_o (Projection.npars p + Projection.arg p), stack'') else s,csts' - |_, ((Stack.App _|Stack.Cst _) :: _|[]) -> s,csts' + |_, ((Stack.App _|Stack.Cst _|Stack.Primitive _) :: _|[]) -> s,csts' in whrec csts s let find_conclusion env sigma = diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli index a1fd610676..fae0b23b83 100644 --- a/pretyping/reductionops.mli +++ b/pretyping/reductionops.mli @@ -77,6 +77,7 @@ module Stack : sig | Case of case_info * 'a * 'a array * Cst_stack.t | Proj of Projection.t * Cst_stack.t | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t + | Primitive of CPrimitives.t * (Constant.t * EInstance.t) * 'a t * CPrimitives.args_red * Cst_stack.t | Cst of cst_member * int (* current foccussed arg *) * int list (* remaining args *) diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 62ad296ecb..a76a203e37 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -143,6 +143,7 @@ let retype ?(polyprop=true) sigma = with Invalid_argument _ -> retype_error BadRecursiveType) | Cast (c,_, t) -> t | Sort _ | Prod _ -> mkSort (sort_of env cstr) + | Int _ -> EConstr.of_constr (Typeops.type_of_int env) and sort_of env t = match EConstr.kind sigma t with diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 2e7176a6b3..2308a541fb 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -48,8 +48,8 @@ let error_not_evaluable r = spc () ++ str "to an evaluable reference.") let is_evaluable_const env cst = - is_transparent env (ConstKey cst) && - evaluable_constant cst env + is_transparent env (ConstKey cst) && + (evaluable_constant cst env || is_primitive env cst) let is_evaluable_var env id = is_transparent env (VarKey id) && evaluable_named id env diff --git a/pretyping/typing.ml b/pretyping/typing.ml index b5729d7574..e8d935fcbb 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -306,6 +306,9 @@ let type_of_constructor env sigma ((ind,_ as ctor),u) = let sigma = Evd.add_constraints sigma csts in sigma, (EConstr.of_constr (rename_type ty (Names.GlobRef.ConstructRef ctor))) +let judge_of_int env v = + Termops.on_judgment EConstr.of_constr (judge_of_int 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 = @@ -408,6 +411,9 @@ let rec execute env sigma cstr = let sigma, tj = type_judgment env sigma tj in judge_of_cast env sigma cj k tj + | Int i -> + sigma, judge_of_int env i + 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 79f2941554..1ea16bbf34 100644 --- a/pretyping/typing.mli +++ b/pretyping/typing.mli @@ -55,3 +55,4 @@ val judge_of_abstraction : Environ.env -> Name.t -> 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 diff --git a/pretyping/unification.ml b/pretyping/unification.ml index f0cd5c5f6a..e4d96da0a6 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -550,7 +550,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 _ -> true + | Construct _ | Int _ -> true | Fix _ | CoFix _ -> true | Rel _ | Var _ | Meta _ | Evar _ | Sort _ | Cast (_, _, _) | Prod (_, _, _) | Lambda (_, _, _) | LetIn (_, _, _, _) | App (_, _) | Case (_, _, _, _) @@ -641,7 +641,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 _ -> false + | Lambda _ | LetIn _ | Construct _ | CoFix _ | Int _ -> false | Sort _ | Cast (_, _, _) | Prod (_, _, _) | Ind _ -> false (* Really? *) | Fix _ -> false (* This is an approximation *) | App _ -> assert false @@ -1799,7 +1799,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 _ -> user_err Pp.(str "Match_subterm"))) + | Construct _ | Int _ -> user_err Pp.(str "Match_subterm"))) in try matchrec cl with ex when precatchable_exception ex -> @@ -1868,7 +1868,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 _ -> fail "Match_subterm")) + | Construct _ | Int _ -> fail "Match_subterm")) in let res = matchrec cl [] in diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml index 113aac25da..083661a64b 100644 --- a/pretyping/vnorm.ml +++ b/pretyping/vnorm.ml @@ -57,7 +57,7 @@ let find_rectype_a env c = let (t, l) = decompose_appvect (whd_all env c) in match kind t with | Ind ind -> (ind, l) - | _ -> assert false + | _ -> raise Not_found (* Instantiate inductives and parameters in constructor type *) @@ -75,25 +75,18 @@ let type_constructor mind mib u typ params = let construct_of_constr const env tag typ = - let ((mind,_ as ind), u) as indu, allargs = find_rectype_a env typ in - (* spiwack : here be a branch for specific decompilation handled by retroknowledge *) - try - if const then - ((Retroknowledge.get_vm_decompile_constant_info env.retroknowledge (GlobRef.IndRef ind) tag), - typ) (*spiwack: this may need to be changed in case there are parameters in the - type which may cause a constant value to have an arity. - (type_constructor seems to be all about parameters actually) - but it shouldn't really matter since constant values don't use - their ctyp in the rest of the code.*) - else - raise Not_found (* No retroknowledge function (yet) for block decompilation *) - with Not_found -> + let (t, allargs) = decompose_appvect (whd_all env typ) in + match Constr.kind t with + | Ind ((mind,_ as ind), u as indu) -> let mib,mip = lookup_mind_specif env ind in let nparams = mib.mind_nparams in let i = invert_tag const tag mip.mind_reloc_tbl in let params = Array.sub allargs 0 nparams in let ctyp = type_constructor mind mib u (mip.mind_nf_lc.(i-1)) params in - (mkApp(mkConstructUi(indu,i), params), ctyp) + (mkApp(mkConstructUi(indu,i), params), ctyp) + | _ -> + assert (Constr.equal t (Typeops.type_of_int env)); + (mkInt (Uint63.of_int tag), t) let construct_of_constr_const env tag typ = fst (construct_of_constr true env tag typ) @@ -169,6 +162,7 @@ and nf_whd env sigma whd typ = let capp,ctyp = construct_of_constr_block env tag typ in let args = nf_bargs env sigma b ofs ctyp in mkApp(capp,args) + | Vint64 i -> i |> Uint63.of_int64 |> mkInt | Vatom_stk(Aid idkey, stk) -> constr_type_of_idkey env sigma idkey stk | Vatom_stk(Aind ((mi,i) as ind), stk) -> @@ -344,9 +338,9 @@ and nf_fun env sigma f typ = let name,dom,codom = try decompose_prod env typ with DestKO -> - (* 27/2/13: Turned this into an anomaly *) CErrors.anomaly - (Pp.strbrk "Returned a functional value in a type not recognized as a product type.") + Pp.(strbrk "Returned a functional value in type " ++ + Termops.Internal.print_constr_env env sigma (EConstr.of_constr typ)) in let body = nf_val (push_rel (LocalAssum (name,dom)) env) sigma vb codom in mkLambda(name,dom,body) diff --git a/printing/prettyp.ml b/printing/prettyp.ml index 408bd5f60b..e0403a9f97 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -193,7 +193,7 @@ let opacity env = | ConstRef cst -> let cb = Environ.lookup_constant cst env in (match cb.const_body with - | Undef _ -> None + | Undef _ | Primitive _ -> None | OpaqueDef _ -> Some FullyOpaque | Def _ -> Some (TransparentMaybeOpacified @@ -558,7 +558,7 @@ let print_constant with_values sep sp udecl = let open Univ in let otab = Global.opaque_tables () in match cb.const_body with - | Undef _ | Def _ -> cb.const_universes + | Undef _ | Def _ | Primitive _ -> cb.const_universes | OpaqueDef o -> let body_uctxs = Opaqueproof.force_constraints otab o in match cb.const_universes with @@ -724,7 +724,10 @@ let print_full_pure_context env sigma = | Def c -> str "Definition " ++ print_basename con ++ cut () ++ str " : " ++ pr_ltype_env env sigma typ ++ cut () ++ str " := " ++ - pr_lconstr_env env sigma (Mod_subst.force_constr c)) + pr_lconstr_env env sigma (Mod_subst.force_constr c) + | Primitive _ -> + str "Primitive " ++ + print_basename con ++ str " : " ++ cut () ++ pr_ltype_env env sigma typ) ++ str "." ++ fnl () ++ fnl () | "INDUCTIVE" -> let mind = Global.mind_of_delta_kn kn in diff --git a/stm/vernac_classifier.ml b/stm/vernac_classifier.ml index 292e3966a1..42540af991 100644 --- a/stm/vernac_classifier.ml +++ b/stm/vernac_classifier.ml @@ -130,6 +130,8 @@ let classify_vernac e = | VernacAssumption (_,_,l) -> let ids = List.flatten (List.map (fun (_,(l,_)) -> List.map (fun (id, _) -> id.v) l) l) in VtSideff ids, VtLater + | VernacPrimitive (id,_,_) -> + VtSideff [id.CAst.v], VtLater | VernacDefinition (_,({v=id},_),DefineBody _) -> VtSideff (idents_of_name id), VtLater | VernacInductive (_, _,_,l) -> let ids = List.map (fun (((_,({v=id},_)),_,_,_,cl),_) -> id :: match cl with diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml index e273891500..e8a66f1889 100644 --- a/tactics/term_dnet.ml +++ b/tactics/term_dnet.ml @@ -44,6 +44,7 @@ struct | DCase of case_info * 't * 't * 't array | DFix of int array * int * 't array * 't array | DCoFix of int * 't array * 't array + | DInt of Uint63.t (* special constructors only inside the left-hand side of DCtx or DApp. Used to encode lists of foralls/letins/apps as contexts *) @@ -61,6 +62,7 @@ struct | DCase (_,t1,t2,ta) -> str "case" | DFix _ -> str "fix" | DCoFix _ -> str "cofix" + | DInt _ -> str "INT" | DCons ((t,dopt),tl) -> f t ++ (match dopt with Some t' -> str ":=" ++ f t' | None -> str "") ++ spc() ++ str "::" ++ spc() ++ f tl @@ -72,7 +74,7 @@ struct *) let map f = function - | (DRel | DSort | DNil | DRef _) as c -> c + | (DRel | DSort | DNil | DRef _ | DInt _) 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) @@ -145,6 +147,10 @@ struct else c | DCoFix _, _ -> -1 | _, DCoFix _ -> 1 + | DInt i1, DInt i2 -> Uint63.compare i1 i2 + + | DInt _, _ -> -1 | _, DInt _ -> 1 + | DCons ((t1, ot1), u1), DCons ((t2, ot2), u2) -> let c = cmp t1 t2 in if Int.equal c 0 then @@ -157,7 +163,7 @@ struct | DNil, DNil -> 0 let fold f acc = function - | (DRel | DNil | DSort | DRef _) -> acc + | (DRel | DNil | DSort | DRef _ | DInt _) -> 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 @@ -169,7 +175,7 @@ struct | DCons ((t,topt),u) -> f (Option.fold_left f (f acc t) topt) u let choose f = function - | (DRel | DSort | DNil | DRef _) -> invalid_arg "choose" + | (DRel | DSort | DNil | DRef _ | DInt _) -> invalid_arg "choose" | DCtx (ctx,c) -> f ctx | DLambda (t,c) -> f t | DApp (t,u) -> f u @@ -185,7 +191,8 @@ struct if not (Int.equal (compare dummy_cmp (head c1) (head c2)) 0) then invalid_arg "fold2:compare" else match c1,c2 with - | (DRel, DRel | DNil, DNil | DSort, DSort | DRef _, DRef _) -> acc + | (DRel, DRel | DNil, DNil | DSort, DSort | DRef _, DRef _ + | DInt _, DInt _) -> 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 @@ -198,14 +205,15 @@ 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 _), _ -> assert false + | DFix _ | DCoFix _ | DCons _ | DInt _), _ -> assert false let map2 (f:'a -> 'b -> 'c) (c1:'a t) (c2:'b t) : 'c t = let head w = map (fun _ -> ()) w in if not (Int.equal (compare dummy_cmp (head c1) (head c2)) 0) then invalid_arg "map2_t:compare" else match c1,c2 with - | (DRel, DRel | DSort, DSort | DNil, DNil | DRef _, DRef _) as cc -> + | (DRel, DRel | DSort, DSort | DNil, DNil | DRef _, DRef _ + | DInt _, DInt _) 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) @@ -219,10 +227,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 _), _ -> assert false + | DFix _ | DCoFix _ | DCons _ | DInt _), _ -> assert false let terminal = function - | (DRel | DSort | DNil | DRef _) -> true + | (DRel | DSort | DNil | DRef _ | DInt _) -> true | DLambda _ | DApp _ | DCase _ | DFix _ | DCoFix _ | DCtx _ | DCons _ -> false @@ -315,6 +323,7 @@ struct (pat_of_constr f) (Array.map pat_of_constr ca) | Proj (p,c) -> Term (DApp (Term (DRef (ConstRef (Projection.constant p))), pat_of_constr c)) + | Int i -> Term (DInt i) 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 cafc9a744c..68acb6f04d 100644 --- a/test-suite/Makefile +++ b/test-suite/Makefile @@ -95,7 +95,7 @@ INTERACTIVE := interactive UNIT_TESTS := unit-tests VSUBSYSTEMS := prerequisite success failure $(BUGS) output \ output-modulo-time $(INTERACTIVE) micromega $(COMPLEXITY) modules stm \ - coqdoc ssr + coqdoc ssr arithmetic # All subsystems SUBSYSTEMS := $(VSUBSYSTEMS) misc bugs ide vio coqchk coqwc coq-makefile tools $(UNIT_TESTS) @@ -176,6 +176,7 @@ summary: $(call summary_dir, "Coqdoc tests", coqdoc); \ $(call summary_dir, "tools/ tests", tools); \ $(call summary_dir, "Unit tests", unit-tests); \ + $(call summary_dir, "Machine arithmetic tests", arithmetic); \ nb_success=`find . -name '*.log' -exec tail -n2 '{}' \; | grep -e $(log_success) | wc -l`; \ nb_failure=`find . -name '*.log' -exec tail -n2 '{}' \; | grep -e $(log_failure) | wc -l`; \ nb_tests=`expr $$nb_success + $$nb_failure`; \ @@ -308,7 +309,7 @@ $(addsuffix .log,$(wildcard prerequisite/*.v)): %.v.log: %.v $(HIDE)$(coqchk) -norec TestSuite.$(shell basename $< .v) > $(shell dirname $<)/$(shell basename $< .v).chk.log 2>&1 ssr: $(wildcard ssr/*.v:%.v=%.v.log) -$(addsuffix .log,$(wildcard ssr/*.v success/*.v micromega/*.v modules/*.v)): %.v.log: %.v $(PREREQUISITELOG) +$(addsuffix .log,$(wildcard ssr/*.v success/*.v micromega/*.v modules/*.v arithmetic/*.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/arithmetic/add.v b/test-suite/arithmetic/add.v new file mode 100644 index 0000000000..fb7eb1d53c --- /dev/null +++ b/test-suite/arithmetic/add.v @@ -0,0 +1,18 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 2 + 3 = 5). +Check (eq_refl 5 <: 2 + 3 = 5). +Check (eq_refl 5 <<: 2 + 3 = 5). + +Definition compute1 := Eval compute in 2 + 3. +Check (eq_refl compute1 : 5 = 5). + +Check (eq_refl : 9223372036854775807 + 1 = 0). +Check (eq_refl 0 <: 9223372036854775807 + 1 = 0). +Check (eq_refl 0 <<: 9223372036854775807 + 1 = 0). +Definition compute2 := Eval compute in 9223372036854775807 + 1. +Check (eq_refl compute2 : 0 = 0). diff --git a/test-suite/arithmetic/addc.v b/test-suite/arithmetic/addc.v new file mode 100644 index 0000000000..432aec0291 --- /dev/null +++ b/test-suite/arithmetic/addc.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 2 +c 3 = C0 5). +Check (eq_refl (C0 5) <: 2 +c 3 = C0 5). +Check (eq_refl (C0 5) <<: 2 +c 3 = C0 5). +Definition compute1 := Eval compute in 2 +c 3. +Check (eq_refl compute1 : C0 5 = C0 5). + +Check (eq_refl : 9223372036854775807 +c 2 = C1 1). +Check (eq_refl (C1 1) <: 9223372036854775807 +c 2 = C1 1). +Check (eq_refl (C1 1) <<: 9223372036854775807 +c 2 = C1 1). +Definition compute2 := Eval compute in 9223372036854775807 +c 2. +Check (eq_refl compute2 : C1 1 = C1 1). diff --git a/test-suite/arithmetic/addcarryc.v b/test-suite/arithmetic/addcarryc.v new file mode 100644 index 0000000000..a4430769ca --- /dev/null +++ b/test-suite/arithmetic/addcarryc.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : addcarryc 2 3 = C0 6). +Check (eq_refl (C0 6) <: addcarryc 2 3 = C0 6). +Check (eq_refl (C0 6) <<: addcarryc 2 3 = C0 6). +Definition compute1 := Eval compute in addcarryc 2 3. +Check (eq_refl compute1 : C0 6 = C0 6). + +Check (eq_refl : addcarryc 9223372036854775807 2 = C1 2). +Check (eq_refl (C1 2) <: addcarryc 9223372036854775807 2 = C1 2). +Check (eq_refl (C1 2) <<: addcarryc 9223372036854775807 2 = C1 2). +Definition compute2 := Eval compute in addcarryc 9223372036854775807 2. +Check (eq_refl compute2 : C1 2 = C1 2). diff --git a/test-suite/arithmetic/addmuldiv.v b/test-suite/arithmetic/addmuldiv.v new file mode 100644 index 0000000000..72b0164b49 --- /dev/null +++ b/test-suite/arithmetic/addmuldiv.v @@ -0,0 +1,12 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : addmuldiv 32 3 5629499534213120 = 12887523328). +Check (eq_refl 12887523328 <: addmuldiv 32 3 5629499534213120 = 12887523328). +Check (eq_refl 12887523328 <<: addmuldiv 32 3 5629499534213120 = 12887523328). + +Definition compute2 := Eval compute in addmuldiv 32 3 5629499534213120. +Check (eq_refl compute2 : 12887523328 = 12887523328). diff --git a/test-suite/arithmetic/compare.v b/test-suite/arithmetic/compare.v new file mode 100644 index 0000000000..a8d1ea1226 --- /dev/null +++ b/test-suite/arithmetic/compare.v @@ -0,0 +1,23 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 1 ?= 1 = Eq). +Check (eq_refl Eq <: 1 ?= 1 = Eq). +Check (eq_refl Eq <<: 1 ?= 1 = Eq). +Definition compute1 := Eval compute in 1 ?= 1. +Check (eq_refl compute1 : Eq = Eq). + +Check (eq_refl : 1 ?= 2 = Lt). +Check (eq_refl Lt <: 1 ?= 2 = Lt). +Check (eq_refl Lt <<: 1 ?= 2 = Lt). +Definition compute2 := Eval compute in 1 ?= 2. +Check (eq_refl compute2 : Lt = Lt). + +Check (eq_refl : 9223372036854775807 ?= 0 = Gt). +Check (eq_refl Gt <: 9223372036854775807 ?= 0 = Gt). +Check (eq_refl Gt <<: 9223372036854775807 ?= 0 = Gt). +Definition compute3 := Eval compute in 9223372036854775807 ?= 0. +Check (eq_refl compute3 : Gt = Gt). diff --git a/test-suite/arithmetic/div.v b/test-suite/arithmetic/div.v new file mode 100644 index 0000000000..0ee0b91580 --- /dev/null +++ b/test-suite/arithmetic/div.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 6 / 3 = 2). +Check (eq_refl 2 <: 6 / 3 = 2). +Check (eq_refl 2 <<: 6 / 3 = 2). +Definition compute1 := Eval compute in 6 / 3. +Check (eq_refl compute1 : 2 = 2). + +Check (eq_refl : 3 / 2 = 1). +Check (eq_refl 1 <: 3 / 2 = 1). +Check (eq_refl 1 <<: 3 / 2 = 1). +Definition compute2 := Eval compute in 3 / 2. +Check (eq_refl compute2 : 1 = 1). diff --git a/test-suite/arithmetic/diveucl.v b/test-suite/arithmetic/diveucl.v new file mode 100644 index 0000000000..8f88a0f356 --- /dev/null +++ b/test-suite/arithmetic/diveucl.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : diveucl 6 3 = (2,0)). +Check (eq_refl (2,0) <: diveucl 6 3 = (2,0)). +Check (eq_refl (2,0) <<: diveucl 6 3 = (2,0)). +Definition compute1 := Eval compute in diveucl 6 3. +Check (eq_refl compute1 : (2,0) = (2,0)). + +Check (eq_refl : diveucl 5 3 = (1,2)). +Check (eq_refl (1,2) <: diveucl 5 3 = (1,2)). +Check (eq_refl (1,2) <<: diveucl 5 3 = (1,2)). +Definition compute2 := Eval compute in diveucl 5 3. +Check (eq_refl compute2 : (1,2) = (1,2)). diff --git a/test-suite/arithmetic/diveucl_21.v b/test-suite/arithmetic/diveucl_21.v new file mode 100644 index 0000000000..7e12a08906 --- /dev/null +++ b/test-suite/arithmetic/diveucl_21.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : diveucl_21 1 1 2 = (4611686018427387904,1)). +Check (eq_refl (4611686018427387904,1) <: diveucl_21 1 1 2 = (4611686018427387904,1)). +Check (eq_refl (4611686018427387904,1) <<: diveucl_21 1 1 2 = (4611686018427387904,1)). +Definition compute1 := Eval compute in diveucl_21 1 1 2. +Check (eq_refl compute1 : (4611686018427387904,1) = (4611686018427387904,1)). + +Check (eq_refl : diveucl_21 3 1 2 = (4611686018427387904, 1)). +Check (eq_refl (4611686018427387904, 1) <: diveucl_21 3 1 2 = (4611686018427387904, 1)). +Check (eq_refl (4611686018427387904, 1) <<: diveucl_21 3 1 2 = (4611686018427387904, 1)). +Definition compute2 := Eval compute in diveucl_21 3 1 2. +Check (eq_refl compute2 : (4611686018427387904, 1) = (4611686018427387904, 1)). diff --git a/test-suite/arithmetic/eqb.v b/test-suite/arithmetic/eqb.v new file mode 100644 index 0000000000..dcc0b71f6d --- /dev/null +++ b/test-suite/arithmetic/eqb.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 1 == 1 = true). +Check (eq_refl true <: 1 == 1 = true). +Check (eq_refl true <<: 1 == 1 = true). +Definition compute1 := Eval compute in 1 == 1. +Check (eq_refl compute1 : true = true). + +Check (eq_refl : 9223372036854775807 == 0 = false). +Check (eq_refl false <: 9223372036854775807 == 0 = false). +Check (eq_refl false <<: 9223372036854775807 == 0 = false). +Definition compute2 := Eval compute in 9223372036854775807 == 0. +Check (eq_refl compute2 : false = false). diff --git a/test-suite/arithmetic/head0.v b/test-suite/arithmetic/head0.v new file mode 100644 index 0000000000..f4234d2605 --- /dev/null +++ b/test-suite/arithmetic/head0.v @@ -0,0 +1,23 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : head0 3 = 61). +Check (eq_refl 61 <: head0 3 = 61). +Check (eq_refl 61 <<: head0 3 = 61). +Definition compute1 := Eval compute in head0 3. +Check (eq_refl compute1 : 61 = 61). + +Check (eq_refl : head0 4611686018427387904 = 0). +Check (eq_refl 0 <: head0 4611686018427387904 = 0). +Check (eq_refl 0 <<: head0 4611686018427387904 = 0). +Definition compute2 := Eval compute in head0 4611686018427387904. +Check (eq_refl compute2 : 0 = 0). + +Check (eq_refl : head0 0 = 63). +Check (eq_refl 63 <: head0 0 = 63). +Check (eq_refl 63 <<: head0 0 = 63). +Definition compute3 := Eval compute in head0 0. +Check (eq_refl compute3 : 63 = 63). diff --git a/test-suite/arithmetic/isint.v b/test-suite/arithmetic/isint.v new file mode 100644 index 0000000000..c215caa878 --- /dev/null +++ b/test-suite/arithmetic/isint.v @@ -0,0 +1,50 @@ +(* This file tests the check that arithmetic operations use to know if their +arguments are ground. The various test cases correspond to possible +optimizations of these tests made by the compiler. *) +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Section test. + +Variable m n : int. + +Check (eq_refl : (fun x => x + 3) m = m + 3). +Check (eq_refl (m + 3) <: (fun x => x + 3) m = m + 3). +Check (eq_refl (m + 3) <<: (fun x => x + 3) m = m + 3). +Definition compute1 := Eval compute in (fun x => x + 3) m. +Check (eq_refl compute1 : m + 3 = m + 3). + +Check (eq_refl : (fun x => 3 + x) m = 3 + m). +Check (eq_refl (3 + m) <: (fun x => 3 + x) m = 3 + m). +Check (eq_refl (3 + m) <<: (fun x => 3 + x) m = 3 + m). +Definition compute2 := Eval compute in (fun x => 3 + x) m. +Check (eq_refl compute2 : 3 + m = 3 + m). + +Check (eq_refl : (fun x y => x + y) m n = m + n). +Check (eq_refl (m + n) <: (fun x y => x + y) m n = m + n). +Check (eq_refl (m + n) <<: (fun x y => x + y) m n = m + n). +Definition compute3 := Eval compute in (fun x y => x + y) m n. +Check (eq_refl compute3 : m + n = m + n). + +Check (eq_refl : (fun x y => x + y) 2 3 = 5). +Check (eq_refl 5 <: (fun x y => x + y) 2 3 = 5). +Check (eq_refl 5 <<: (fun x y => x + y) 2 3 = 5). +Definition compute4 := Eval compute in (fun x y => x + y) 2 3. +Check (eq_refl compute4 : 5 = 5). + +Check (eq_refl : (fun x => x + x) m = m + m). +Check (eq_refl (m + m) <: (fun x => x + x) m = m + m). +Check (eq_refl (m + m) <<: (fun x => x + x) m = m + m). +Definition compute5 := Eval compute in (fun x => x + x) m. +Check (eq_refl compute5 : m + m = m + m). + +Check (eq_refl : (fun x => x + x) 2 = 4). +Check (eq_refl 4 <: (fun x => x + x) 2 = 4). +Check (eq_refl 4 <<: (fun x => x + x) 2 = 4). +Definition compute6 := Eval compute in (fun x => x + x) 2. +Check (eq_refl compute6 : 4 = 4). + +End test. diff --git a/test-suite/arithmetic/land.v b/test-suite/arithmetic/land.v new file mode 100644 index 0000000000..0ea6fd90b6 --- /dev/null +++ b/test-suite/arithmetic/land.v @@ -0,0 +1,29 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 0 land 0 = 0). +Check (eq_refl 0 <: 0 land 0 = 0). +Check (eq_refl 0 <<: 0 land 0 = 0). +Definition compute1 := Eval compute in 0 land 0. +Check (eq_refl compute1 : 0 = 0). + +Check (eq_refl : 9223372036854775807 land 0 = 0). +Check (eq_refl 0 <: 9223372036854775807 land 0 = 0). +Check (eq_refl 0 <<: 9223372036854775807 land 0 = 0). +Definition compute2 := Eval compute in 9223372036854775807 land 0. +Check (eq_refl compute2 : 0 = 0). + +Check (eq_refl : 0 land 9223372036854775807 = 0). +Check (eq_refl 0 <: 0 land 9223372036854775807 = 0). +Check (eq_refl 0 <<: 0 land 9223372036854775807 = 0). +Definition compute3 := Eval compute in 0 land 9223372036854775807. +Check (eq_refl compute3 : 0 = 0). + +Check (eq_refl : 9223372036854775807 land 9223372036854775807 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <: 9223372036854775807 land 9223372036854775807 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <<: 9223372036854775807 land 9223372036854775807 = 9223372036854775807). +Definition compute4 := Eval compute in 9223372036854775807 land 9223372036854775807. +Check (eq_refl compute4 : 9223372036854775807 = 9223372036854775807). diff --git a/test-suite/arithmetic/leb.v b/test-suite/arithmetic/leb.v new file mode 100644 index 0000000000..5354919978 --- /dev/null +++ b/test-suite/arithmetic/leb.v @@ -0,0 +1,23 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 1 <= 1 = true). +Check (eq_refl true <: 1 <= 1 = true). +Check (eq_refl true <<: 1 <= 1 = true). +Definition compute1 := Eval compute in 1 <= 1. +Check (eq_refl compute1 : true = true). + +Check (eq_refl : 1 <= 2 = true). +Check (eq_refl true <: 1 <= 2 = true). +Check (eq_refl true <<: 1 <= 2 = true). +Definition compute2 := Eval compute in 1 <= 2. +Check (eq_refl compute2 : true = true). + +Check (eq_refl : 9223372036854775807 <= 0 = false). +Check (eq_refl false <: 9223372036854775807 <= 0 = false). +Check (eq_refl false <<: 9223372036854775807 <= 0 = false). +Definition compute3 := Eval compute in 9223372036854775807 <= 0. +Check (eq_refl compute3 : false = false). diff --git a/test-suite/arithmetic/lor.v b/test-suite/arithmetic/lor.v new file mode 100644 index 0000000000..9c3b85c054 --- /dev/null +++ b/test-suite/arithmetic/lor.v @@ -0,0 +1,29 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 0 lor 0 = 0). +Check (eq_refl 0 <: 0 lor 0 = 0). +Check (eq_refl 0 <<: 0 lor 0 = 0). +Definition compute1 := Eval compute in 0 lor 0. +Check (eq_refl compute1 : 0 = 0). + +Check (eq_refl : 9223372036854775807 lor 0 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <: 9223372036854775807 lor 0 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <<: 9223372036854775807 lor 0 = 9223372036854775807). +Definition compute2 := Eval compute in 9223372036854775807 lor 0. +Check (eq_refl compute2 : 9223372036854775807 = 9223372036854775807). + +Check (eq_refl : 0 lor 9223372036854775807 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <: 0 lor 9223372036854775807 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <<: 0 lor 9223372036854775807 = 9223372036854775807). +Definition compute3 := Eval compute in 0 lor 9223372036854775807. +Check (eq_refl compute3 : 9223372036854775807 = 9223372036854775807). + +Check (eq_refl : 9223372036854775807 lor 9223372036854775807 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <: 9223372036854775807 lor 9223372036854775807 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <<: 9223372036854775807 lor 9223372036854775807 = 9223372036854775807). +Definition compute4 := Eval compute in 9223372036854775807 lor 9223372036854775807. +Check (eq_refl compute4 : 9223372036854775807 = 9223372036854775807). diff --git a/test-suite/arithmetic/lsl.v b/test-suite/arithmetic/lsl.v new file mode 100644 index 0000000000..70f3b90140 --- /dev/null +++ b/test-suite/arithmetic/lsl.v @@ -0,0 +1,23 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 3 << 61 = 6917529027641081856). +Check (eq_refl 6917529027641081856 <: 3 << 61 = 6917529027641081856). +Check (eq_refl 6917529027641081856 <<: 3 << 61 = 6917529027641081856). +Definition compute1 := Eval compute in 3 << 61. +Check (eq_refl compute1 : 6917529027641081856 = 6917529027641081856). + +Check (eq_refl : 2 << 62 = 0). +Check (eq_refl 0 <: 2 << 62 = 0). +Check (eq_refl 0 <<: 2 << 62 = 0). +Definition compute2 := Eval compute in 2 << 62. +Check (eq_refl compute2 : 0 = 0). + +Check (eq_refl : 9223372036854775807 << 64 = 0). +Check (eq_refl 0 <: 9223372036854775807 << 64 = 0). +Check (eq_refl 0 <<: 9223372036854775807 << 64 = 0). +Definition compute3 := Eval compute in 9223372036854775807 << 64. +Check (eq_refl compute3 : 0 = 0). diff --git a/test-suite/arithmetic/lsr.v b/test-suite/arithmetic/lsr.v new file mode 100644 index 0000000000..c36c24e237 --- /dev/null +++ b/test-suite/arithmetic/lsr.v @@ -0,0 +1,23 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 6917529027641081856 >> 61 = 3). +Check (eq_refl 3 <: 6917529027641081856 >> 61 = 3). +Check (eq_refl 3 <<: 6917529027641081856 >> 61 = 3). +Definition compute1 := Eval compute in 6917529027641081856 >> 61. +Check (eq_refl compute1 : 3 = 3). + +Check (eq_refl : 2305843009213693952 >> 62 = 0). +Check (eq_refl 0 <: 2305843009213693952 >> 62 = 0). +Check (eq_refl 0 <<: 2305843009213693952 >> 62 = 0). +Definition compute2 := Eval compute in 2305843009213693952 >> 62. +Check (eq_refl compute2 : 0 = 0). + +Check (eq_refl : 9223372036854775807 >> 64 = 0). +Check (eq_refl 0 <: 9223372036854775807 >> 64 = 0). +Check (eq_refl 0 <<: 9223372036854775807 >> 64 = 0). +Definition compute3 := Eval compute in 9223372036854775807 >> 64. +Check (eq_refl compute3 : 0 = 0). diff --git a/test-suite/arithmetic/ltb.v b/test-suite/arithmetic/ltb.v new file mode 100644 index 0000000000..7ae5ac6493 --- /dev/null +++ b/test-suite/arithmetic/ltb.v @@ -0,0 +1,23 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 1 < 1 = false). +Check (eq_refl false <: 1 < 1 = false). +Check (eq_refl false <<: 1 < 1 = false). +Definition compute1 := Eval compute in 1 < 1. +Check (eq_refl compute1 : false = false). + +Check (eq_refl : 1 < 2 = true). +Check (eq_refl true <: 1 < 2 = true). +Check (eq_refl true <<: 1 < 2 = true). +Definition compute2 := Eval compute in 1 < 2. +Check (eq_refl compute2 : true = true). + +Check (eq_refl : 9223372036854775807 < 0 = false). +Check (eq_refl false <: 9223372036854775807 < 0 = false). +Check (eq_refl false <<: 9223372036854775807 < 0 = false). +Definition compute3 := Eval compute in 9223372036854775807 < 0. +Check (eq_refl compute3 : false = false). diff --git a/test-suite/arithmetic/lxor.v b/test-suite/arithmetic/lxor.v new file mode 100644 index 0000000000..b453fc7697 --- /dev/null +++ b/test-suite/arithmetic/lxor.v @@ -0,0 +1,29 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 0 lxor 0 = 0). +Check (eq_refl 0 <: 0 lxor 0 = 0). +Check (eq_refl 0 <<: 0 lxor 0 = 0). +Definition compute1 := Eval compute in 0 lxor 0. +Check (eq_refl compute1 : 0 = 0). + +Check (eq_refl : 9223372036854775807 lxor 0 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <: 9223372036854775807 lxor 0 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <<: 9223372036854775807 lxor 0 = 9223372036854775807). +Definition compute2 := Eval compute in 9223372036854775807 lxor 0. +Check (eq_refl compute2 : 9223372036854775807 = 9223372036854775807). + +Check (eq_refl : 0 lxor 9223372036854775807 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <: 0 lxor 9223372036854775807 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <<: 0 lxor 9223372036854775807 = 9223372036854775807). +Definition compute3 := Eval compute in 0 lxor 9223372036854775807. +Check (eq_refl compute3 : 9223372036854775807 = 9223372036854775807). + +Check (eq_refl : 9223372036854775807 lxor 9223372036854775807 = 0). +Check (eq_refl 0 <: 9223372036854775807 lxor 9223372036854775807 = 0). +Check (eq_refl 0 <<: 9223372036854775807 lxor 9223372036854775807 = 0). +Definition compute4 := Eval compute in 9223372036854775807 lxor 9223372036854775807. +Check (eq_refl compute4 : 0 = 0). diff --git a/test-suite/arithmetic/mod.v b/test-suite/arithmetic/mod.v new file mode 100644 index 0000000000..5307eed493 --- /dev/null +++ b/test-suite/arithmetic/mod.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 6 \% 3 = 0). +Check (eq_refl 0 <: 6 \% 3 = 0). +Check (eq_refl 0 <<: 6 \% 3 = 0). +Definition compute1 := Eval compute in 6 \% 3. +Check (eq_refl compute1 : 0 = 0). + +Check (eq_refl : 5 \% 3 = 2). +Check (eq_refl 2 <: 5 \% 3 = 2). +Check (eq_refl 2 <<: 5 \% 3 = 2). +Definition compute2 := Eval compute in 5 \% 3. +Check (eq_refl compute2 : 2 = 2). diff --git a/test-suite/arithmetic/mul.v b/test-suite/arithmetic/mul.v new file mode 100644 index 0000000000..9480e8cd46 --- /dev/null +++ b/test-suite/arithmetic/mul.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 2 * 3 = 6). +Check (eq_refl 6 <: 2 * 3 = 6). +Check (eq_refl 6 <<: 2 * 3 = 6). +Definition compute1 := Eval compute in 2 * 3. +Check (eq_refl compute1 : 6 = 6). + +Check (eq_refl : 9223372036854775807 * 2 = 9223372036854775806). +Check (eq_refl 9223372036854775806 <: 9223372036854775807 * 2 = 9223372036854775806). +Check (eq_refl 9223372036854775806 <<: 9223372036854775807 * 2 = 9223372036854775806). +Definition compute2 := Eval compute in 9223372036854775807 * 2. +Check (eq_refl compute2 : 9223372036854775806 = 9223372036854775806). diff --git a/test-suite/arithmetic/mulc.v b/test-suite/arithmetic/mulc.v new file mode 100644 index 0000000000..e10855bafa --- /dev/null +++ b/test-suite/arithmetic/mulc.v @@ -0,0 +1,22 @@ +Require Import Cyclic63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 2 *c 3 = WW 0 6). +Check (eq_refl (WW 0 6) <: 2 *c 3 = WW 0 6). +Check (eq_refl (WW 0 6) <<: 2 *c 3 = WW 0 6). +Definition compute1 := Eval compute in 2 *c 3. +Check (eq_refl compute1 : WW 0 6 = WW 0 6). + +Check (eq_refl : 9223372036854775807 *c 2 = WW 1 9223372036854775806). +Check (eq_refl (WW 1 9223372036854775806) <: 9223372036854775807 *c 2 = WW 1 9223372036854775806). +Check (eq_refl (WW 1 9223372036854775806) <<: 9223372036854775807 *c 2 = WW 1 9223372036854775806). + +Definition compute2 := Eval compute in 9223372036854775807 *c 2. +Check (eq_refl compute2 : WW 1 9223372036854775806 = WW 1 9223372036854775806). + +Check (eq_refl : 0 *c 0 = W0). +Check (eq_refl (@W0 int) <: 0 *c 0 = W0). +Check (eq_refl (@W0 int) <<: 0 *c 0 = W0). diff --git a/test-suite/arithmetic/primitive.v b/test-suite/arithmetic/primitive.v new file mode 100644 index 0000000000..f62f6109e1 --- /dev/null +++ b/test-suite/arithmetic/primitive.v @@ -0,0 +1,12 @@ +Section S. + Variable A : Type. + Fail Primitive int : let x := A in Set := #int63_type. + Fail Primitive add := #int63_add. +End S. + +(* [Primitive] should be forbidden in sections, otherwise its type after cooking +will be incorrect: + +Check int. + +*) diff --git a/test-suite/arithmetic/reduction.v b/test-suite/arithmetic/reduction.v new file mode 100644 index 0000000000..00e067ac5a --- /dev/null +++ b/test-suite/arithmetic/reduction.v @@ -0,0 +1,28 @@ +Require Import Int63. + +Open Scope int63_scope. + +Definition div_eucl_plus_one i1 i2 := + let (q,r) := diveucl i1 i2 in + (q+1, r+1)%int63. + +Definition rcbn := Eval cbn in div_eucl_plus_one 3 2. +Check (eq_refl : rcbn = (2, 2)). + +Definition rcbv := Eval cbv in div_eucl_plus_one 3 2. +Check (eq_refl : rcbv = (2, 2)). + +Definition rvmc := Eval vm_compute in div_eucl_plus_one 3 2. +Check (eq_refl : rvmc = (2, 2)). + +Definition f n m := + match (n ?= 42)%int63 with + | Lt => (n + m)%int63 + | _ => (2*m)%int63 + end. + +Goal forall n, (n ?= 42)%int63 = Gt -> f n 256 = 512%int63. + intros. unfold f. + cbn. Undo. cbv. (* Test reductions under match clauses *) + rewrite H. reflexivity. +Qed. diff --git a/test-suite/arithmetic/sub.v b/test-suite/arithmetic/sub.v new file mode 100644 index 0000000000..1606fd2aa1 --- /dev/null +++ b/test-suite/arithmetic/sub.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 3 - 2 = 1). +Check (eq_refl 1 <: 3 - 2 = 1). +Check (eq_refl 1 <<: 3 - 2 = 1). +Definition compute1 := Eval compute in 3 - 2. +Check (eq_refl compute1 : 1 = 1). + +Check (eq_refl : 0 - 1 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <: 0 - 1 = 9223372036854775807). +Check (eq_refl 9223372036854775807 <<: 0 - 1 = 9223372036854775807). +Definition compute2 := Eval compute in 0 - 1. +Check (eq_refl compute2 : 9223372036854775807 = 9223372036854775807). diff --git a/test-suite/arithmetic/subc.v b/test-suite/arithmetic/subc.v new file mode 100644 index 0000000000..fc4067e48b --- /dev/null +++ b/test-suite/arithmetic/subc.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : 3 -c 2 = C0 1). +Check (eq_refl (C0 1) <: 3 -c 2 = C0 1). +Check (eq_refl (C0 1) <<: 3 -c 2 = C0 1). +Definition compute1 := Eval compute in 3 -c 2. +Check (eq_refl compute1 : C0 1 = C0 1). + +Check (eq_refl : 0 -c 1 = C1 9223372036854775807). +Check (eq_refl (C1 9223372036854775807) <: 0 -c 1 = C1 9223372036854775807). +Check (eq_refl (C1 9223372036854775807) <<: 0 -c 1 = C1 9223372036854775807). +Definition compute2 := Eval compute in 0 -c 1. +Check (eq_refl compute2 : C1 9223372036854775807 = C1 9223372036854775807). diff --git a/test-suite/arithmetic/subcarryc.v b/test-suite/arithmetic/subcarryc.v new file mode 100644 index 0000000000..e81b6536b2 --- /dev/null +++ b/test-suite/arithmetic/subcarryc.v @@ -0,0 +1,17 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : subcarryc 3 1 = C0 1). +Check (eq_refl (C0 1) <: subcarryc 3 1 = C0 1). +Check (eq_refl (C0 1) <<: subcarryc 3 1 = C0 1). +Definition compute1 := Eval compute in subcarryc 3 1. +Check (eq_refl compute1 : C0 1 = C0 1). + +Check (eq_refl : subcarryc 0 1 = C1 9223372036854775806). +Check (eq_refl (C1 9223372036854775806) <: subcarryc 0 1 = C1 9223372036854775806). +Check (eq_refl (C1 9223372036854775806) <<: subcarryc 0 1 = C1 9223372036854775806). +Definition compute2 := Eval compute in subcarryc 0 1. +Check (eq_refl compute2 : C1 9223372036854775806 = C1 9223372036854775806). diff --git a/test-suite/arithmetic/tail0.v b/test-suite/arithmetic/tail0.v new file mode 100644 index 0000000000..c9d426087a --- /dev/null +++ b/test-suite/arithmetic/tail0.v @@ -0,0 +1,23 @@ +Require Import Int63. + +Set Implicit Arguments. + +Open Scope int63_scope. + +Check (eq_refl : tail0 2305843009213693952 = 61). +Check (eq_refl 61 <: tail0 2305843009213693952 = 61). +Check (eq_refl 61 <<: tail0 2305843009213693952 = 61). +Definition compute1 := Eval compute in tail0 2305843009213693952. +Check (eq_refl compute1 : 61 = 61). + +Check (eq_refl : tail0 1 = 0). +Check (eq_refl 0 <: tail0 1 = 0). +Check (eq_refl 0 <<: tail0 1 = 0). +Definition compute2 := Eval compute in tail0 1. +Check (eq_refl compute2 : 0 = 0). + +Check (eq_refl : tail0 0 = 63). +Check (eq_refl 63 <: tail0 0 = 63). +Check (eq_refl 63 <<: tail0 0 = 63). +Definition compute3 := Eval compute in tail0 0. +Check (eq_refl compute3 : 63 = 63). diff --git a/test-suite/arithmetic/unsigned.v b/test-suite/arithmetic/unsigned.v new file mode 100644 index 0000000000..82920bd201 --- /dev/null +++ b/test-suite/arithmetic/unsigned.v @@ -0,0 +1,18 @@ +(* This file checks that operations over int63 are unsigned. *) +Require Import Int63. + +Open Scope int63_scope. + +(* (0-1) must be the maximum integer value and not negative 1 *) + +Check (eq_refl : 1/(0-1) = 0). +Check (eq_refl 0 <: 1/(0-1) = 0). +Check (eq_refl 0 <<: 1/(0-1) = 0). +Definition compute1 := Eval compute in 1/(0-1). +Check (eq_refl compute1 : 0 = 0). + +Check (eq_refl : 3 \% (0-1) = 3). +Check (eq_refl 3 <: 3 \% (0-1) = 3). +Check (eq_refl 3 <<: 3 \% (0-1) = 3). +Definition compute2 := Eval compute in 3 \% (0-1). +Check (eq_refl compute2 : 3 = 3). diff --git a/test-suite/failure/int31.v b/test-suite/failure/int63.v index ed4c9f9c78..0bb87c6548 100644 --- a/test-suite/failure/int31.v +++ b/test-suite/failure/int63.v @@ -1,17 +1,16 @@ -Require Import Int31 Cyclic31. +Require Import Int63 Cyclic63. -Open Scope int31_scope. +Open Scope int63_scope. (* This used to go through because of an unbalanced stack bug in the bytecode interpreter *) Lemma bad : False. assert (1 = 2). -change 1 with (add31 (addmuldiv31 65 (add31 1 1) 2) 1). +change 1 with (Int63.add (Int63.addmuldiv 129 (Int63.add 1 1) 2) 1). Fail vm_compute; reflexivity. (* discriminate. Qed. *) Abort. - diff --git a/test-suite/output/Int31Syntax.out b/test-suite/output/Int31Syntax.out index 4e8796c14b..0d6504f5f8 100644 --- a/test-suite/output/Int31Syntax.out +++ b/test-suite/output/Int31Syntax.out @@ -12,3 +12,5 @@ I31 : int31 = 710436486 : int31 +The command has indeed failed with message: +Cannot interpret this number as a value of type int31 diff --git a/test-suite/output/Int31Syntax.v b/test-suite/output/Int31Syntax.v index 83be3b976b..48889a26ef 100644 --- a/test-suite/output/Int31Syntax.v +++ b/test-suite/output/Int31Syntax.v @@ -3,11 +3,12 @@ Require Import Int31 Cyclic31. Open Scope int31_scope. Check I31. (* Would be nice to have I31 : digits->digits->...->int31 For the moment, I31 : digits31 int31, which is better - than (fix nfun .....) size int31 *) + than (fix nfun .....) size int31 *) Check 2. Check 1000000000000000000. (* = 660865024, after modulo 2^31 *) Check (add31 2 2). Check (2+2). Eval vm_compute in 2+2. Eval vm_compute in 65675757 * 565675998. +Fail Check -1. Close Scope int31_scope. diff --git a/test-suite/output/Int63Syntax.out b/test-suite/output/Int63Syntax.out new file mode 100644 index 0000000000..fdd5599565 --- /dev/null +++ b/test-suite/output/Int63Syntax.out @@ -0,0 +1,16 @@ +2 + : int +9223372036854775807 + : int +2 + 2 + : int +2 + 2 + : int + = 4 + : int + = 37151199385380486 + : int +The command has indeed failed with message: +int63 are only non-negative numbers. +The command has indeed failed with message: +overflow in int63 literal: 9223372036854775808 diff --git a/test-suite/output/Int63Syntax.v b/test-suite/output/Int63Syntax.v new file mode 100644 index 0000000000..3dc364eddb --- /dev/null +++ b/test-suite/output/Int63Syntax.v @@ -0,0 +1,12 @@ +Require Import Int63 Cyclic63. + +Open Scope int63_scope. +Check 2. +Check 9223372036854775807. +Check (Int63.add 2 2). +Check (2+2). +Eval vm_compute in 2+2. +Eval vm_compute in 65675757 * 565675998. +Fail Check -1. +Fail Check 9223372036854775808. +Close Scope int63_scope. diff --git a/test-suite/output/Notations4.out b/test-suite/output/Notations4.out index 7a64b7eb45..efa895d709 100644 --- a/test-suite/output/Notations4.out +++ b/test-suite/output/Notations4.out @@ -51,3 +51,5 @@ r 2 3 : Prop Notation Cn := Foo.FooCn Expands to: Notation Notations4.J.Mfoo.Foo.Bar.Cn +let v := 0%test17 in v : myint63 + : myint63 diff --git a/test-suite/output/Notations4.v b/test-suite/output/Notations4.v index 90babf9c55..b4c65ce196 100644 --- a/test-suite/output/Notations4.v +++ b/test-suite/output/Notations4.v @@ -197,3 +197,16 @@ End Mfoo. About Cn. End J. + +Require Import Coq.Numbers.Cyclic.Int63.Int63. +Module NumeralNotations. + Module Test17. + (** Test int63 *) + Declare Scope test17_scope. + Delimit Scope test17_scope with test17. + Local Set Primitive Projections. + Record myint63 := of_int { to_int : int }. + Numeral Notation myint63 of_int to_int : test17_scope. + Check let v := 0%test17 in v : myint63. + End Test17. +End NumeralNotations. diff --git a/test-suite/output/sint63Notation.out b/test-suite/output/sint63Notation.out new file mode 100644 index 0000000000..9d325b38c7 --- /dev/null +++ b/test-suite/output/sint63Notation.out @@ -0,0 +1,24 @@ + = 0 + : uint + = 1 + : uint + = 9223372036854775807 + : uint +let v := 0 in v : uint + : uint +let v := 1 in v : uint + : uint +let v := 9223372036854775807 in v : uint + : uint + = 0 + : sint + = 1 + : sint + = -1 + : sint +let v := 0 in v : sint + : sint +let v := 1 in v : sint + : sint +let v := -1 in v : sint + : sint diff --git a/test-suite/output/sint63Notation.v b/test-suite/output/sint63Notation.v new file mode 100644 index 0000000000..331d74ed3d --- /dev/null +++ b/test-suite/output/sint63Notation.v @@ -0,0 +1,37 @@ +From Coq +Require Import Int63. +Import ZArith. + +Declare Scope uint_scope. +Declare Scope sint_scope. +Delimit Scope uint_scope with uint. +Delimit Scope sint_scope with sint. + +Record uint := wrapu { unwrapu : int }. +Record sint := wraps { unwraps : int }. + +Definition uof_Z (v : Z) := wrapu (of_Z v). +Definition uto_Z (v : uint) := to_Z (unwrapu v). + +Definition sof_Z (v : Z) := wraps (of_Z (v mod (2 ^ 31))). +Definition as_signed (bw : Z) (v : Z) := + (((2 ^ (bw - 1) + v) mod (2 ^ bw)) - 2 ^ (bw - 1))%Z. + +Definition sto_Z (v : sint) := as_signed 31 (to_Z (unwraps v)). +Numeral Notation uint uof_Z uto_Z : uint_scope. +Numeral Notation sint sof_Z sto_Z : sint_scope. +Open Scope uint_scope. +Compute uof_Z 0. +Compute uof_Z 1. +Compute uof_Z (-1). +Check let v := 0 in v : uint. +Check let v := 1 in v : uint. +Check let v := -1 in v : uint. +Close Scope uint_scope. +Open Scope sint_scope. +Compute sof_Z 0. +Compute sof_Z 1. +Compute sof_Z (-1). +Check let v := 0 in v : sint. +Check let v := 1 in v : sint. +Check let v := -1 in v : sint. diff --git a/test-suite/success/Compat88.v b/test-suite/success/Compat88.v index e2045900d5..1233a4b8f5 100644 --- a/test-suite/success/Compat88.v +++ b/test-suite/success/Compat88.v @@ -4,7 +4,7 @@ Require Coq.Strings.Ascii Coq.Strings.String. Require Coq.ZArith.BinIntDef Coq.PArith.BinPosDef Coq.NArith.BinNatDef. Require Coq.Reals.Rdefinitions. -Require Coq.Numbers.Cyclic.Int31.Cyclic31. +Require Coq.Numbers.Cyclic.Int63.Cyclic63. Require Import Coq.Compat.Coq88. (* XXX FIXME Should not need [Require], see https://github.com/coq/coq/issues/8311 *) @@ -15,4 +15,4 @@ Check BinNat.N.eqb 1 1. Check BinInt.Z.eqb 1 1. Check BinPos.Pos.eqb 1 1. Check Rdefinitions.Rplus 1 1. -Check Int31.iszero 1. +Check Int63.is_zero 1. diff --git a/test-suite/success/NumeralNotations.v b/test-suite/success/NumeralNotations.v index 47ef381270..7b857c70c5 100644 --- a/test-suite/success/NumeralNotations.v +++ b/test-suite/success/NumeralNotations.v @@ -300,3 +300,14 @@ Module Test16. (** Ideally this should work, but it should definitely not anomaly *) Fail Check let v := 0%test16 in v : Foo. End Test16. + +Require Import Coq.Numbers.Cyclic.Int63.Int63. +Module Test17. + (** Test int63 *) + Declare Scope test17_scope. + Delimit Scope test17_scope with test17. + Local Set Primitive Projections. + Record myint63 := of_int { to_int : int }. + Numeral Notation myint63 of_int to_int : test17_scope. + Check let v := 0%test17 in v : myint63. +End Test17. diff --git a/theories/Compat/Coq88.v b/theories/Compat/Coq88.v index 989072940a..e4a8df1e93 100644 --- a/theories/Compat/Coq88.v +++ b/theories/Compat/Coq88.v @@ -20,9 +20,11 @@ Require Coq.Strings.Ascii Coq.Strings.String. Export String.StringSyntax Ascii.AsciiSyntax. Require Coq.ZArith.BinIntDef Coq.PArith.BinPosDef Coq.NArith.BinNatDef. Require Coq.Reals.Rdefinitions. +Require Coq.Numbers.Cyclic.Int63.Int63. Require Coq.Numbers.Cyclic.Int31.Int31. Declare ML Module "r_syntax_plugin". -Declare ML Module "int31_syntax_plugin". +Declare ML Module "int63_syntax_plugin". Numeral Notation BinNums.Z BinIntDef.Z.of_int BinIntDef.Z.to_int : Z_scope. Numeral Notation BinNums.positive BinPosDef.Pos.of_int BinPosDef.Pos.to_int : positive_scope. Numeral Notation BinNums.N BinNatDef.N.of_int BinNatDef.N.to_int : N_scope. +Numeral Notation Int31.int31 Int31.phi_inv_nonneg Int31.phi : int31_scope. diff --git a/theories/Numbers/Cyclic/Abstract/DoubleType.v b/theories/Numbers/Cyclic/Abstract/DoubleType.v index b6441bb76a..9547a642df 100644 --- a/theories/Numbers/Cyclic/Abstract/DoubleType.v +++ b/theories/Numbers/Cyclic/Abstract/DoubleType.v @@ -12,7 +12,7 @@ Set Implicit Arguments. -Require Import ZArith. +Require Import BinInt. Local Open Scope Z_scope. Definition base digits := Z.pow 2 (Zpos digits). @@ -23,7 +23,7 @@ Section Carry. Variable A : Type. #[universes(template)] - Inductive carry := + Variant carry := | C0 : A -> carry | C1 : A -> carry. @@ -46,7 +46,7 @@ Section Zn2Z. *) #[universes(template)] - Inductive zn2z := + Variant zn2z := | W0 : zn2z | WW : znz -> znz -> zn2z. diff --git a/theories/Numbers/Cyclic/Int31/Cyclic31.v b/theories/Numbers/Cyclic/Int31/Cyclic31.v index 4a1f24b95e..4b0bda3d44 100644 --- a/theories/Numbers/Cyclic/Int31/Cyclic31.v +++ b/theories/Numbers/Cyclic/Int31/Cyclic31.v @@ -8,6 +8,8 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) +(** This library has been deprecated since Coq version 8.10. *) + (** * Int31 numbers defines indeed a cyclic structure : Z/(2^31)Z *) (** @@ -1274,7 +1276,7 @@ Section Int31_Specs. Qed. Lemma spec_add_carry : - forall x y, [|x+y+1|] = ([|x|] + [|y|] + 1) mod wB. + forall x y, [|x+y+1|] = ([|x|] + [|y|] + 1) mod wB. Proof. unfold add31; intros. repeat rewrite phi_phi_inv. @@ -1776,7 +1778,7 @@ Section Int31_Specs. Qed. Lemma spec_head0 : forall x, 0 < [|x|] -> - wB/ 2 <= 2 ^ ([|head031 x|]) * [|x|] < wB. + wB/ 2 <= 2 ^ ([|head031 x|]) * [|x|] < wB. Proof. intros. rewrite head031_equiv. diff --git a/theories/Numbers/Cyclic/Int31/Int31.v b/theories/Numbers/Cyclic/Int31/Int31.v index ce540775e3..b9185c9ca0 100644 --- a/theories/Numbers/Cyclic/Int31/Int31.v +++ b/theories/Numbers/Cyclic/Int31/Int31.v @@ -10,6 +10,8 @@ (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) +(** This library has been deprecated since Coq version 8.10. *) + Require Import NaryFunctions. Require Import Wf_nat. Require Export ZArith. @@ -44,18 +46,11 @@ Definition digits31 t := Eval compute in nfun digits size t. Inductive int31 : Type := I31 : digits31 int31. -(* spiwack: Registration of the type of integers, so that the matchs in - the functions below perform dynamic decompilation (otherwise some segfault - occur when they are applied to one non-closed term and one closed term). *) -Register digits as int31.bits. -Register int31 as int31.type. - Scheme int31_ind := Induction for int31 Sort Prop. Scheme int31_rec := Induction for int31 Sort Set. Scheme int31_rect := Induction for int31 Sort Type. Declare Scope int31_scope. -Declare ML Module "int31_syntax_plugin". Delimit Scope int31_scope with int31. Bind Scope int31_scope with int31. Local Open Scope int31_scope. @@ -208,6 +203,13 @@ Definition phi_inv : Z -> int31 := fun n => | Zneg p => incr (complement_negative p) end. +(** [phi_inv_nonneg] returns [None] if the [Z] is negative; this matches the old behavior of parsing int31 numerals *) +Definition phi_inv_nonneg : Z -> option int31 := fun n => + match n with + | Zneg _ => None + | _ => Some (phi_inv n) + end. + (** [phi_inv2] is similar to [phi_inv] but returns a double word [zn2z int31] *) @@ -351,22 +353,6 @@ Definition lor31 n m := phi_inv (Z.lor (phi n) (phi m)). Definition land31 n m := phi_inv (Z.land (phi n) (phi m)). Definition lxor31 n m := phi_inv (Z.lxor (phi n) (phi m)). -Register add31 as int31.plus. -Register add31c as int31.plusc. -Register add31carryc as int31.pluscarryc. -Register sub31 as int31.minus. -Register sub31c as int31.minusc. -Register sub31carryc as int31.minuscarryc. -Register mul31 as int31.times. -Register mul31c as int31.timesc. -Register div3121 as int31.div21. -Register div31 as int31.diveucl. -Register compare31 as int31.compare. -Register addmuldiv31 as int31.addmuldiv. -Register lor31 as int31.lor. -Register land31 as int31.land. -Register lxor31 as int31.lxor. - Definition lnot31 n := lxor31 Tn n. Definition ldiff31 n m := land31 n (lnot31 m). @@ -491,5 +477,4 @@ Definition tail031 (i:int31) := end) i On. -Register head031 as int31.head0. -Register tail031 as int31.tail0. +Numeral Notation int31 phi_inv_nonneg phi : int31_scope. diff --git a/theories/Numbers/Cyclic/Int31/Ring31.v b/theories/Numbers/Cyclic/Int31/Ring31.v index b693529451..eb47141cab 100644 --- a/theories/Numbers/Cyclic/Int31/Ring31.v +++ b/theories/Numbers/Cyclic/Int31/Ring31.v @@ -8,6 +8,8 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) +(** This library has been deprecated since Coq version 8.10. *) + (** * Int31 numbers defines Z/(2^31)Z, and can hence be equipped with a ring structure and a ring tactic *) @@ -101,4 +103,3 @@ Let test : forall x y, 1 + x*y + x*x + 1 = 1*1 + 1 + y*x + 1*x*x. intros. ring. Qed. End TestRing. - diff --git a/theories/Numbers/Cyclic/Int63/Cyclic63.v b/theories/Numbers/Cyclic/Int63/Cyclic63.v new file mode 100644 index 0000000000..3b431d5b47 --- /dev/null +++ b/theories/Numbers/Cyclic/Int63/Cyclic63.v @@ -0,0 +1,330 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(** * Int63 numbers defines indeed a cyclic structure : Z/(2^31)Z *) + +(** +Author: Arnaud Spiwack (+ Pierre Letouzey) +*) +Require Import CyclicAxioms. +Require Export ZArith. +Require Export Int63. +Import Zpow_facts. +Import Utf8. +Import Lia. + +Local Open Scope int63_scope. +(** {2 Operators } **) + +Definition Pdigits := Eval compute in P_of_succ_nat (size - 1). + +Fixpoint positive_to_int_rec (n:nat) (p:positive) := + match n, p with + | O, _ => (Npos p, 0) + | S n, xH => (0%N, 1) + | S n, xO p => + let (N,i) := positive_to_int_rec n p in + (N, i << 1) + | S n, xI p => + let (N,i) := positive_to_int_rec n p in + (N, (i << 1) + 1) + end. + +Definition positive_to_int := positive_to_int_rec size. + +Definition mulc_WW x y := + let (h, l) := mulc x y in + if is_zero h then + if is_zero l then W0 + else WW h l + else WW h l. +Notation "n '*c' m" := (mulc_WW n m) (at level 40, no associativity) : int63_scope. + +Definition pos_mod p x := + if p <= digits then + let p := digits - p in + (x << p) >> p + else x. + +Notation pos_mod_int := pos_mod. + +Import ZnZ. + +Instance int_ops : ZnZ.Ops int := +{| + digits := Pdigits; (* number of digits *) + zdigits := Int63.digits; (* number of digits *) + to_Z := Int63.to_Z; (* conversion to Z *) + of_pos := positive_to_int; (* positive -> N*int63 : p => N,i + where p = N*2^31+phi i *) + head0 := Int63.head0; (* number of head 0 *) + tail0 := Int63.tail0; (* number of tail 0 *) + zero := 0; + one := 1; + minus_one := Int63.max_int; + compare := Int63.compare; + eq0 := Int63.is_zero; + opp_c := Int63.oppc; + opp := Int63.opp; + opp_carry := Int63.oppcarry; + succ_c := Int63.succc; + add_c := Int63.addc; + add_carry_c := Int63.addcarryc; + succ := Int63.succ; + add := Int63.add; + add_carry := Int63.addcarry; + pred_c := Int63.predc; + sub_c := Int63.subc; + sub_carry_c := Int63.subcarryc; + pred := Int63.pred; + sub := Int63.sub; + sub_carry := Int63.subcarry; + mul_c := mulc_WW; + mul := Int63.mul; + square_c := fun x => mulc_WW x x; + div21 := diveucl_21; + div_gt := diveucl; (* this is supposed to be the special case of + division a/b where a > b *) + div := diveucl; + modulo_gt := Int63.mod; + modulo := Int63.mod; + gcd_gt := Int63.gcd; + gcd := Int63.gcd; + add_mul_div := Int63.addmuldiv; + pos_mod := pos_mod_int; + is_even := Int63.is_even; + sqrt2 := Int63.sqrt2; + sqrt := Int63.sqrt; + ZnZ.lor := Int63.lor; + ZnZ.land := Int63.land; + ZnZ.lxor := Int63.lxor +|}. + +Local Open Scope Z_scope. + +Lemma is_zero_spec_aux : forall x : int, is_zero x = true -> [|x|] = 0%Z. +Proof. + intros x;rewrite is_zero_spec;intros H;rewrite H;trivial. +Qed. + +Lemma positive_to_int_spec : + forall p : positive, + Zpos p = + Z_of_N (fst (positive_to_int p)) * wB + to_Z (snd (positive_to_int p)). +Proof. + assert (H: (wB <= wB) -> forall p : positive, + Zpos p = Z_of_N (fst (positive_to_int p)) * wB + [|snd (positive_to_int p)|] /\ + [|snd (positive_to_int p)|] < wB). + 2: intros p; case (H (Z.le_refl wB) p); auto. + unfold positive_to_int, wB at 1 3 4. + elim size. + intros _ p; simpl; + rewrite to_Z_0, Pmult_1_r; split; auto with zarith; apply refl_equal. + intros n; rewrite inj_S; unfold Z.succ; rewrite Zpower_exp, Z.pow_1_r; auto with zarith. + intros IH Hle p. + assert (F1: 2 ^ Z_of_nat n <= wB); auto with zarith. + assert (0 <= 2 ^ Z_of_nat n); auto with zarith. + case p; simpl. + intros p1. + generalize (IH F1 p1); case positive_to_int_rec; simpl. + intros n1 i (H1,H2). + rewrite Zpos_xI, H1. + replace [|i << 1 + 1|] with ([|i|] * 2 + 1). + split; auto with zarith; ring. + rewrite add_spec, lsl_spec, Zplus_mod_idemp_l, to_Z_1, Z.pow_1_r, Zmod_small; auto. + case (to_Z_bounded i); split; auto with zarith. + intros p1. + generalize (IH F1 p1); case positive_to_int_rec; simpl. + intros n1 i (H1,H2). + rewrite Zpos_xO, H1. + replace [|i << 1|] with ([|i|] * 2). + split; auto with zarith; ring. + rewrite lsl_spec, to_Z_1, Z.pow_1_r, Zmod_small; auto. + case (to_Z_bounded i); split; auto with zarith. + rewrite to_Z_1; assert (0 < 2^ Z_of_nat n); auto with zarith. +Qed. + +Lemma mulc_WW_spec : + forall x y,[|| x *c y ||] = [|x|] * [|y|]. +Proof. + intros x y;unfold mulc_WW. + generalize (mulc_spec x y);destruct (mulc x y);simpl;intros Heq;rewrite Heq. + case_eq (is_zero i);intros;trivial. + apply is_zero_spec in H;rewrite H, to_Z_0. + case_eq (is_zero i0);intros;trivial. + apply is_zero_spec in H0;rewrite H0, to_Z_0, Zmult_comm;trivial. +Qed. + +Lemma squarec_spec : + forall x, + [||x *c x||] = [|x|] * [|x|]. +Proof (fun x => mulc_WW_spec x x). + +Lemma diveucl_spec_aux : forall a b, 0 < [|b|] -> + let (q,r) := diveucl a b in + [|a|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. +Proof. + intros a b H;assert (W:= diveucl_spec a b). + assert ([|b|]>0) by (auto with zarith). + generalize (Z_div_mod [|a|] [|b|] H0). + destruct (diveucl a b);destruct (Z.div_eucl [|a|] [|b|]). + inversion W;rewrite Zmult_comm;trivial. +Qed. + +Lemma diveucl_21_spec_aux : forall a1 a2 b, + wB/2 <= [|b|] -> + [|a1|] < [|b|] -> + let (q,r) := diveucl_21 a1 a2 b in + [|a1|] *wB+ [|a2|] = [|q|] * [|b|] + [|r|] /\ + 0 <= [|r|] < [|b|]. +Proof. + intros a1 a2 b H1 H2;assert (W:= diveucl_21_spec a1 a2 b). + assert (W1:= to_Z_bounded a1). + assert ([|b|]>0) by (auto with zarith). + generalize (Z_div_mod ([|a1|]*wB+[|a2|]) [|b|] H). + destruct (diveucl_21 a1 a2 b);destruct (Z.div_eucl ([|a1|]*wB+[|a2|]) [|b|]). + inversion W;rewrite (Zmult_comm [|b|]);trivial. +Qed. + +Lemma shift_unshift_mod_2 : forall n p a, 0 <= p <= n -> + ((a * 2 ^ (n - p)) mod (2^n) / 2 ^ (n - p)) mod (2^n) = + a mod 2 ^ p. + Proof. + intros n p a H. + rewrite Zmod_small. + - rewrite Zmod_eq by auto with zarith. + unfold Zminus at 1. + rewrite Zdiv.Z_div_plus_full_l by auto with zarith. + replace (2 ^ n) with (2 ^ (n - p) * 2 ^ p) by (rewrite <- Zpower_exp; [ f_equal | | ]; lia). + rewrite <- Zdiv_Zdiv, Z_div_mult by auto with zarith. + rewrite (Zmult_comm (2^(n-p))), Zmult_assoc. + rewrite Zopp_mult_distr_l. + rewrite Z_div_mult by auto with zarith. + symmetry; apply Zmod_eq; auto with zarith. + - remember (a * 2 ^ (n - p)) as b. + destruct (Z_mod_lt b (2^n)); auto with zarith. + split. + apply Z_div_pos; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + apply Z.lt_le_trans with (2^n); auto with zarith. + generalize (pow2_pos (n - p)); nia. + Qed. + +Lemma div_le_0 : forall p x, 0 <= x -> 0 <= x / 2 ^ p. + Proof. + intros p x Hle;destruct (Z_le_gt_dec 0 p). + apply Zdiv_le_lower_bound;auto with zarith. + replace (2^p) with 0. + destruct x;compute;intro;discriminate. + destruct p;trivial;discriminate. + Qed. + +Lemma div_lt : forall p x y, 0 <= x < y -> x / 2^p < y. + Proof. + intros p x y H;destruct (Z_le_gt_dec 0 p). + apply Zdiv_lt_upper_bound;auto with zarith. + apply Z.lt_le_trans with y;auto with zarith. + rewrite <- (Zmult_1_r y);apply Zmult_le_compat;auto with zarith. + assert (0 < 2^p);auto with zarith. + replace (2^p) with 0. + destruct x;change (0<y);auto with zarith. + destruct p;trivial;discriminate. + Qed. + +Lemma P (A B C: Prop) : + A → (B → C) → (A → B) → C. +Proof. tauto. Qed. + +Lemma shift_unshift_mod_3: + forall n p a : Z, + 0 <= p <= n -> + (a * 2 ^ (n - p)) mod 2 ^ n / 2 ^ (n - p) = a mod 2 ^ p. +Proof. + intros;rewrite <- (shift_unshift_mod_2 n p a);[ | auto with zarith]. + symmetry;apply Zmod_small. + generalize (a * 2 ^ (n - p));intros w. + generalize (2 ^ (n - p)) (pow2_pos (n - p)); intros x; apply P. lia. intros hx. + generalize (2 ^ n) (pow2_pos n); intros y; apply P. lia. intros hy. + elim_div. intros q r. apply P. lia. + elim_div. intros z t. refine (P _ _ _ _ _). lia. + intros [ ? [ ht | ] ]; [ | lia ]; subst w. + intros [ ? [ hr | ] ]; [ | lia ]; subst t. + nia. +Qed. + +Lemma pos_mod_spec w p : φ(pos_mod p w) = φ(w) mod (2 ^ φ(p)). +Proof. + simpl. unfold pos_mod_int. + assert (W:=to_Z_bounded p);assert (W':=to_Z_bounded Int63.digits);assert (W'' := to_Z_bounded w). + case lebP; intros hle. + 2: { + symmetry; apply Zmod_small. + assert (2 ^ [|Int63.digits|] < 2 ^ [|p|]); [ apply Zpower_lt_monotone; auto with zarith | ]. + change wB with (2 ^ [|Int63.digits|]) in *; auto with zarith. } + rewrite <- (shift_unshift_mod_3 [|Int63.digits|] [|p|] [|w|]) by auto with zarith. + replace ([|Int63.digits|] - [|p|]) with [|Int63.digits - p|] by (rewrite sub_spec, Zmod_small; auto with zarith). + rewrite lsr_spec, lsl_spec; reflexivity. +Qed. + +(** {2 Specification and proof} **) +Global Instance int_specs : ZnZ.Specs int_ops := { + spec_to_Z := to_Z_bounded; + spec_of_pos := positive_to_int_spec; + spec_zdigits := refl_equal _; + spec_more_than_1_digit:= refl_equal _; + spec_0 := to_Z_0; + spec_1 := to_Z_1; + spec_m1 := refl_equal _; + spec_compare := compare_spec; + spec_eq0 := is_zero_spec_aux; + spec_opp_c := oppc_spec; + spec_opp := opp_spec; + spec_opp_carry := oppcarry_spec; + spec_succ_c := succc_spec; + spec_add_c := addc_spec; + spec_add_carry_c := addcarryc_spec; + spec_succ := succ_spec; + spec_add := add_spec; + spec_add_carry := addcarry_spec; + spec_pred_c := predc_spec; + spec_sub_c := subc_spec; + spec_sub_carry_c := subcarryc_spec; + spec_pred := pred_spec; + spec_sub := sub_spec; + spec_sub_carry := subcarry_spec; + spec_mul_c := mulc_WW_spec; + spec_mul := mul_spec; + spec_square_c := squarec_spec; + spec_div21 := diveucl_21_spec_aux; + spec_div_gt := fun a b _ => diveucl_spec_aux a b; + spec_div := diveucl_spec_aux; + spec_modulo_gt := fun a b _ _ => mod_spec a b; + spec_modulo := fun a b _ => mod_spec a b; + spec_gcd_gt := fun a b _ => gcd_spec a b; + spec_gcd := gcd_spec; + spec_head00 := head00_spec; + spec_head0 := head0_spec; + spec_tail00 := tail00_spec; + spec_tail0 := tail0_spec; + spec_add_mul_div := addmuldiv_spec; + spec_pos_mod := pos_mod_spec; + spec_is_even := is_even_spec; + spec_sqrt2 := sqrt2_spec; + spec_sqrt := sqrt_spec; + spec_land := land_spec'; + spec_lor := lor_spec'; + spec_lxor := lxor_spec' }. + + + +Module Int63Cyclic <: CyclicType. + Definition t := int. + Definition ops := int_ops. + Definition specs := int_specs. +End Int63Cyclic. diff --git a/theories/Numbers/Cyclic/Int63/Int63.v b/theories/Numbers/Cyclic/Int63/Int63.v new file mode 100644 index 0000000000..eac26add03 --- /dev/null +++ b/theories/Numbers/Cyclic/Int63/Int63.v @@ -0,0 +1,1918 @@ +(************************************************************************) +(* * 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) *) +(************************************************************************) +(* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) +(************************************************************************) + +Require Import Utf8. +Require Export DoubleType. +Require Import Lia. +Require Import Zpow_facts. +Require Import Zgcd_alt. +Import Znumtheory. + +Register bool as kernel.ind_bool. +Register prod as kernel.ind_pair. +Register carry as kernel.ind_carry. +Register comparison as kernel.ind_cmp. + +Definition size := 63%nat. + +Primitive int := #int63_type. +Register int as num.int63.type. +Declare Scope int63_scope. +Definition id_int : int -> int := fun x => x. +Declare ML Module "int63_syntax_plugin". + +Delimit Scope int63_scope with int63. +Bind Scope int63_scope with int. + +(* Logical operations *) +Primitive lsl := #int63_lsl. +Infix "<<" := lsl (at level 30, no associativity) : int63_scope. + +Primitive lsr := #int63_lsr. +Infix ">>" := lsr (at level 30, no associativity) : int63_scope. + +Primitive land := #int63_land. +Infix "land" := land (at level 40, left associativity) : int63_scope. + +Primitive lor := #int63_lor. +Infix "lor" := lor (at level 40, left associativity) : int63_scope. + +Primitive lxor := #int63_lxor. +Infix "lxor" := lxor (at level 40, left associativity) : int63_scope. + +(* Arithmetic modulo operations *) +Primitive add := #int63_add. +Notation "n + m" := (add n m) : int63_scope. + +Primitive sub := #int63_sub. +Notation "n - m" := (sub n m) : int63_scope. + +Primitive mul := #int63_mul. +Notation "n * m" := (mul n m) : int63_scope. + +Primitive mulc := #int63_mulc. + +Primitive div := #int63_div. +Notation "n / m" := (div n m) : int63_scope. + +Primitive mod := #int63_mod. +Notation "n '\%' m" := (mod n m) (at level 40, left associativity) : int63_scope. + +(* Comparisons *) +Primitive eqb := #int63_eq. +Notation "m '==' n" := (eqb m n) (at level 70, no associativity) : int63_scope. + +Primitive ltb := #int63_lt. +Notation "m < n" := (ltb m n) : int63_scope. + +Primitive leb := #int63_le. +Notation "m <= n" := (leb m n) : int63_scope. +Notation "m ≤ n" := (leb m n) (at level 70, no associativity) : int63_scope. + +Local Open Scope int63_scope. + +(** The number of digits as a int *) +Definition digits := 63. + +(** The bigger int *) +Definition max_int := Eval vm_compute in 0 - 1. +Register Inline max_int. + +(** Access to the nth digits *) +Definition get_digit x p := (0 < (x land (1 << p))). + +Definition set_digit x p (b:bool) := + if if 0 <= p then p < digits else false then + if b then x lor (1 << p) + else x land (max_int lxor (1 << p)) + else x. + +(** Equality to 0 *) +Definition is_zero (i:int) := i == 0. +Register Inline is_zero. + +(** Parity *) +Definition is_even (i:int) := is_zero (i land 1). +Register Inline is_even. + +(** Bit *) + +Definition bit i n := negb (is_zero ((i >> n) << (digits - 1))). +(* Register bit as PrimInline. *) + +(** Extra modulo operations *) +Definition opp (i:int) := 0 - i. +Register Inline opp. +Notation "- x" := (opp x) : int63_scope. + +Definition oppcarry i := max_int - i. +Register Inline oppcarry. + +Definition succ i := i + 1. +Register Inline succ. + +Definition pred i := i - 1. +Register Inline pred. + +Definition addcarry i j := i + j + 1. +Register Inline addcarry. + +Definition subcarry i j := i - j - 1. +Register Inline subcarry. + +(** Exact arithmetic operations *) + +Definition addc_def x y := + let r := x + y in + if r < x then C1 r else C0 r. +(* the same but direct implementation for effeciancy *) +Primitive addc := #int63_addc. +Notation "n '+c' m" := (addc n m) (at level 50, no associativity) : int63_scope. + +Definition addcarryc_def x y := + let r := addcarry x y in + if r <= x then C1 r else C0 r. +(* the same but direct implementation for effeciancy *) +Primitive addcarryc := #int63_addcarryc. + +Definition subc_def x y := + if y <= x then C0 (x - y) else C1 (x - y). +(* the same but direct implementation for effeciancy *) +Primitive subc := #int63_subc. +Notation "n '-c' m" := (subc n m) (at level 50, no associativity) : int63_scope. + +Definition subcarryc_def x y := + if y < x then C0 (x - y - 1) else C1 (x - y - 1). +(* the same but direct implementation for effeciancy *) +Primitive subcarryc := #int63_subcarryc. + +Definition diveucl_def x y := (x/y, x\%y). +(* the same but direct implementation for effeciancy *) +Primitive diveucl := #int63_diveucl. + +Primitive diveucl_21 := #int63_div21. + +Definition addmuldiv_def p x y := + (x << p) lor (y >> (digits - p)). +Primitive addmuldiv := #int63_addmuldiv. + +Definition oppc (i:int) := 0 -c i. +Register Inline oppc. + +Definition succc i := i +c 1. +Register Inline succc. + +Definition predc i := i -c 1. +Register Inline predc. + +(** Comparison *) +Definition compare_def x y := + if x < y then Lt + else if (x == y) then Eq else Gt. + +Primitive compare := #int63_compare. +Notation "n ?= m" := (compare n m) (at level 70, no associativity) : int63_scope. + +Import Bool ZArith. +(** Translation to Z *) +Fixpoint to_Z_rec (n:nat) (i:int) := + match n with + | O => 0%Z + | S n => + (if is_even i then Z.double else Zdouble_plus_one) (to_Z_rec n (i >> 1)) + end. + +Definition to_Z := to_Z_rec size. + +Fixpoint of_pos_rec (n:nat) (p:positive) := + match n, p with + | O, _ => 0 + | S n, xH => 1 + | S n, xO p => (of_pos_rec n p) << 1 + | S n, xI p => (of_pos_rec n p) << 1 lor 1 + end. + +Definition of_pos := of_pos_rec size. + +Definition of_Z z := + match z with + | Zpos p => of_pos p + | Z0 => 0 + | Zneg p => - (of_pos p) + end. + +Notation "[| x |]" := (to_Z x) (at level 0, x at level 99) : int63_scope. + +Definition wB := (2 ^ (Z.of_nat size))%Z. + +Lemma to_Z_rec_bounded size : forall x, (0 <= to_Z_rec size x < 2 ^ Z.of_nat size)%Z. +Proof. + elim size. simpl; auto with zarith. + intros n ih x; rewrite inj_S; simpl; assert (W := ih (x >> 1)%int63). + rewrite Z.pow_succ_r; auto with zarith. + destruct (is_even x). + rewrite Zdouble_mult; auto with zarith. + rewrite Zdouble_plus_one_mult; auto with zarith. +Qed. + +Corollary to_Z_bounded : forall x, (0 <= [| x |] < wB)%Z. +Proof. apply to_Z_rec_bounded. Qed. + +(* =================================================== *) +Local Open Scope Z_scope. +(* General arithmetic results *) +Lemma Z_lt_div2 x y : x < 2 * y -> x / 2 < y. +Proof. apply Z.div_lt_upper_bound; reflexivity. Qed. + +Theorem Zmod_le_first a b : 0 <= a -> 0 < b -> 0 <= a mod b <= a. +Proof. + intros ha hb; case (Z_mod_lt a b); [ auto with zarith | ]; intros p q; apply (conj p). + case (Z.le_gt_cases b a). lia. + intros hlt; rewrite Zmod_small; lia. +Qed. + +Theorem Zmod_distr: forall a b r t, 0 <= a <= b -> 0 <= r -> 0 <= t < 2 ^a -> + (2 ^a * r + t) mod (2 ^ b) = (2 ^a * r) mod (2 ^ b) + t. +Proof. + intros a b r t (H1, H2) H3 (H4, H5). + assert (t < 2 ^ b). + apply Z.lt_le_trans with (1:= H5); auto with zarith. + apply Zpower_le_monotone; auto with zarith. + rewrite Zplus_mod; auto with zarith. + rewrite -> Zmod_small with (a := t); auto with zarith. + apply Zmod_small; auto with zarith. + split; auto with zarith. + assert (0 <= 2 ^a * r); auto with zarith. + apply Z.add_nonneg_nonneg; auto with zarith. + match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end; + auto with zarith. + pattern (2 ^ b) at 2; replace (2 ^ b) with ((2 ^ b - 2 ^a) + 2 ^ a); + try ring. + apply Z.add_le_lt_mono; auto with zarith. + replace b with ((b - a) + a); try ring. + rewrite Zpower_exp; auto with zarith. + pattern (2 ^a) at 4; rewrite <- (Z.mul_1_l (2 ^a)); + try rewrite <- Z.mul_sub_distr_r. + rewrite (Z.mul_comm (2 ^(b - a))); rewrite Zmult_mod_distr_l; + auto with zarith. + rewrite (Z.mul_comm (2 ^a)); apply Z.mul_le_mono_nonneg_r; auto with zarith. + match goal with |- context [?X mod ?Y] => case (Z_mod_lt X Y) end. + apply Z.lt_gt; auto with zarith. + auto with zarith. +Qed. + +(* Results about pow2 *) +Lemma pow2_pos n : 0 <= n → 2 ^ n > 0. +Proof. intros h; apply Z.lt_gt, Zpower_gt_0; lia. Qed. + +Lemma pow2_nz n : 0 <= n → 2 ^ n ≠ 0. +Proof. intros h; generalize (pow2_pos _ h); lia. Qed. + +Hint Resolve pow2_pos pow2_nz : zarith. + +(* =================================================== *) + +(** Trivial lemmas without axiom *) + +Lemma wB_diff_0 : wB <> 0. +Proof. exact (fun x => let 'eq_refl := x in idProp). Qed. + +Lemma wB_pos : 0 < wB. +Proof. reflexivity. Qed. + +Lemma to_Z_0 : [|0|] = 0. +Proof. reflexivity. Qed. + +Lemma to_Z_1 : [|1|] = 1. +Proof. reflexivity. Qed. + +(* Notations *) +Local Open Scope Z_scope. + +Notation "[+| c |]" := + (interp_carry 1 wB to_Z c) (at level 0, c at level 99) : int63_scope. + +Notation "[-| c |]" := + (interp_carry (-1) wB to_Z c) (at level 0, c at level 99) : int63_scope. + +Notation "[|| x ||]" := + (zn2z_to_Z wB to_Z x) (at level 0, x at level 99) : int63_scope. + +(* Bijection : int63 <-> Bvector size *) + +Axiom of_to_Z : forall x, of_Z [| x |] = x. + +Notation "'φ' x" := [| x |] (at level 0) : int63_scope. + +Lemma can_inj {rT aT} {f: aT -> rT} {g: rT -> aT} (K: forall a, g (f a) = a) {a a'} (e: f a = f a') : a = a'. +Proof. generalize (K a) (K a'). congruence. Qed. + +Lemma to_Z_inj x y : φ x = φ y → x = y. +Proof. exact (λ e, can_inj of_to_Z e). Qed. + +(** Specification of logical operations *) +Local Open Scope Z_scope. +Axiom lsl_spec : forall x p, [| x << p |] = [| x |] * 2 ^ [| p |] mod wB. + +Axiom lsr_spec : forall x p, [|x >> p|] = [|x|] / 2 ^ [|p|]. + +Axiom land_spec: forall x y i , bit (x land y) i = bit x i && bit y i. + +Axiom lor_spec: forall x y i, bit (x lor y) i = bit x i || bit y i. + +Axiom lxor_spec: forall x y i, bit (x lxor y) i = xorb (bit x i) (bit y i). + +(** Specification of basic opetations *) + +(* Arithmetic modulo operations *) + +(* Remarque : les axiomes seraient plus simple si on utilise of_Z a la place : + exemple : add_spec : forall x y, of_Z (x + y) = of_Z x + of_Z y. *) + +Axiom add_spec : forall x y, [|x + y|] = ([|x|] + [|y|]) mod wB. + +Axiom sub_spec : forall x y, [|x - y|] = ([|x|] - [|y|]) mod wB. + +Axiom mul_spec : forall x y, [| x * y |] = [|x|] * [|y|] mod wB. + +Axiom mulc_spec : forall x y, [|x|] * [|y|] = [|fst (mulc x y)|] * wB + [|snd (mulc x y)|]. + +Axiom div_spec : forall x y, [|x / y|] = [|x|] / [|y|]. + +Axiom mod_spec : forall x y, [|x \% y|] = [|x|] mod [|y|]. + +(* Comparisons *) +Axiom eqb_correct : forall i j, (i == j)%int63 = true -> i = j. + +Axiom eqb_refl : forall x, (x == x)%int63 = true. + +Axiom ltb_spec : forall x y, (x < y)%int63 = true <-> [|x|] < [|y|]. + +Axiom leb_spec : forall x y, (x <= y)%int63 = true <-> [|x|] <= [|y|]. + +(** Exotic operations *) + +(** I should add the definition (like for compare) *) +Primitive head0 := #int63_head0. +Primitive tail0 := #int63_tail0. + +(** Axioms on operations which are just short cut *) + +Axiom compare_def_spec : forall x y, compare x y = compare_def x y. + +Axiom head0_spec : forall x, 0 < [|x|] -> + wB/ 2 <= 2 ^ ([|head0 x|]) * [|x|] < wB. + +Axiom tail0_spec : forall x, 0 < [|x|] -> + (exists y, 0 <= y /\ [|x|] = (2 * y + 1) * (2 ^ [|tail0 x|]))%Z. + +Axiom addc_def_spec : forall x y, (x +c y)%int63 = addc_def x y. + +Axiom addcarryc_def_spec : forall x y, addcarryc x y = addcarryc_def x y. + +Axiom subc_def_spec : forall x y, (x -c y)%int63 = subc_def x y. + +Axiom subcarryc_def_spec : forall x y, subcarryc x y = subcarryc_def x y. + +Axiom diveucl_def_spec : forall x y, diveucl x y = diveucl_def x y. + +Axiom diveucl_21_spec : forall a1 a2 b, + let (q,r) := diveucl_21 a1 a2 b in + ([|q|],[|r|]) = Z.div_eucl ([|a1|] * wB + [|a2|]) [|b|]. + +Axiom addmuldiv_def_spec : forall p x y, + addmuldiv p x y = addmuldiv_def p x y. + +(** Square root functions using newton iteration **) +Local Open Scope int63_scope. + +Definition sqrt_step (rec: int -> int -> int) (i j: int) := + let quo := i / j in + if quo < j then rec i ((j + quo) >> 1) + else j. + +Definition iter_sqrt := + Eval lazy beta delta [sqrt_step] in + fix iter_sqrt (n: nat) (rec: int -> int -> int) + (i j: int) {struct n} : int := + sqrt_step + (fun i j => match n with + O => rec i j + | S n => (iter_sqrt n (iter_sqrt n rec)) i j + end) i j. + +Definition sqrt i := + match compare 1 i with + Gt => 0 + | Eq => 1 + | Lt => iter_sqrt size (fun i j => j) i (i >> 1) + end. + +Definition high_bit := 1 << (digits - 1). + +Definition sqrt2_step (rec: int -> int -> int -> int) + (ih il j: int) := + if ih < j then + let (quo,_) := diveucl_21 ih il j in + if quo < j then + match j +c quo with + | C0 m1 => rec ih il (m1 >> 1) + | C1 m1 => rec ih il ((m1 >> 1) + high_bit) + end + else j + else j. + +Definition iter2_sqrt := + Eval lazy beta delta [sqrt2_step] in + fix iter2_sqrt (n: nat) + (rec: int -> int -> int -> int) + (ih il j: int) {struct n} : int := + sqrt2_step + (fun ih il j => + match n with + | O => rec ih il j + | S n => (iter2_sqrt n (iter2_sqrt n rec)) ih il j + end) ih il j. + +Definition sqrt2 ih il := + let s := iter2_sqrt size (fun ih il j => j) ih il max_int in + let (ih1, il1) := mulc s s in + match il -c il1 with + | C0 il2 => + if ih1 < ih then (s, C1 il2) else (s, C0 il2) + | C1 il2 => + if ih1 < (ih - 1) then (s, C1 il2) else (s, C0 il2) + end. + +(** Gcd **) +Fixpoint gcd_rec (guard:nat) (i j:int) {struct guard} := + match guard with + | O => 1 + | S p => if j == 0 then i else gcd_rec p j (i \% j) + end. + +Definition gcd := gcd_rec (2*size). + +(** equality *) +Lemma eqb_complete : forall x y, x = y -> (x == y) = true. +Proof. + intros x y H; rewrite -> H, eqb_refl;trivial. +Qed. + +Lemma eqb_spec : forall x y, (x == y) = true <-> x = y. +Proof. + split;auto using eqb_correct, eqb_complete. +Qed. + +Lemma eqb_false_spec : forall x y, (x == y) = false <-> x <> y. +Proof. + intros;rewrite <- not_true_iff_false, eqb_spec;split;trivial. +Qed. + +Lemma eqb_false_complete : forall x y, x <> y -> (x == y) = false. +Proof. + intros x y;rewrite eqb_false_spec;trivial. +Qed. + +Lemma eqb_false_correct : forall x y, (x == y) = false -> x <> y. +Proof. + intros x y;rewrite eqb_false_spec;trivial. +Qed. + +Definition eqs (i j : int) : {i = j} + { i <> j } := + (if i == j as b return ((b = true -> i = j) -> (b = false -> i <> j) -> {i=j} + {i <> j} ) + then fun (Heq : true = true -> i = j) _ => left _ (Heq (eq_refl true)) + else fun _ (Hdiff : false = false -> i <> j) => right _ (Hdiff (eq_refl false))) + (eqb_correct i j) + (eqb_false_correct i j). + +Lemma eq_dec : forall i j:int, i = j \/ i <> j. +Proof. + intros i j;destruct (eqs i j);auto. +Qed. + +(* Extra function on equality *) + +Definition cast i j := + (if i == j as b return ((b = true -> i = j) -> option (forall P : int -> Type, P i -> P j)) + then fun Heq : true = true -> i = j => + Some + (fun (P : int -> Type) (Hi : P i) => + match Heq (eq_refl true) in (_ = y) return (P y) with + | eq_refl => Hi + end) + else fun _ : false = true -> i = j => None) (eqb_correct i j). + +Lemma cast_refl : forall i, cast i i = Some (fun P H => H). +Proof. + unfold cast;intros. + generalize (eqb_correct i i). + rewrite eqb_refl;intros. + rewrite (Eqdep_dec.eq_proofs_unicity eq_dec (e (eq_refl true)) (eq_refl i));trivial. +Qed. + +Lemma cast_diff : forall i j, i == j = false -> cast i j = None. +Proof. + intros;unfold cast;intros; generalize (eqb_correct i j). + rewrite H;trivial. +Qed. + +Definition eqo i j := + (if i == j as b return ((b = true -> i = j) -> option (i=j)) + then fun Heq : true = true -> i = j => + Some (Heq (eq_refl true)) + else fun _ : false = true -> i = j => None) (eqb_correct i j). + +Lemma eqo_refl : forall i, eqo i i = Some (eq_refl i). +Proof. + unfold eqo;intros. + generalize (eqb_correct i i). + rewrite eqb_refl;intros. + rewrite (Eqdep_dec.eq_proofs_unicity eq_dec (e (eq_refl true)) (eq_refl i));trivial. +Qed. + +Lemma eqo_diff : forall i j, i == j = false -> eqo i j = None. +Proof. + unfold eqo;intros; generalize (eqb_correct i j). + rewrite H;trivial. +Qed. + +(** Comparison *) + +Lemma eqbP x y : reflect ([| x |] = [| y |]) (x == y). +Proof. apply iff_reflect; rewrite eqb_spec; split; [ apply to_Z_inj | apply f_equal ]. Qed. + +Lemma ltbP x y : reflect ([| x |] < [| y |])%Z (x < y). +Proof. apply iff_reflect; symmetry; apply ltb_spec. Qed. + +Lemma lebP x y : reflect ([| x |] <= [| y |])%Z (x ≤ y). +Proof. apply iff_reflect; symmetry; apply leb_spec. Qed. + +Lemma compare_spec x y : compare x y = ([|x|] ?= [|y|])%Z. +Proof. + rewrite compare_def_spec; unfold compare_def. + case ltbP; [ auto using Z.compare_lt_iff | intros hge ]. + case eqbP; [ now symmetry; apply Z.compare_eq_iff | intros hne ]. + symmetry; apply Z.compare_gt_iff; lia. +Qed. + +Lemma is_zero_spec x : is_zero x = true <-> x = 0%int63. +Proof. apply eqb_spec. Qed. + +Lemma diveucl_spec x y : + let (q,r) := diveucl x y in + ([| q |], [| r |]) = Z.div_eucl [| x |] [| y |]. +Proof. + rewrite diveucl_def_spec; unfold diveucl_def; rewrite div_spec, mod_spec; unfold Z.div, Zmod. + destruct (Z.div_eucl [| x |] [| y |]); trivial. +Qed. + +Local Open Scope Z_scope. +(** Addition *) +Lemma addc_spec x y : [+| x +c y |] = [| x |] + [| y |]. +Proof. + rewrite addc_def_spec; unfold addc_def, interp_carry. + pose proof (to_Z_bounded x); pose proof (to_Z_bounded y). + case ltbP; rewrite add_spec. + case (Z_lt_ge_dec ([| x |] + [| y |]) wB). + intros k; rewrite Zmod_small; lia. + intros hge; rewrite <- (Zmod_unique _ _ 1 ([| x |] + [| y |] - wB)); lia. + case (Z_lt_ge_dec ([| x |] + [| y |]) wB). + intros k; rewrite Zmod_small; lia. + intros hge; rewrite <- (Zmod_unique _ _ 1 ([| x |] + [| y |] - wB)); lia. +Qed. + +Lemma succ_spec x : [| succ x |] = ([| x |] + 1) mod wB. +Proof. apply add_spec. Qed. + +Lemma succc_spec x : [+| succc x |] = [| x |] + 1. +Proof. apply addc_spec. Qed. + +Lemma addcarry_spec x y : [| addcarry x y |] = ([| x |] + [| y |] + 1) mod wB. +Proof. unfold addcarry; rewrite -> !add_spec, Zplus_mod_idemp_l; trivial. Qed. + +Lemma addcarryc_spec x y : [+| addcarryc x y |] = [| x |] + [| y |] + 1. +Proof. + rewrite addcarryc_def_spec; unfold addcarryc_def, interp_carry. + pose proof (to_Z_bounded x); pose proof (to_Z_bounded y). + case lebP; rewrite addcarry_spec. + case (Z_lt_ge_dec ([| x |] + [| y |] + 1) wB). + intros hlt; rewrite Zmod_small; lia. + intros hge; rewrite <- (Zmod_unique _ _ 1 ([| x |] + [| y |] + 1 - wB)); lia. + case (Z_lt_ge_dec ([| x |] + [| y |] + 1) wB). + intros hlt; rewrite Zmod_small; lia. + intros hge; rewrite <- (Zmod_unique _ _ 1 ([| x |] + [| y |] + 1 - wB)); lia. +Qed. + +(** Subtraction *) +Lemma subc_spec x y : [-| x -c y |] = [| x |] - [| y |]. +Proof. + rewrite subc_def_spec; unfold subc_def; unfold interp_carry. + pose proof (to_Z_bounded x); pose proof (to_Z_bounded y). + case lebP. + intros hle; rewrite sub_spec, Z.mod_small; lia. + intros hgt; rewrite sub_spec, <- (Zmod_unique _ wB (-1) ([| x |] - [| y |] + wB)); lia. +Qed. + +Lemma pred_spec x : [| pred x |] = ([| x |] - 1) mod wB. +Proof. apply sub_spec. Qed. + +Lemma predc_spec x : [-| predc x |] = [| x |] - 1. +Proof. apply subc_spec. Qed. + +Lemma oppc_spec x : [-| oppc x |] = - [| x |]. +Proof. unfold oppc; rewrite -> subc_spec, to_Z_0; trivial. Qed. + +Lemma opp_spec x : [|- x |] = - [| x |] mod wB. +Proof. unfold opp; rewrite -> sub_spec, to_Z_0; trivial. Qed. + +Lemma oppcarry_spec x : [| oppcarry x |] = wB - [| x |] - 1. +Proof. + unfold oppcarry; rewrite sub_spec. + rewrite <- Zminus_plus_distr, Zplus_comm, Zminus_plus_distr. + apply Zmod_small. + generalize (to_Z_bounded x); auto with zarith. +Qed. + +Lemma subcarry_spec x y : [| subcarry x y |] = ([| x |] - [| y |] - 1) mod wB. +Proof. unfold subcarry; rewrite !sub_spec, Zminus_mod_idemp_l; trivial. Qed. + +Lemma subcarryc_spec x y : [-| subcarryc x y |] = [| x |] - [| y |] - 1. +Proof. + rewrite subcarryc_def_spec; unfold subcarryc_def, interp_carry; fold (subcarry x y). + pose proof (to_Z_bounded x); pose proof (to_Z_bounded y). + case ltbP; rewrite subcarry_spec. + intros hlt; rewrite Zmod_small; lia. + intros hge; rewrite <- (Zmod_unique _ _ (-1) ([| x |] - [| y |] - 1 + wB)); lia. +Qed. + +(** GCD *) +Lemma to_Z_gcd : forall i j, [| gcd i j |] = Zgcdn (2 * size) [| j |] [| i |]. +Proof. + unfold gcd. + elim (2*size)%nat. reflexivity. + intros n ih i j; simpl. + pose proof (to_Z_bounded j) as hj; pose proof (to_Z_bounded i). + case eqbP; rewrite to_Z_0. + intros ->; rewrite Z.abs_eq; lia. + intros hne; rewrite ih; clear ih. + rewrite <- mod_spec. + revert hj hne; case [| j |]; intros; lia. +Qed. + +Lemma gcd_spec a b : Zis_gcd [| a |] [| b |] [| gcd a b |]. +Proof. + rewrite to_Z_gcd. + apply Zis_gcd_sym. + apply Zgcdn_is_gcd. + unfold Zgcd_bound. + generalize (to_Z_bounded b). + destruct [|b|]. + unfold size; auto with zarith. + intros (_,H). + cut (Psize p <= size)%nat; [ lia | rewrite <- Zpower2_Psize; auto]. + intros (H,_); compute in H; elim H; auto. +Qed. + +(** Head0, Tail0 *) +Lemma head00_spec x : [| x |] = 0 -> [| head0 x |] = [| digits |]. +Proof. now intros h; rewrite (to_Z_inj _ 0 h). Qed. + +Lemma tail00_spec x : [| x |] = 0 -> [|tail0 x|] = [|digits|]. +Proof. now intros h; rewrite (to_Z_inj _ 0 h). Qed. + +Infix "≡" := (eqm wB) (at level 80) : int63_scope. + +Lemma eqm_mod x y : x mod wB ≡ y mod wB → x ≡ y. +Proof. + intros h. + eapply (eqm_trans). + apply eqm_sym; apply Zmod_eqm. + apply (eqm_trans _ _ _ _ h). + apply Zmod_eqm. +Qed. + +Lemma eqm_sub x y : x ≡ y → x - y ≡ 0. +Proof. intros h; unfold eqm; rewrite Zminus_mod, h, Z.sub_diag; reflexivity. Qed. + +Lemma eqmE x y : x ≡ y → ∃ k, x - y = k * wB. +Proof. + intros h. + exact (Zmod_divide (x - y) wB (λ e, let 'eq_refl := e in I) (eqm_sub _ _ h)). +Qed. + +Lemma eqm_subE x y : x ≡ y ↔ x - y ≡ 0. +Proof. + split. apply eqm_sub. + intros h; case (eqmE _ _ h); clear h; intros q h. + assert (y = x - q * wB) by lia. + clear h; subst y. + unfold eqm; rewrite Zminus_mod, Z_mod_mult, Z.sub_0_r, Zmod_mod; reflexivity. +Qed. + +Lemma int_eqm x y : x = y → φ x ≡ φ y. +Proof. unfold eqm; intros ->; reflexivity. Qed. + +Lemma eqmI x y : φ x ≡ φ y → x = y. +Proof. + unfold eqm. + repeat rewrite Zmod_small by apply to_Z_bounded. + apply to_Z_inj. +Qed. + +(* ADD *) +Lemma add_assoc x y z: (x + (y + z) = (x + y) + z)%int63. +Proof. + apply to_Z_inj; rewrite !add_spec. + rewrite -> Zplus_mod_idemp_l, Zplus_mod_idemp_r, Zplus_assoc; auto. +Qed. + +Lemma add_comm x y: (x + y = y + x)%int63. +Proof. + apply to_Z_inj; rewrite -> !add_spec, Zplus_comm; auto. +Qed. + +Lemma add_le_r m n: + if (n <= m + n)%int63 then ([|m|] + [|n|] < wB)%Z else (wB <= [|m|] + [|n|])%Z. +Proof. + case (to_Z_bounded m); intros H1m H2m. + case (to_Z_bounded n); intros H1n H2n. + case (Zle_or_lt wB ([|m|] + [|n|])); intros H. + assert (H1: ([| m + n |] = [|m|] + [|n|] - wB)%Z). + rewrite add_spec. + replace (([|m|] + [|n|]) mod wB)%Z with (((([|m|] + [|n|]) - wB) + wB) mod wB)%Z. + rewrite -> Zplus_mod, Z_mod_same_full, Zplus_0_r, !Zmod_small; auto with zarith. + rewrite !Zmod_small; auto with zarith. + apply f_equal2 with (f := Zmod); auto with zarith. + case_eq (n <= m + n)%int63; auto. + rewrite leb_spec, H1; auto with zarith. + assert (H1: ([| m + n |] = [|m|] + [|n|])%Z). + rewrite add_spec, Zmod_small; auto with zarith. + replace (n <= m + n)%int63 with true; auto. + apply sym_equal; rewrite leb_spec, H1; auto with zarith. +Qed. + +Lemma add_cancel_l x y z : (x + y = x + z)%int63 -> y = z. +Proof. + intros h; apply int_eqm in h; rewrite !add_spec in h; apply eqm_mod, eqm_sub in h. + replace (_ + _ - _) with (φ(y) - φ(z)) in h by lia. + rewrite <- eqm_subE in h. + apply eqmI, h. +Qed. + +Lemma add_cancel_r x y z : (y + x = z + x)%int63 -> y = z. +Proof. + rewrite !(fun t => add_comm t x); intros Hl; apply (add_cancel_l x); auto. +Qed. + +Coercion b2i (b: bool) : int := if b then 1%int63 else 0%int63. + +(* LSR *) +Lemma lsr0 i : 0 >> i = 0%int63. +Proof. apply to_Z_inj; rewrite lsr_spec; reflexivity. Qed. + +Lemma lsr_0_r i: i >> 0 = i. +Proof. apply to_Z_inj; rewrite lsr_spec, Zdiv_1_r; exact eq_refl. Qed. + +Lemma lsr_1 n : 1 >> n = (n == 0). +Proof. + case eqbP. + intros h; rewrite (to_Z_inj _ _ h), lsr_0_r; reflexivity. + intros Hn. + assert (H1n : (1 >> n = 0)%int63); auto. + apply to_Z_inj; rewrite lsr_spec. + apply Zdiv_small; rewrite to_Z_1; split; auto with zarith. + change 1%Z with (2^0)%Z. + apply Zpower_lt_monotone; split; auto with zarith. + rewrite to_Z_0 in Hn. + generalize (to_Z_bounded n). + lia. +Qed. + +Lemma lsr_add i m n: ((i >> m) >> n = if n <= m + n then i >> (m + n) else 0)%int63. +Proof. + case (to_Z_bounded m); intros H1m H2m. + case (to_Z_bounded n); intros H1n H2n. + case (to_Z_bounded i); intros H1i H2i. + generalize (add_le_r m n); case (n <= m + n)%int63; intros H. + apply to_Z_inj; rewrite -> !lsr_spec, Zdiv_Zdiv, <- Zpower_exp; auto with zarith. + rewrite add_spec, Zmod_small; auto with zarith. + apply to_Z_inj; rewrite -> !lsr_spec, Zdiv_Zdiv, <- Zpower_exp; auto with zarith. + apply Zdiv_small. split; [ auto with zarith | ]. + eapply Z.lt_le_trans; [ | apply Zpower2_le_lin ]; auto with zarith. +Qed. + +Lemma lsr_add_distr x y n: (x + y) << n = ((x << n) + (y << n))%int63. +Proof. + apply to_Z_inj. + rewrite -> add_spec, !lsl_spec, add_spec. + rewrite -> Zmult_mod_idemp_l, <-Zplus_mod. + apply f_equal2 with (f := Zmod); auto with zarith. +Qed. + +(* LSL *) +Lemma lsl0 i: 0 << i = 0%int63. +Proof. + apply to_Z_inj. + generalize (lsl_spec 0 i). + rewrite to_Z_0, Zmult_0_l, Zmod_0_l; auto. +Qed. + +Lemma lsl0_r i : i << 0 = i. +Proof. + apply to_Z_inj. + rewrite -> lsl_spec, to_Z_0, Z.mul_1_r. + apply Zmod_small; apply (to_Z_bounded i). +Qed. + +Lemma lsl_add_distr x y n: (x + y) << n = ((x << n) + (y << n))%int63. +Proof. + apply to_Z_inj; rewrite -> !lsl_spec, !add_spec, Zmult_mod_idemp_l. + rewrite -> !lsl_spec, <-Zplus_mod. + apply f_equal2 with (f := Zmod); auto with zarith. +Qed. + +Lemma lsr_M_r x i (H: (digits <= i = true)%int63) : x >> i = 0%int63. +Proof. + apply to_Z_inj. + rewrite lsr_spec, to_Z_0. + case (to_Z_bounded x); intros H1x H2x. + case (to_Z_bounded digits); intros H1d H2d. + rewrite -> leb_spec in H. + apply Zdiv_small; split; [ auto | ]. + apply (Z.lt_le_trans _ _ _ H2x). + unfold wB; change (Z_of_nat size) with [|digits|]. + apply Zpower_le_monotone; auto with zarith. +Qed. + +(* BIT *) +Lemma bit_0_spec i: [|bit i 0|] = [|i|] mod 2. +Proof. + unfold bit, is_zero. rewrite lsr_0_r. + assert (Hbi: ([|i|] mod 2 < 2)%Z). + apply Z_mod_lt; auto with zarith. + case (to_Z_bounded i); intros H1i H2i. + case (Zmod_le_first [|i|] 2); auto with zarith; intros H3i H4i. + assert (H2b: (0 < 2 ^ [|digits - 1|])%Z). + apply Zpower_gt_0; auto with zarith. + case (to_Z_bounded (digits -1)); auto with zarith. + assert (H: [|i << (digits -1)|] = ([|i|] mod 2 * 2^ [|digits -1|])%Z). + rewrite lsl_spec. + rewrite -> (Z_div_mod_eq [|i|] 2) at 1; auto with zarith. + rewrite -> Zmult_plus_distr_l, <-Zplus_mod_idemp_l. + rewrite -> (Zmult_comm 2), <-Zmult_assoc. + replace (2 * 2 ^ [|digits - 1|])%Z with wB; auto. + rewrite Z_mod_mult, Zplus_0_l; apply Zmod_small. + split; auto with zarith. + replace wB with (2 * 2 ^ [|digits -1|])%Z; auto. + apply Zmult_lt_compat_r; auto with zarith. + case (Zle_lt_or_eq 0 ([|i|] mod 2)); auto with zarith; intros Hi. + 2: generalize H; rewrite <-Hi, Zmult_0_l. + 2: replace 0%Z with [|0|]; auto. + 2: now case eqbP. + generalize H; replace ([|i|] mod 2) with 1%Z; auto with zarith. + rewrite Zmult_1_l. + intros H1. + assert (H2: [|i << (digits - 1)|] <> [|0|]). + replace [|0|] with 0%Z; auto with zarith. + now case eqbP. +Qed. + +Lemma bit_split i : ( i = (i >> 1 ) << 1 + bit i 0)%int63. +Proof. + apply to_Z_inj. + rewrite -> add_spec, lsl_spec, lsr_spec, bit_0_spec, Zplus_mod_idemp_l. + replace (2 ^ [|1|]) with 2%Z; auto with zarith. + rewrite -> Zmult_comm, <-Z_div_mod_eq; auto with zarith. + rewrite Zmod_small; auto; case (to_Z_bounded i); auto. +Qed. + +Lemma bit_lsr x i j : + (bit (x >> i) j = if j <= i + j then bit x (i + j) else false)%int63. +Proof. + unfold bit; rewrite lsr_add; case (_ ≤ _); auto. +Qed. + +Lemma bit_b2i (b: bool) i : bit b i = (i == 0) && b. +Proof. + case b; unfold bit; simpl b2i. + rewrite lsr_1; case (i == 0); auto. + rewrite lsr0, lsl0, andb_false_r; auto. +Qed. + +Lemma bit_1 n : bit 1 n = (n == 0). +Proof. + unfold bit; rewrite lsr_1. + case (_ == _); simpl; auto. +Qed. + +Local Hint Resolve Z.lt_gt Z.div_pos : zarith. + +Lemma to_Z_split x : [|x|] = [|(x >> 1)|] * 2 + [|bit x 0|]. +Proof. + case (to_Z_bounded x); intros H1x H2x. + case (to_Z_bounded (bit x 0)); intros H1b H2b. + assert (F1: 0 <= [|x >> 1|] < wB/2). + rewrite -> lsr_spec, to_Z_1, Z.pow_1_r. split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + rewrite -> (bit_split x) at 1. + rewrite -> add_spec, Zmod_small, lsl_spec, to_Z_1, Z.pow_1_r, Zmod_small; + split; auto with zarith. + change wB with ((wB/2)*2); auto with zarith. + rewrite -> lsl_spec, to_Z_1, Z.pow_1_r, Zmod_small; auto with zarith. + change wB with ((wB/2)*2); auto with zarith. + rewrite -> lsl_spec, to_Z_1, Z.pow_1_r, Zmod_small; auto with zarith. + 2: change wB with ((wB/2)*2); auto with zarith. + change wB with (((wB/2 - 1) * 2 + 1) + 1). + assert ([|bit x 0|] <= 1); auto with zarith. + case bit; discriminate. +Qed. + +Lemma bit_M i n (H: (digits <= n = true)%int63): bit i n = false. +Proof. unfold bit; rewrite lsr_M_r; auto. Qed. + +Lemma bit_half i n (H: (n < digits = true)%int63) : bit (i>>1) n = bit i (n+1). +Proof. + unfold bit. + rewrite lsr_add. + case_eq (n <= (1 + n))%int63. + replace (1+n)%int63 with (n+1)%int63; [auto|idtac]. + apply to_Z_inj; rewrite !add_spec, Zplus_comm; auto. + intros H1; assert (H2: n = max_int). + 2: generalize H; rewrite H2; discriminate. + case (to_Z_bounded n); intros H1n H2n. + case (Zle_lt_or_eq [|n|] (wB - 1)); auto with zarith; + intros H2; apply to_Z_inj; auto. + generalize (add_le_r 1 n); rewrite H1. + change [|max_int|] with (wB - 1)%Z. + replace [|1|] with 1%Z; auto with zarith. +Qed. + +Lemma bit_ext i j : (forall n, bit i n = bit j n) -> i = j. +Proof. + case (to_Z_bounded j); case (to_Z_bounded i). + unfold wB; revert i j; elim size. + simpl; intros i j ???? _; apply to_Z_inj; lia. + intros n ih i j. + rewrite Nat2Z.inj_succ, Z.pow_succ_r by auto with zarith. + intros hi1 hi2 hj1 hj2 hext. + rewrite (bit_split i), (bit_split j), hext. + do 2 f_equal; apply ih; clear ih. + 1, 3: apply to_Z_bounded. + 1, 2: rewrite lsr_spec; auto using Z_lt_div2. + intros b. + case (Zle_or_lt [|digits|] [|b|]). + rewrite <- leb_spec; intros; rewrite !bit_M; auto. + rewrite <- ltb_spec; intros; rewrite !bit_half; auto. +Qed. + +Lemma bit_lsl x i j : bit (x << i) j = +(if (j < i) || (digits <= j) then false else bit x (j - i))%int63. +Proof. + assert (F1: 1 >= 0) by discriminate. + case_eq (digits <= j)%int63; intros H. + rewrite orb_true_r, bit_M; auto. + set (d := [|digits|]). + case (Zle_or_lt d [|j|]); intros H1. + case (leb_spec digits j); rewrite H; auto with zarith. + intros _ HH; generalize (HH H1); discriminate. + clear H. + generalize (ltb_spec j i); case ltb; intros H2; unfold bit; simpl. + assert (F2: ([|j|] < [|i|])%Z) by (case H2; auto); clear H2. + replace (is_zero (((x << i) >> j) << (digits - 1))) with true; auto. + case (to_Z_bounded j); intros H1j H2j. + apply sym_equal; rewrite is_zero_spec; apply to_Z_inj. + rewrite lsl_spec, lsr_spec, lsl_spec. + replace wB with (2^d); auto. + pattern d at 1; replace d with ((d - ([|j|] + 1)) + ([|j|] + 1))%Z by ring. + rewrite Zpower_exp; auto with zarith. + replace [|i|] with (([|i|] - ([|j|] + 1)) + ([|j|] + 1))%Z by ring. + rewrite -> Zpower_exp, Zmult_assoc; auto with zarith. + rewrite Zmult_mod_distr_r. + rewrite -> Zplus_comm, Zpower_exp, !Zmult_assoc; auto with zarith. + rewrite -> Z_div_mult_full; auto with zarith. + rewrite <-Zmult_assoc, <-Zpower_exp; auto with zarith. + replace (1 + [|digits - 1|])%Z with d; auto with zarith. + rewrite Z_mod_mult; auto. + case H2; intros _ H3; case (Zle_or_lt [|i|] [|j|]); intros F2. + 2: generalize (H3 F2); discriminate. + clear H2 H3. + apply f_equal with (f := negb). + apply f_equal with (f := is_zero). + apply to_Z_inj. + rewrite -> !lsl_spec, !lsr_spec, !lsl_spec. + pattern wB at 2 3; replace wB with (2^(1+ [|digits - 1|])); auto. + rewrite -> Zpower_exp, Z.pow_1_r; auto with zarith. + rewrite !Zmult_mod_distr_r. + apply f_equal2 with (f := Zmult); auto. + replace wB with (2^ d); auto with zarith. + replace d with ((d - [|i|]) + [|i|])%Z by ring. + case (to_Z_bounded i); intros H1i H2i. + rewrite Zpower_exp; auto with zarith. + rewrite Zmult_mod_distr_r. + case (to_Z_bounded j); intros H1j H2j. + replace [|j - i|] with ([|j|] - [|i|])%Z. + 2: rewrite sub_spec, Zmod_small; auto with zarith. + set (d1 := (d - [|i|])%Z). + set (d2 := ([|j|] - [|i|])%Z). + pattern [|j|] at 1; + replace [|j|] with (d2 + [|i|])%Z. + 2: unfold d2; ring. + rewrite -> Zpower_exp; auto with zarith. + rewrite -> Zdiv_mult_cancel_r. + 2: generalize (Zpower2_lt_lin [| i |] H1i); auto with zarith. + rewrite -> (Z_div_mod_eq [|x|] (2^d1)) at 2; auto with zarith. + pattern d1 at 2; + replace d1 with (d2 + (1+ (d - [|j|] - 1)))%Z + by (unfold d1, d2; ring). + rewrite Zpower_exp; auto with zarith. + rewrite <-Zmult_assoc, Zmult_comm. + rewrite Zdiv.Z_div_plus_full_l; auto with zarith. + rewrite Zpower_exp, Z.pow_1_r; auto with zarith. + rewrite <-Zplus_mod_idemp_l. + rewrite <-!Zmult_assoc, Zmult_comm, Z_mod_mult, Zplus_0_l; auto. +Qed. + +(* LOR *) +Lemma lor_lsr i1 i2 i: (i1 lor i2) >> i = (i1 >> i) lor (i2 >> i). +Proof. + apply bit_ext; intros n. + rewrite -> lor_spec, !bit_lsr, lor_spec. + case (_ <= _)%int63; auto. +Qed. + +Lemma lor_le x y : (y <= x lor y)%int63 = true. +Proof. + generalize x y (to_Z_bounded x) (to_Z_bounded y); clear x y. + unfold wB; elim size. + replace (2^Z_of_nat 0) with 1%Z; auto with zarith. + intros x y Hx Hy; replace x with 0%int63. + replace y with 0%int63; auto. + apply to_Z_inj; rewrite to_Z_0; auto with zarith. + apply to_Z_inj; rewrite to_Z_0; auto with zarith. + intros n IH x y; rewrite inj_S. + unfold Z.succ; rewrite -> Zpower_exp, Z.pow_1_r; auto with zarith. + intros Hx Hy. + rewrite leb_spec. + rewrite -> (to_Z_split y) at 1; rewrite (to_Z_split (x lor y)). + assert ([|y>>1|] <= [|(x lor y) >> 1|]). + rewrite -> lor_lsr, <-leb_spec; apply IH. + rewrite -> lsr_spec, to_Z_1, Z.pow_1_r; split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + rewrite -> lsr_spec, to_Z_1, Z.pow_1_r; split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + assert ([|bit y 0|] <= [|bit (x lor y) 0|]); auto with zarith. + rewrite lor_spec; do 2 case bit; try discriminate. +Qed. + +Lemma bit_0 n : bit 0 n = false. +Proof. unfold bit; rewrite lsr0; auto. Qed. + +Lemma bit_add_or x y: + (forall n, bit x n = true -> bit y n = true -> False) <-> (x + y)%int63= x lor y. +Proof. + generalize x y (to_Z_bounded x) (to_Z_bounded y); clear x y. + unfold wB; elim size. + replace (2^Z_of_nat 0) with 1%Z; auto with zarith. + intros x y Hx Hy; replace x with 0%int63. + replace y with 0%int63. + split; auto; intros _ n; rewrite !bit_0; discriminate. + apply to_Z_inj; rewrite to_Z_0; auto with zarith. + apply to_Z_inj; rewrite to_Z_0; auto with zarith. + intros n IH x y; rewrite inj_S. + unfold Z.succ; rewrite Zpower_exp, Z.pow_1_r; auto with zarith. + intros Hx Hy. + split. + intros Hn. + assert (F1: ((x >> 1) + (y >> 1))%int63 = (x >> 1) lor (y >> 1)). + apply IH. + rewrite lsr_spec, Z.pow_1_r; split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + rewrite lsr_spec, Z.pow_1_r; split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + intros m H1 H2. + case_eq (digits <= m)%int63; [idtac | rewrite <- not_true_iff_false]; + intros Heq. + rewrite bit_M in H1; auto; discriminate. + rewrite leb_spec in Heq. + apply (Hn (m + 1)%int63); + rewrite <-bit_half; auto; rewrite ltb_spec; auto with zarith. + rewrite (bit_split (x lor y)), lor_lsr, <- F1, lor_spec. + replace (b2i (bit x 0 || bit y 0)) with (bit x 0 + bit y 0)%int63. + 2: generalize (Hn 0%int63); do 2 case bit; auto; intros [ ]; auto. + rewrite lsl_add_distr. + rewrite (bit_split x) at 1; rewrite (bit_split y) at 1. + rewrite <-!add_assoc; apply f_equal2 with (f := add); auto. + rewrite add_comm, <-!add_assoc; apply f_equal2 with (f := add); auto. + rewrite add_comm; auto. + intros Heq. + generalize (add_le_r x y); rewrite Heq, lor_le; intro Hb. + generalize Heq; rewrite (bit_split x) at 1; rewrite (bit_split y )at 1; clear Heq. + rewrite (fun y => add_comm y (bit x 0)), <-!add_assoc, add_comm, + <-!add_assoc, (add_comm (bit y 0)), add_assoc, <-lsr_add_distr. + rewrite (bit_split (x lor y)), lor_spec. + intros Heq. + assert (F: (bit x 0 + bit y 0)%int63 = (bit x 0 || bit y 0)). + assert (F1: (2 | wB)) by (apply Zpower_divide; apply refl_equal). + assert (F2: 0 < wB) by (apply refl_equal). + assert (F3: [|bit x 0 + bit y 0|] mod 2 = [|bit x 0 || bit y 0|] mod 2). + apply trans_equal with (([|(x>>1 + y>>1) << 1|] + [|bit x 0 + bit y 0|]) mod 2). + rewrite lsl_spec, Zplus_mod, <-Zmod_div_mod; auto with zarith. + rewrite Z.pow_1_r, Z_mod_mult, Zplus_0_l, Zmod_mod; auto with zarith. + rewrite (Zmod_div_mod 2 wB), <-add_spec, Heq; auto with zarith. + rewrite add_spec, <-Zmod_div_mod; auto with zarith. + rewrite lsl_spec, Zplus_mod, <-Zmod_div_mod; auto with zarith. + rewrite Z.pow_1_r, Z_mod_mult, Zplus_0_l, Zmod_mod; auto with zarith. + generalize F3; do 2 case bit; try discriminate; auto. + case (IH (x >> 1) (y >> 1)). + rewrite lsr_spec, to_Z_1, Z.pow_1_r; split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + rewrite lsr_spec, to_Z_1, Z.pow_1_r; split; auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + intros _ HH m; case (to_Z_bounded m); intros H1m H2m. + case_eq (digits <= m)%int63. + intros Hlm; rewrite bit_M; auto; discriminate. + rewrite <- not_true_iff_false, leb_spec; intros Hlm. + case (Zle_lt_or_eq 0 [|m|]); auto; intros Hm. + replace m with ((m -1) + 1)%int63. + rewrite <-(bit_half x), <-(bit_half y); auto with zarith. + apply HH. + rewrite <-lor_lsr. + assert (0 <= [|bit (x lor y) 0|] <= 1) by (case bit; split; discriminate). + rewrite F in Heq; generalize (add_cancel_r _ _ _ Heq). + intros Heq1; apply to_Z_inj. + generalize (f_equal to_Z Heq1); rewrite lsl_spec, to_Z_1, Z.pow_1_r, Zmod_small. + rewrite lsl_spec, to_Z_1, Z.pow_1_r, Zmod_small; auto with zarith. + case (to_Z_bounded (x lor y)); intros H1xy H2xy. + rewrite lsr_spec, to_Z_1, Z.pow_1_r; auto with zarith. + change wB with ((wB/2)*2); split; auto with zarith. + assert ([|x lor y|] / 2 < wB / 2); auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + split. + case (to_Z_bounded (x >> 1 + y >> 1)); auto with zarith. + rewrite add_spec. + apply Z.le_lt_trans with (([|x >> 1|] + [|y >> 1|]) * 2); auto with zarith. + case (Zmod_le_first ([|x >> 1|] + [|y >> 1|]) wB); auto with zarith. + case (to_Z_bounded (x >> 1)); case (to_Z_bounded (y >> 1)); auto with zarith. + generalize Hb; rewrite (to_Z_split x) at 1; rewrite (to_Z_split y) at 1. + case (to_Z_bounded (bit x 0)); case (to_Z_bounded (bit y 0)); auto with zarith. + rewrite ltb_spec, sub_spec, to_Z_1, Zmod_small; auto with zarith. + rewrite ltb_spec, sub_spec, to_Z_1, Zmod_small; auto with zarith. + apply to_Z_inj. + rewrite add_spec, sub_spec, Zplus_mod_idemp_l, to_Z_1, Zmod_small; auto with zarith. + pose proof (to_Z_inj 0 _ Hm); clear Hm; subst m. + intros hx hy; revert F; rewrite hx, hy; intros F. generalize (f_equal to_Z F). vm_compute. lia. +Qed. + +Lemma addmuldiv_spec x y p : + [| p |] <= [| digits |] -> + [| addmuldiv p x y |] = ([| x |] * (2 ^ [| p |]) + [| y |] / (2 ^ ([| digits |] - [| p |]))) mod wB. +Proof. + intros H. + assert (Fp := to_Z_bounded p); assert (Fd := to_Z_bounded digits). + rewrite addmuldiv_def_spec; unfold addmuldiv_def. + case (bit_add_or (x << p) (y >> (digits - p))); intros HH _. + rewrite <-HH, add_spec, lsl_spec, lsr_spec, Zplus_mod_idemp_l, sub_spec. + rewrite (fun x y => Zmod_small (x - y)); auto with zarith. + intros n; rewrite -> bit_lsl, bit_lsr. + generalize (add_le_r (digits - p) n). + case (_ ≤ _); try discriminate. + rewrite -> sub_spec, Zmod_small; auto with zarith; intros H1. + case_eq (n < p)%int63; try discriminate. + rewrite <- not_true_iff_false, ltb_spec; intros H2. + case (_ ≤ _); try discriminate. + intros _; rewrite bit_M; try discriminate. + rewrite -> leb_spec, add_spec, Zmod_small, sub_spec, Zmod_small; auto with zarith. + rewrite -> sub_spec, Zmod_small; auto with zarith. +Qed. + +(* is_even *) +Lemma is_even_bit i : is_even i = negb (bit i 0). +Proof. + unfold is_even. + replace (i land 1) with (b2i (bit i 0)). + case bit; auto. + apply bit_ext; intros n. + rewrite bit_b2i, land_spec, bit_1. + generalize (eqb_spec n 0). + case (n == 0); auto. + intros(H,_); rewrite andb_true_r, H; auto. + rewrite andb_false_r; auto. +Qed. + +Lemma is_even_spec x : if is_even x then [|x|] mod 2 = 0 else [|x|] mod 2 = 1. +Proof. +rewrite is_even_bit. +generalize (bit_0_spec x); case bit; simpl; auto. +Qed. + +Lemma is_even_0 : is_even 0 = true. +Proof. apply refl_equal. Qed. + +Lemma is_even_lsl_1 i : is_even (i << 1) = true. +Proof. + rewrite is_even_bit, bit_lsl; auto. +Qed. + +(* Sqrt *) + + (* Direct transcription of an old proof + of a fortran program in boyer-moore *) + +Ltac elim_div := + unfold Z.div, Z.modulo; + match goal with + | H : context[ Z.div_eucl ?X ?Y ] |- _ => + generalize dependent H; generalize (Z_div_mod_full X Y) ; case (Z.div_eucl X Y) + | |- context[ Z.div_eucl ?X ?Y ] => + generalize (Z_div_mod_full X Y) ; case (Z.div_eucl X Y) + end; unfold Remainder. + +Lemma quotient_by_2 a: a - 1 <= (a/2) + (a/2). +Proof. + case (Z_mod_lt a 2); auto with zarith. + intros H1; rewrite Zmod_eq_full; auto with zarith. +Qed. + +Lemma sqrt_main_trick j k: 0 <= j -> 0 <= k -> + (j * k) + j <= ((j + k)/2 + 1) ^ 2. +Proof. + intros Hj; generalize Hj k; pattern j; apply natlike_ind; + auto; clear k j Hj. + intros _ k Hk; repeat rewrite Zplus_0_l. + apply Zmult_le_0_compat; generalize (Z_div_pos k 2); auto with zarith. + intros j Hj Hrec _ k Hk; pattern k; apply natlike_ind; auto; clear k Hk. + rewrite -> Zmult_0_r, Zplus_0_r, Zplus_0_l. + generalize (sqr_pos (Z.succ j / 2)) (quotient_by_2 (Z.succ j)); + unfold Z.succ. + rewrite Z.pow_2_r, Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r. + auto with zarith. + intros k Hk _. + replace ((Z.succ j + Z.succ k) / 2) with ((j + k)/2 + 1). + generalize (Hrec Hj k Hk) (quotient_by_2 (j + k)). + unfold Z.succ; repeat rewrite Z.pow_2_r; + repeat rewrite Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r. + repeat rewrite Zmult_1_l; repeat rewrite Zmult_1_r. + auto with zarith. + rewrite -> Zplus_comm, <- Z_div_plus_full_l; auto with zarith. + apply f_equal2; auto with zarith. +Qed. + +Lemma sqrt_main i j: 0 <= i -> 0 < j -> i < ((j + (i/j))/2 + 1) ^ 2. +Proof. + intros Hi Hj. + assert (Hij: 0 <= i/j) by (apply Z_div_pos; auto with zarith). + refine (Z.lt_le_trans _ _ _ _ (sqrt_main_trick _ _ (Zlt_le_weak _ _ Hj) Hij)). + pattern i at 1; rewrite -> (Z_div_mod_eq i j); case (Z_mod_lt i j); auto with zarith. +Qed. + +Lemma sqrt_test_false i j: 0 <= i -> 0 < j -> i/j < j -> (j + (i/j))/2 < j. +Proof. + intros Hi Hj; elim_div; intros q r [ ? hr ]; [ lia | subst i ]. + elim_div; intros a b [ h [ hb | ] ]; lia. +Qed. + +Lemma sqrt_test_true i j: 0 <= i -> 0 < j -> i/j >= j -> j ^ 2 <= i. +Proof. + intros Hi Hj Hd; rewrite Z.pow_2_r. + apply Z.le_trans with (j * (i/j)); auto with zarith. + apply Z_mult_div_ge; auto with zarith. +Qed. + +Lemma sqrt_step_correct rec i j: + 0 < [|i|] -> 0 < [|j|] -> [|i|] < ([|j|] + 1) ^ 2 -> + 2 * [|j|] < wB -> + (forall j1 : int, + 0 < [|j1|] < [|j|] -> [|i|] < ([|j1|] + 1) ^ 2 -> + [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> + [|sqrt_step rec i j|] ^ 2 <= [|i|] < ([|sqrt_step rec i j|] + 1) ^ 2. +Proof. + assert (Hp2: 0 < [|2|]) by exact (refl_equal Lt). + intros Hi Hj Hij H31 Hrec. + unfold sqrt_step. + case ltbP; rewrite div_spec. + - intros hlt. + assert ([| j + i / j|] = [|j|] + [|i|]/[|j|]) as hj. + rewrite add_spec, Zmod_small;rewrite div_spec; auto with zarith. + apply Hrec; rewrite lsr_spec, hj, to_Z_1; change (2 ^ 1) with 2. + + split; [ | apply sqrt_test_false;auto with zarith]. + replace ([|j|] + [|i|]/[|j|]) with (1 * 2 + (([|j|] - 2) + [|i|] / [|j|])) by ring. + rewrite Z_div_plus_full_l; auto with zarith. + assert (0 <= [|i|]/ [|j|]) by (apply Z_div_pos; auto with zarith). + assert (0 <= ([|j|] - 2 + [|i|] / [|j|]) / 2) ; auto with zarith. + apply Z.div_pos; [ | lia ]. + case (Zle_lt_or_eq 1 [|j|]); auto with zarith; intros Hj1. + rewrite <- Hj1, Zdiv_1_r; lia. + + apply sqrt_main;auto with zarith. + - split;[apply sqrt_test_true | ];auto with zarith. +Qed. + +Lemma iter_sqrt_correct n rec i j: 0 < [|i|] -> 0 < [|j|] -> + [|i|] < ([|j|] + 1) ^ 2 -> 2 * [|j|] < wB -> + (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] -> + [|i|] < ([|j1|] + 1) ^ 2 -> 2 * [|j1|] < wB -> + [|rec i j1|] ^ 2 <= [|i|] < ([|rec i j1|] + 1) ^ 2) -> + [|iter_sqrt n rec i j|] ^ 2 <= [|i|] < ([|iter_sqrt n rec i j|] + 1) ^ 2. +Proof. + revert rec i j; elim n; unfold iter_sqrt; fold iter_sqrt; clear n. + intros rec i j Hi Hj Hij H31 Hrec; apply sqrt_step_correct; auto with zarith. + intros; apply Hrec; auto with zarith. + rewrite Zpower_0_r; auto with zarith. + intros n Hrec rec i j Hi Hj Hij H31 HHrec. + apply sqrt_step_correct; auto. + intros j1 Hj1 Hjp1; apply Hrec; auto with zarith. + intros j2 Hj2 H2j2 Hjp2 Hj31; apply Hrec; auto with zarith. + intros j3 Hj3 Hpj3. + apply HHrec; auto. + rewrite -> inj_S, Z.pow_succ_r. + apply Z.le_trans with (2 ^Z_of_nat n + [|j2|]); auto with zarith. + apply Zle_0_nat. +Qed. + +Lemma sqrt_init i: 1 < i -> i < (i/2 + 1) ^ 2. +Proof. + intros Hi. + assert (H1: 0 <= i - 2) by auto with zarith. + assert (H2: 1 <= (i / 2) ^ 2); auto with zarith. + replace i with (1* 2 + (i - 2)); auto with zarith. + rewrite Z.pow_2_r, Z_div_plus_full_l; auto with zarith. + generalize (sqr_pos ((i - 2)/ 2)) (Z_div_pos (i - 2) 2). + rewrite Zmult_plus_distr_l; repeat rewrite Zmult_plus_distr_r. + auto with zarith. + generalize (quotient_by_2 i). + rewrite -> Z.pow_2_r in H2 |- *; + repeat (rewrite Zmult_plus_distr_l || + rewrite Zmult_plus_distr_r || + rewrite Zmult_1_l || rewrite Zmult_1_r). + auto with zarith. +Qed. + +Lemma sqrt_spec : forall x, + [|sqrt x|] ^ 2 <= [|x|] < ([|sqrt x|] + 1) ^ 2. +Proof. + intros i; unfold sqrt. + rewrite compare_spec. case Z.compare_spec; rewrite to_Z_1; + intros Hi; auto with zarith. + repeat rewrite Z.pow_2_r; auto with zarith. + apply iter_sqrt_correct; auto with zarith; + rewrite lsr_spec, to_Z_1; change (2^1) with 2; auto with zarith. + replace [|i|] with (1 * 2 + ([|i|] - 2))%Z; try ring. + assert (0 <= ([|i|] - 2)/2)%Z by (apply Z_div_pos; auto with zarith). + rewrite Z_div_plus_full_l; auto with zarith. + apply sqrt_init; auto. + assert (W:= Z_mult_div_ge [|i|] 2);assert (W':= to_Z_bounded i);auto with zarith. + intros j2 H1 H2; contradict H2; apply Zlt_not_le. + fold wB;assert (W:=to_Z_bounded i). + apply Z.le_lt_trans with ([|i|]); auto with zarith. + assert (0 <= [|i|]/2)%Z by (apply Z_div_pos; auto with zarith). + apply Z.le_trans with (2 * ([|i|]/2)); auto with zarith. + apply Z_mult_div_ge; auto with zarith. + case (to_Z_bounded i); repeat rewrite Z.pow_2_r; auto with zarith. +Qed. + +(* sqrt2 *) +Lemma sqrt2_step_def rec ih il j: + sqrt2_step rec ih il j = + if (ih < j)%int63 then + let quo := fst (diveucl_21 ih il j) in + if (quo < j)%int63 then + let m := + match j +c quo with + | C0 m1 => m1 >> 1 + | C1 m1 => (m1 >> 1 + 1 << (digits -1))%int63 + end in + rec ih il m + else j + else j. +Proof. + unfold sqrt2_step; case diveucl_21; intros;simpl. + case (j +c i);trivial. +Qed. + +Lemma sqrt2_lower_bound ih il j: + [|| WW ih il||] < ([|j|] + 1) ^ 2 -> [|ih|] <= [|j|]. +Proof. + intros H1. + case (to_Z_bounded j); intros Hbj _. + case (to_Z_bounded il); intros Hbil _. + case (to_Z_bounded ih); intros Hbih Hbih1. + assert (([|ih|] < [|j|] + 1)%Z); auto with zarith. + apply Zlt_square_simpl; auto with zarith. + simpl zn2z_to_Z in H1. + repeat rewrite <-Z.pow_2_r. + refine (Z.le_lt_trans _ _ _ _ H1). + apply Z.le_trans with ([|ih|] * wB)%Z;try rewrite Z.pow_2_r; auto with zarith. +Qed. + +Lemma div2_phi ih il j: + [|fst (diveucl_21 ih il j)|] = [|| WW ih il||] /[|j|]. +Proof. + generalize (diveucl_21_spec ih il j). + case diveucl_21; intros q r Heq. + simpl zn2z_to_Z;unfold Z.div;rewrite <- Heq;trivial. +Qed. + +Lemma sqrt2_step_correct rec ih il j: + 2 ^ (Z_of_nat (size - 2)) <= [|ih|] -> + 0 < [|j|] -> [|| WW ih il||] < ([|j|] + 1) ^ 2 -> + (forall j1, 0 < [|j1|] < [|j|] -> [|| WW ih il||] < ([|j1|] + 1) ^ 2 -> + [|rec ih il j1|] ^ 2 <= [||WW ih il||] < ([|rec ih il j1|] + 1) ^ 2) -> + [|sqrt2_step rec ih il j|] ^ 2 <= [||WW ih il ||] + < ([|sqrt2_step rec ih il j|] + 1) ^ 2. +Proof. + assert (Hp2: (0 < [|2|])%Z) by exact (refl_equal Lt). + intros Hih Hj Hij Hrec; rewrite sqrt2_step_def. + assert (H1: ([|ih|] <= [|j|])%Z) by (apply sqrt2_lower_bound with il; auto). + case (to_Z_bounded ih); intros Hih1 _. + case (to_Z_bounded il); intros Hil1 _. + case (to_Z_bounded j); intros _ Hj1. + assert (Hp3: (0 < [||WW ih il||])). + simpl zn2z_to_Z;apply Z.lt_le_trans with ([|ih|] * wB)%Z; auto with zarith. + apply Zmult_lt_0_compat; auto with zarith. + refine (Z.lt_le_trans _ _ _ _ Hih); auto with zarith. + cbv zeta. + case_eq (ih < j)%int63;intros Heq. + rewrite -> ltb_spec in Heq. + 2: rewrite <-not_true_iff_false, ltb_spec in Heq. + 2: split; auto. + 2: apply sqrt_test_true; auto with zarith. + 2: unfold zn2z_to_Z; replace [|ih|] with [|j|]; auto with zarith. + 2: assert (0 <= [|il|]/[|j|]) by (apply Z_div_pos; auto with zarith). + 2: rewrite Zmult_comm, Z_div_plus_full_l; unfold base; auto with zarith. + case (Zle_or_lt (2^(Z_of_nat size -1)) [|j|]); intros Hjj. + case_eq (fst (diveucl_21 ih il j) < j)%int63;intros Heq0. + 2: rewrite <-not_true_iff_false, ltb_spec, div2_phi in Heq0. + 2: split; auto; apply sqrt_test_true; auto with zarith. + rewrite -> ltb_spec, div2_phi in Heq0. + match goal with |- context[rec _ _ ?X] => + set (u := X) + end. + assert (H: [|u|] = ([|j|] + ([||WW ih il||])/([|j|]))/2). + unfold u; generalize (addc_spec j (fst (diveucl_21 ih il j))); + case addc;unfold interp_carry;rewrite div2_phi;simpl zn2z_to_Z. + intros i H;rewrite lsr_spec, H;trivial. + intros i H;rewrite <- H. + case (to_Z_bounded i); intros H1i H2i. + rewrite -> add_spec, Zmod_small, lsr_spec. + change (1 * wB) with ([|(1 << (digits -1))|] * 2)%Z. + rewrite Z_div_plus_full_l; auto with zarith. + change wB with (2 * (wB/2))%Z; auto. + replace [|(1 << (digits - 1))|] with (wB/2); auto. + rewrite lsr_spec; auto. + replace (2^[|1|]) with 2%Z; auto. + split; auto with zarith. + assert ([|i|]/2 < wB/2); auto with zarith. + apply Zdiv_lt_upper_bound; auto with zarith. + apply Hrec; rewrite H; clear u H. + assert (Hf1: 0 <= [||WW ih il||]/ [|j|]) by (apply Z_div_pos; auto with zarith). + case (Zle_lt_or_eq 1 ([|j|])); auto with zarith; intros Hf2. + 2: contradict Heq0; apply Zle_not_lt; rewrite <- Hf2, Zdiv_1_r; auto with zarith. + split. + replace ([|j|] + [||WW ih il||]/ [|j|])%Z with + (1 * 2 + (([|j|] - 2) + [||WW ih il||] / [|j|])) by lia. + rewrite Z_div_plus_full_l; auto with zarith. + assert (0 <= ([|j|] - 2 + [||WW ih il||] / [|j|]) / 2) ; auto with zarith. + apply sqrt_test_false; auto with zarith. + apply sqrt_main; auto with zarith. + contradict Hij; apply Zle_not_lt. + assert ((1 + [|j|]) <= 2 ^ (Z_of_nat size - 1)); auto with zarith. + apply Z.le_trans with ((2 ^ (Z_of_nat size - 1)) ^2); auto with zarith. + assert (0 <= 1 + [|j|]); auto with zarith. + apply Zmult_le_compat; auto with zarith. + change ((2 ^ (Z_of_nat size - 1))^2) with (2 ^ (Z_of_nat size - 2) * wB). + apply Z.le_trans with ([|ih|] * wB); auto with zarith. + unfold zn2z_to_Z, wB; auto with zarith. +Qed. + +Lemma iter2_sqrt_correct n rec ih il j: + 2^(Z_of_nat (size - 2)) <= [|ih|] -> 0 < [|j|] -> [||WW ih il||] < ([|j|] + 1) ^ 2 -> + (forall j1, 0 < [|j1|] -> 2^(Z_of_nat n) + [|j1|] <= [|j|] -> + [||WW ih il||] < ([|j1|] + 1) ^ 2 -> + [|rec ih il j1|] ^ 2 <= [||WW ih il||] < ([|rec ih il j1|] + 1) ^ 2) -> + [|iter2_sqrt n rec ih il j|] ^ 2 <= [||WW ih il||] + < ([|iter2_sqrt n rec ih il j|] + 1) ^ 2. +Proof. + revert rec ih il j; elim n; unfold iter2_sqrt; fold iter2_sqrt; clear n. + intros rec ih il j Hi Hj Hij Hrec; apply sqrt2_step_correct; auto with zarith. + intros; apply Hrec; auto with zarith. + rewrite Zpower_0_r; auto with zarith. + intros n Hrec rec ih il j Hi Hj Hij HHrec. + apply sqrt2_step_correct; auto. + intros j1 Hj1 Hjp1; apply Hrec; auto with zarith. + intros j2 Hj2 H2j2 Hjp2; apply Hrec; auto with zarith. + intros j3 Hj3 Hpj3. + apply HHrec; auto. + rewrite -> inj_S, Z.pow_succ_r. + apply Z.le_trans with (2 ^Z_of_nat n + [|j2|])%Z; auto with zarith. + apply Zle_0_nat. +Qed. + +Lemma sqrt2_spec : forall x y, + wB/ 4 <= [|x|] -> + let (s,r) := sqrt2 x y in + [||WW x y||] = [|s|] ^ 2 + [+|r|] /\ + [+|r|] <= 2 * [|s|]. + Proof. + intros ih il Hih; unfold sqrt2. + change [||WW ih il||] with ([||WW ih il||]). + assert (Hbin: forall s, s * s + 2* s + 1 = (s + 1) ^ 2) by + (intros s; ring). + assert (Hb: 0 <= wB) by (red; intros HH; discriminate). + assert (Hi2: [||WW ih il ||] < ([|max_int|] + 1) ^ 2). + apply Z.le_lt_trans with ((wB - 1) * wB + (wB - 1)); auto with zarith. + 2: apply refl_equal. + case (to_Z_bounded ih); case (to_Z_bounded il); intros H1 H2 H3 H4. + unfold zn2z_to_Z; auto with zarith. + case (iter2_sqrt_correct size (fun _ _ j => j) ih il max_int); auto with zarith. + apply refl_equal. + intros j1 _ HH; contradict HH. + apply Zlt_not_le. + case (to_Z_bounded j1); auto with zarith. + change (2 ^ Z_of_nat size) with ([|max_int|]+1)%Z; auto with zarith. + set (s := iter2_sqrt size (fun _ _ j : int=> j) ih il max_int). + intros Hs1 Hs2. + generalize (mulc_spec s s); case mulc. + simpl fst; simpl snd; intros ih1 il1 Hihl1. + generalize (subc_spec il il1). + case subc; intros il2 Hil2. + simpl interp_carry in Hil2. + case_eq (ih1 < ih)%int63; [idtac | rewrite <- not_true_iff_false]; + rewrite ltb_spec; intros Heq. + unfold interp_carry; rewrite Zmult_1_l. + rewrite -> Z.pow_2_r, Hihl1, Hil2. + case (Zle_lt_or_eq ([|ih1|] + 1) ([|ih|])); auto with zarith. + intros H2; contradict Hs2; apply Zle_not_lt. + replace (([|s|] + 1) ^ 2) with ([||WW ih1 il1||] + 2 * [|s|] + 1). + unfold zn2z_to_Z. + case (to_Z_bounded il); intros Hpil _. + assert (Hl1l: [|il1|] <= [|il|]). + case (to_Z_bounded il2); rewrite Hil2; auto with zarith. + assert ([|ih1|] * wB + 2 * [|s|] + 1 <= [|ih|] * wB); auto with zarith. + case (to_Z_bounded s); intros _ Hps. + case (to_Z_bounded ih1); intros Hpih1 _; auto with zarith. + apply Z.le_trans with (([|ih1|] + 2) * wB); auto with zarith. + rewrite Zmult_plus_distr_l. + assert (2 * [|s|] + 1 <= 2 * wB); auto with zarith. + unfold zn2z_to_Z; rewrite <-Hihl1, Hbin; auto. + intros H2; split. + unfold zn2z_to_Z; rewrite <- H2; ring. + replace (wB + ([|il|] - [|il1|])) with ([||WW ih il||] - ([|s|] * [|s|])). + rewrite <-Hbin in Hs2; auto with zarith. + rewrite Hihl1; unfold zn2z_to_Z; rewrite <- H2; ring. + unfold interp_carry. + case (Zle_lt_or_eq [|ih|] [|ih1|]); auto with zarith; intros H. + contradict Hs1. + apply Zlt_not_le; rewrite Z.pow_2_r, Hihl1. + unfold zn2z_to_Z. + case (to_Z_bounded il); intros _ H2. + apply Z.lt_le_trans with (([|ih|] + 1) * wB + 0). + rewrite Zmult_plus_distr_l, Zplus_0_r; auto with zarith. + case (to_Z_bounded il1); intros H3 _. + apply Zplus_le_compat; auto with zarith. + split. + rewrite Z.pow_2_r, Hihl1. + unfold zn2z_to_Z; ring[Hil2 H]. + replace [|il2|] with ([||WW ih il||] - [||WW ih1 il1||]). + unfold zn2z_to_Z at 2; rewrite <-Hihl1. + rewrite <-Hbin in Hs2; auto with zarith. + unfold zn2z_to_Z; rewrite H, Hil2; ring. + unfold interp_carry in Hil2 |- *. + assert (Hsih: [|ih - 1|] = [|ih|] - 1). + rewrite sub_spec, Zmod_small; auto; replace [|1|] with 1; auto. + case (to_Z_bounded ih); intros H1 H2. + split; auto with zarith. + apply Z.le_trans with (wB/4 - 1); auto with zarith. + case_eq (ih1 < ih - 1)%int63; [idtac | rewrite <- not_true_iff_false]; + rewrite ltb_spec, Hsih; intros Heq. + rewrite Z.pow_2_r, Hihl1. + case (Zle_lt_or_eq ([|ih1|] + 2) [|ih|]); auto with zarith. + intros H2; contradict Hs2; apply Zle_not_lt. + replace (([|s|] + 1) ^ 2) with ([||WW ih1 il1||] + 2 * [|s|] + 1). + unfold zn2z_to_Z. + assert ([|ih1|] * wB + 2 * [|s|] + 1 <= [|ih|] * wB + ([|il|] - [|il1|])); + auto with zarith. + rewrite <-Hil2. + case (to_Z_bounded il2); intros Hpil2 _. + apply Z.le_trans with ([|ih|] * wB + - wB); auto with zarith. + case (to_Z_bounded s); intros _ Hps. + assert (2 * [|s|] + 1 <= 2 * wB); auto with zarith. + apply Z.le_trans with ([|ih1|] * wB + 2 * wB); auto with zarith. + assert (Hi: ([|ih1|] + 3) * wB <= [|ih|] * wB); auto with zarith. + rewrite Zmult_plus_distr_l in Hi; auto with zarith. + unfold zn2z_to_Z; rewrite <-Hihl1, Hbin; auto. + intros H2; unfold zn2z_to_Z; rewrite <-H2. + split. + replace [|il|] with (([|il|] - [|il1|]) + [|il1|]); try ring. + rewrite <-Hil2; ring. + replace (1 * wB + [|il2|]) with ([||WW ih il||] - [||WW ih1 il1||]). + unfold zn2z_to_Z at 2; rewrite <-Hihl1. + rewrite <-Hbin in Hs2; auto with zarith. + unfold zn2z_to_Z; rewrite <-H2. + replace [|il|] with (([|il|] - [|il1|]) + [|il1|]); try ring. + rewrite <-Hil2; ring. + case (Zle_lt_or_eq ([|ih|] - 1) ([|ih1|])); auto with zarith; intros H1. + assert (He: [|ih|] = [|ih1|]). + apply Zle_antisym; auto with zarith. + case (Zle_or_lt [|ih1|] [|ih|]); auto; intros H2. + contradict Hs1; apply Zlt_not_le; rewrite Z.pow_2_r, Hihl1. + unfold zn2z_to_Z. + case (to_Z_bounded il); intros _ Hpil1. + apply Z.lt_le_trans with (([|ih|] + 1) * wB). + rewrite Zmult_plus_distr_l, Zmult_1_l; auto with zarith. + case (to_Z_bounded il1); intros Hpil2 _. + apply Z.le_trans with (([|ih1|]) * wB); auto with zarith. + contradict Hs1; apply Zlt_not_le; rewrite Z.pow_2_r, Hihl1. + unfold zn2z_to_Z; rewrite He. + assert ([|il|] - [|il1|] < 0); auto with zarith. + rewrite <-Hil2. + case (to_Z_bounded il2); auto with zarith. + split. + rewrite Z.pow_2_r, Hihl1. + unfold zn2z_to_Z; rewrite <-H1. + apply trans_equal with ([|ih|] * wB + [|il1|] + ([|il|] - [|il1|])). + ring. + rewrite <-Hil2; ring. + replace [|il2|] with ([||WW ih il||] - [||WW ih1 il1||]). + unfold zn2z_to_Z at 2; rewrite <- Hihl1. + rewrite <-Hbin in Hs2; auto with zarith. + unfold zn2z_to_Z. + rewrite <-H1. + ring_simplify. + apply trans_equal with (wB + ([|il|] - [|il1|])). + ring. + rewrite <-Hil2; ring. +Qed. + +(* of_pos *) +Lemma of_pos_rec_spec (k: nat) : + (k <= size)%nat → + ∀ p, φ(of_pos_rec k p) = Zpos p mod 2 ^ Z.of_nat k. +Proof. + elim k; clear k. + intros _ p; simpl; rewrite to_Z_0, Zmod_1_r; reflexivity. + intros n ih hn. + assert (n <= size)%nat as hn' by lia. + specialize (ih hn'). + intros [ p | p | ]. + 3: { + rewrite Zmod_small. reflexivity. + split. lia. + apply Zpower_gt_1; lia. + } + - simpl. + destruct (bit_add_or (of_pos_rec n p << 1) 1) as (H1, _). + rewrite <- H1;clear H1. + 2: { + intros i; rewrite bit_lsl, bit_1. + case eqbP. + + intros h; apply to_Z_inj in h; subst. exact (λ e _, diff_false_true e). + + exact (λ _ _, diff_false_true). + } + rewrite add_spec, lsl_spec, ih, to_Z_1; clear ih. + rewrite Z.pow_pos_fold, Zpos_P_of_succ_nat. + change (Zpos p~1) with (2 ^ 1 * Zpos p + 1)%Z. + rewrite Zmod_distr by lia. + rewrite Zpower_Zsucc by auto with zarith. + rewrite Zplus_mod_idemp_l. + rewrite Zmod_small. + rewrite Zmult_mod_distr_l; lia. + set (a := Z.of_nat n). + set (b := Zpos p). + change (2 ^ 1) with 2. + pose proof (pow2_pos a (Nat2Z.is_nonneg _)). + elim_div; intros x y [ ? ha]. lia. + destruct ha as [ ha | ]. 2: lia. + split. lia. + apply Z.lt_le_trans with (2 ^ (a + 1)). + 2: apply Z.pow_le_mono_r; subst a; lia. + fold (Z.succ a); rewrite Z.pow_succ_r. lia. + subst a; lia. + - simpl. rewrite lsl_spec, ih, to_Z_1, Zmod_small. + rewrite Z.pow_pos_fold, Zpos_P_of_succ_nat, Zpower_Zsucc by lia. + change (Zpos p~0) with (2 ^ 1 * Zpos p)%Z. + rewrite Z.mul_mod_distr_l; auto with zarith. + set (a := Z.of_nat n). + set (b := Zpos p). + change (2 ^ 1) with 2. + pose proof (pow2_pos a (Nat2Z.is_nonneg _)). + elim_div; intros x y [ ? ha]. lia. + destruct ha as [ ha | ]. 2: lia. + split. lia. + apply Z.lt_le_trans with (2 ^ (a + 1)). + 2: apply Z.pow_le_mono_r; subst a; lia. + fold (Z.succ a); rewrite Z.pow_succ_r. lia. + subst a; lia. +Qed. + +Lemma is_int n : + 0 <= n < 2 ^ φ digits → + n = φ (of_Z n). +Proof. + destruct n. reflexivity. 2: lia. + intros [_ h]. simpl. + unfold of_pos. rewrite of_pos_rec_spec by lia. + symmetry; apply Z.mod_small. split. lia. exact h. +Qed. + +Lemma of_Z_spec n : [| of_Z n |] = n mod wB. +Proof. + destruct n. reflexivity. + { now simpl; unfold of_pos; rewrite of_pos_rec_spec by lia. } + simpl; unfold of_pos; rewrite opp_spec. + rewrite of_pos_rec_spec; [ |auto]; fold wB. + now rewrite <-(Z.sub_0_l), Zminus_mod_idemp_r. +Qed. + +(* General lemmas *) +Lemma negbE a b : a = negb b → negb a = b. +Proof. intros ->; apply negb_involutive. Qed. + +Lemma Z_oddE a : Z.odd a = (a mod 2 =? 1)%Z. +Proof. rewrite Zmod_odd; case Z.odd; reflexivity. Qed. + +Lemma Z_evenE a : Z.even a = (a mod 2 =? 0)%Z. +Proof. rewrite Zmod_even; case Z.even; reflexivity. Qed. + +(* is_zero *) +Lemma is_zeroE n : is_zero n = Z.eqb (φ n) 0. +Proof. + case Z.eqb_spec. + - intros h; apply (to_Z_inj n 0) in h; subst n; reflexivity. + - generalize (proj1 (is_zero_spec n)). + case is_zero; auto; intros ->; auto; destruct 1; reflexivity. +Qed. + +(* bit *) +Lemma bitE i j : bit i j = Z.testbit φ(i) φ(j). +Proof. + apply negbE; rewrite is_zeroE, lsl_spec, lsr_spec. + generalize (φ i) (to_Z_bounded i) (φ j) (to_Z_bounded j); clear i j; + intros i [hi hi'] j [hj hj']. + rewrite Z.testbit_eqb by auto; rewrite <- Z_oddE, Z.negb_odd, Z_evenE. + remember (i / 2 ^ j) as k. + change wB with (2 * 2 ^ φ (digits - 1)). + unfold Z.modulo at 2. + generalize (Z_div_mod_full k 2 (λ k, let 'eq_refl := k in I)); unfold Remainder. + destruct Z.div_eucl as [ p q ]; intros [hk [ hq | ]]. 2: lia. + rewrite hk. + remember φ (digits - 1) as m. + replace ((_ + _) * _) with (q * 2 ^ m + p * (2 * 2 ^ m)) by ring. + rewrite Z_mod_plus by (subst m; reflexivity). + assert (q = 0 ∨ q = 1) as D by lia. + destruct D; subst; reflexivity. +Qed. + +(* land, lor, lxor *) +Lemma lt_pow_lt_log d k n : + 0 < d <= n → + 0 <= k < 2 ^ d → + Z.log2 k < n. +Proof. + intros [hd hdn] [hk hkd]. + assert (k = 0 ∨ 0 < k) as D by lia. + clear hk; destruct D as [ hk | hk ]. + - subst k; simpl; lia. + - apply Z.log2_lt_pow2. lia. + eapply Z.lt_le_trans. eassumption. + apply Z.pow_le_mono_r; lia. +Qed. + +Lemma land_spec' x y : φ (x land y) = Z.land φ(x) φ(y). +Proof. + apply Z.bits_inj'; intros n hn. + destruct (to_Z_bounded (x land y)) as [ hxy hxy' ]. + destruct (to_Z_bounded x) as [ hx hx' ]. + destruct (to_Z_bounded y) as [ hy hy' ]. + case (Z_lt_le_dec n (φ digits)); intros hd. + 2: { + rewrite !Z.bits_above_log2; auto. + - apply Z.land_nonneg; auto. + - eapply Z.le_lt_trans. + apply Z.log2_land; assumption. + apply Z.min_lt_iff. + left. apply (lt_pow_lt_log φ digits). exact (conj eq_refl hd). + split; assumption. + - apply (lt_pow_lt_log φ digits). exact (conj eq_refl hd). + split; assumption. + } + rewrite (is_int n). + rewrite Z.land_spec, <- !bitE, land_spec; reflexivity. + apply (conj hn). + apply (Z.lt_trans _ _ _ hd). + apply Zpower2_lt_lin. lia. +Qed. + +Lemma lor_spec' x y : φ (x lor y) = Z.lor φ(x) φ(y). +Proof. + apply Z.bits_inj'; intros n hn. + destruct (to_Z_bounded (x lor y)) as [ hxy hxy' ]. + destruct (to_Z_bounded x) as [ hx hx' ]. + destruct (to_Z_bounded y) as [ hy hy' ]. + case (Z_lt_le_dec n (φ digits)); intros hd. + 2: { + rewrite !Z.bits_above_log2; auto. + - apply Z.lor_nonneg; auto. + - rewrite Z.log2_lor by assumption. + apply Z.max_lub_lt; apply (lt_pow_lt_log φ digits); split; assumption || reflexivity. + - apply (lt_pow_lt_log φ digits); split; assumption || reflexivity. + } + rewrite (is_int n). + rewrite Z.lor_spec, <- !bitE, lor_spec; reflexivity. + apply (conj hn). + apply (Z.lt_trans _ _ _ hd). + apply Zpower2_lt_lin. lia. +Qed. + +Lemma lxor_spec' x y : φ (x lxor y) = Z.lxor φ(x) φ(y). +Proof. + apply Z.bits_inj'; intros n hn. + destruct (to_Z_bounded (x lxor y)) as [ hxy hxy' ]. + destruct (to_Z_bounded x) as [ hx hx' ]. + destruct (to_Z_bounded y) as [ hy hy' ]. + case (Z_lt_le_dec n (φ digits)); intros hd. + 2: { + rewrite !Z.bits_above_log2; auto. + - apply Z.lxor_nonneg; split; auto. + - eapply Z.le_lt_trans. + apply Z.log2_lxor; assumption. + apply Z.max_lub_lt; apply (lt_pow_lt_log φ digits); split; assumption || reflexivity. + - apply (lt_pow_lt_log φ digits); split; assumption || reflexivity. + } + rewrite (is_int n). + rewrite Z.lxor_spec, <- !bitE, lxor_spec; reflexivity. + apply (conj hn). + apply (Z.lt_trans _ _ _ hd). + apply Zpower2_lt_lin. lia. +Qed. + +Lemma landC i j : i land j = j land i. +Proof. + apply bit_ext; intros n. + rewrite !land_spec, andb_comm; auto. +Qed. + +Lemma landA i j k : i land (j land k) = i land j land k. +Proof. + apply bit_ext; intros n. + rewrite !land_spec, andb_assoc; auto. +Qed. + +Lemma land0 i : 0 land i = 0%int63. +Proof. + apply bit_ext; intros n. + rewrite land_spec, bit_0; auto. +Qed. + +Lemma land0_r i : i land 0 = 0%int63. +Proof. rewrite landC; exact (land0 i). Qed. + +Lemma lorC i j : i lor j = j lor i. +Proof. + apply bit_ext; intros n. + rewrite !lor_spec, orb_comm; auto. +Qed. + +Lemma lorA i j k : i lor (j lor k) = i lor j lor k. +Proof. + apply bit_ext; intros n. + rewrite !lor_spec, orb_assoc; auto. +Qed. + +Lemma lor0 i : 0 lor i = i. +Proof. + apply bit_ext; intros n. + rewrite lor_spec, bit_0; auto. +Qed. + +Lemma lor0_r i : i lor 0 = i. +Proof. rewrite lorC; exact (lor0 i). Qed. + +Lemma lxorC i j : i lxor j = j lxor i. +Proof. + apply bit_ext; intros n. + rewrite !lxor_spec, xorb_comm; auto. +Qed. + +Lemma lxorA i j k : i lxor (j lxor k) = i lxor j lxor k. +Proof. + apply bit_ext; intros n. + rewrite !lxor_spec, xorb_assoc; auto. +Qed. + +Lemma lxor0 i : 0 lxor i = i. +Proof. + apply bit_ext; intros n. + rewrite lxor_spec, bit_0, xorb_false_l; auto. +Qed. + +Lemma lxor0_r i : i lxor 0 = i. +Proof. rewrite lxorC; exact (lxor0 i). Qed. diff --git a/theories/Numbers/Cyclic/Int63/Ring63.v b/theories/Numbers/Cyclic/Int63/Ring63.v new file mode 100644 index 0000000000..d230435378 --- /dev/null +++ b/theories/Numbers/Cyclic/Int63/Ring63.v @@ -0,0 +1,65 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2018 *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(************************************************************************) + +(** * Int63 numbers defines Z/(2^63)Z, and can hence be equipped + with a ring structure and a ring tactic *) + +Require Import Cyclic63 CyclicAxioms. + +Local Open Scope int63_scope. + +(** Detection of constants *) + +Ltac isInt63cst t := + match eval lazy delta [add] in (t + 1)%int63 with + | add _ _ => constr:(false) + | _ => constr:(true) + end. + +Ltac Int63cst t := + match eval lazy delta [add] in (t + 1)%int63 with + | add _ _ => constr:(NotConstant) + | _ => constr:(t) + end. + +(** The generic ring structure inferred from the Cyclic structure *) + +Module Int63ring := CyclicRing Int63Cyclic. + +(** Unlike in the generic [CyclicRing], we can use Leibniz here. *) + +Lemma Int63_canonic : forall x y, to_Z x = to_Z y -> x = y. +Proof to_Z_inj. + +Lemma ring_theory_switch_eq : + forall A (R R':A->A->Prop) zero one add mul sub opp, + (forall x y : A, R x y -> R' x y) -> + ring_theory zero one add mul sub opp R -> + ring_theory zero one add mul sub opp R'. +Proof. +intros A R R' zero one add mul sub opp Impl Ring. +constructor; intros; apply Impl; apply Ring. +Qed. + +Lemma Int63Ring : ring_theory 0 1 add mul sub opp Logic.eq. +Proof. +exact (ring_theory_switch_eq _ _ _ _ _ _ _ _ _ Int63_canonic Int63ring.CyclicRing). +Qed. + +Lemma eq31_correct : forall x y, eqb x y = true -> x=y. +Proof. now apply eqb_spec. Qed. + +Add Ring Int63Ring : Int63Ring + (decidable eq31_correct, + constants [Int63cst]). + +Section TestRing. +Let test : forall x y, 1 + x*y + x*x + 1 = 1*1 + 1 + y*x + 1*x*x. +intros. ring. +Qed. +End TestRing. diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index d9787bc73c..868a6ed3e9 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -245,6 +245,7 @@ let build_beq_scheme mode kn = | Fix _ -> raise (EqUnknown "fix") | Meta _ -> raise (EqUnknown "meta-variable") | Evar _ -> raise (EqUnknown "existential variable") + | Int _ -> raise (EqUnknown "int") in aux t in diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml index 4b8371f5c3..7301e1fff7 100644 --- a/vernac/comAssumption.ml +++ b/vernac/comAssumption.ml @@ -178,3 +178,29 @@ let do_assumptions kind nl l = in subst'@subst, status' && status, next_uctx uctx) ([], true, uctx) l) + +let do_primitive id prim typopt = + if Lib.sections_are_opened () then + CErrors.user_err Pp.(str "Declaring a primitive is not allowed in sections."); + if Dumpglob.dump () then Dumpglob.dump_definition id false "ax"; + let env = Global.env () in + let evd = Evd.from_env env in + let evd, typopt = Option.fold_left_map + (interp_type_evars_impls ~impls:empty_internalization_env env) + evd typopt + in + let evd = Evd.minimize_universes evd in + let uvars, impls, typopt = match typopt with + | None -> Univ.LSet.empty, [], None + | Some (ty,impls) -> + EConstr.universes_of_constr evd ty, impls, Some (EConstr.to_constr evd ty) + in + let evd = Evd.restrict_universe_context evd uvars in + let uctx = UState.check_mono_univ_decl (Evd.evar_universe_context evd) UState.default_univ_decl in + let entry = { prim_entry_type = typopt; + prim_entry_univs = uctx; + prim_entry_content = prim; + } + in + let _kn = declare_constant id.CAst.v (PrimitiveEntry entry,IsPrimitive) in + register_message id.CAst.v diff --git a/vernac/comAssumption.mli b/vernac/comAssumption.mli index 56932360a9..c5bf3725a9 100644 --- a/vernac/comAssumption.mli +++ b/vernac/comAssumption.mli @@ -33,3 +33,5 @@ val declare_assumption : coercion_flag -> assumption_kind -> UnivNames.universe_binders -> Impargs.manual_implicits -> bool (** implicit *) -> Declaremods.inline -> variable CAst.t -> GlobRef.t * Univ.Instance.t * bool + +val do_primitive : lident -> CPrimitives.op_or_type -> constr_expr option -> unit diff --git a/vernac/g_vernac.mlg b/vernac/g_vernac.mlg index 79adefdcf7..0664e18130 100644 --- a/vernac/g_vernac.mlg +++ b/vernac/g_vernac.mlg @@ -219,12 +219,51 @@ GRAMMAR EXTEND Gram { VernacRegister(g, RegisterCoqlib quid) } | IDENT "Register"; IDENT "Inline"; g = global -> { VernacRegister(g, RegisterInline) } + | IDENT "Primitive"; id = identref; typopt = OPT [ ":"; typ = lconstr -> { typ } ]; ":="; r = register_token -> + { VernacPrimitive(id, r, typopt) } | IDENT "Universe"; l = LIST1 identref -> { VernacUniverse l } | IDENT "Universes"; l = LIST1 identref -> { VernacUniverse l } | IDENT "Constraint"; l = LIST1 univ_constraint SEP "," -> { VernacConstraint l } ] ] ; + register_token: + [ [ r = register_prim_token -> { CPrimitives.OT_op r } + | r = register_type_token -> { CPrimitives.OT_type r } ] ] + ; + + register_type_token: + [ [ "#int63_type" -> { CPrimitives.PT_int63 } ] ] + ; + + register_prim_token: + [ [ "#int63_head0" -> { CPrimitives.Int63head0 } + | "#int63_tail0" -> { CPrimitives.Int63tail0 } + | "#int63_add" -> { CPrimitives.Int63add } + | "#int63_sub" -> { CPrimitives.Int63sub } + | "#int63_mul" -> { CPrimitives.Int63mul } + | "#int63_div" -> { CPrimitives.Int63div } + | "#int63_mod" -> { CPrimitives.Int63mod } + | "#int63_lsr" -> { CPrimitives.Int63lsr } + | "#int63_lsl" -> { CPrimitives.Int63lsl } + | "#int63_land" -> { CPrimitives.Int63land } + | "#int63_lor" -> { CPrimitives.Int63lor } + | "#int63_lxor" -> { CPrimitives.Int63lxor } + | "#int63_addc" -> { CPrimitives.Int63addc } + | "#int63_subc" -> { CPrimitives.Int63subc } + | "#int63_addcarryc" -> { CPrimitives.Int63addCarryC } + | "#int63_subcarryc" -> { CPrimitives.Int63subCarryC } + | "#int63_mulc" -> { CPrimitives.Int63mulc } + | "#int63_diveucl" -> { CPrimitives.Int63diveucl } + | "#int63_div21" -> { CPrimitives.Int63div21 } + | "#int63_addmuldiv" -> { CPrimitives.Int63addMulDiv } + | "#int63_eq" -> { CPrimitives.Int63eq } + | "#int63_lt" -> { CPrimitives.Int63lt } + | "#int63_le" -> { CPrimitives.Int63le } + | "#int63_compare" -> { CPrimitives.Int63compare } + ] ] + ; + thm_token: [ [ "Theorem" -> { Theorem } | IDENT "Lemma" -> { Lemma } diff --git a/vernac/ppvernac.ml b/vernac/ppvernac.ml index 5eeeaada2d..d22e52e960 100644 --- a/vernac/ppvernac.ml +++ b/vernac/ppvernac.ml @@ -1182,6 +1182,12 @@ open Pputils hov 2 (keyword "Register Inline" ++ spc() ++ pr_qualid qid) ) + | VernacPrimitive(id,r,typopt) -> + hov 2 + (keyword "Primitive" ++ spc() ++ pr_lident id ++ + (Option.cata (fun ty -> spc() ++ str":" ++ pr_spc_lconstr ty) (mt()) typopt) ++ spc() ++ + str ":=" ++ spc() ++ + str (CPrimitives.op_or_type_to_string r)) | VernacComments l -> return ( hov 2 diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 996fe320f9..7611355100 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -2032,21 +2032,31 @@ let vernac_locate = function let vernac_register qid r = let gr = Smartlocate.global_with_alias qid in - if Proof_global.there_are_pending_proofs () then + if Proof_global.there_are_pending_proofs () then user_err Pp.(str "Cannot register a primitive while in proof editing mode."); match r with | RegisterInline -> - if not (isConstRef gr) then - user_err Pp.(str "Register inline: a constant is expected"); - Global.register_inline (destConstRef gr) + begin match gr with + | ConstRef c -> Global.register_inline c + | _ -> CErrors.user_err (Pp.str "Register Inline: expecting a constant") + end | RegisterCoqlib n -> - let path, id = Libnames.repr_qualid n in - if DirPath.equal path Retroknowledge.int31_path - then - let f = Retroknowledge.(KInt31 (int31_field_of_string (Id.to_string id))) in - Global.register f gr - else - Coqlib.register_ref (Libnames.string_of_qualid n) gr + let ns, id = Libnames.repr_qualid n in + if DirPath.equal (dirpath_of_string "kernel") ns then begin + if Lib.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 + | k -> CErrors.user_err Pp.(str "Register: unknown identifier “" ++ str k ++ str "” in the “kernel” namespace") + in + match gr with + | IndRef ind -> Global.register_inductive ind pind + | _ -> CErrors.user_err (Pp.str "Register in kernel: expecting an inductive type") + end + else Coqlib.register_ref (Libnames.string_of_qualid n) gr (********************) (* Proof management *) @@ -2316,6 +2326,7 @@ let interp ?proof ~atts ~st c = | VernacLocate l -> unsupported_attributes atts; Feedback.msg_notice @@ vernac_locate l | VernacRegister (qid, r) -> unsupported_attributes atts; vernac_register qid r + | VernacPrimitive (id, prim, typopt) -> unsupported_attributes atts; ComAssumption.do_primitive id prim typopt | VernacComments l -> unsupported_attributes atts; Flags.if_verbose Feedback.msg_info (str "Comments ok\n") diff --git a/vernac/vernacexpr.ml b/vernac/vernacexpr.ml index 68a17e316e..2eb901890b 100644 --- a/vernac/vernacexpr.ml +++ b/vernac/vernacexpr.ml @@ -378,6 +378,7 @@ type nonrec vernac_expr = | VernacSearch of searchable * Goal_select.t option * search_restriction | VernacLocate of locatable | VernacRegister of qualid * register_kind + | VernacPrimitive of lident * CPrimitives.op_or_type * constr_expr option | VernacComments of comment list (* Proof management *) |
