aboutsummaryrefslogtreecommitdiff
path: root/interp
diff options
context:
space:
mode:
authorEmilio Jesus Gallego Arias2019-04-05 01:28:47 +0200
committerEmilio Jesus Gallego Arias2019-04-05 01:28:47 +0200
commitbe6f3a6234ee809dd3c290621d80c3280a41355e (patch)
tree8fed697f726193b765c8a2faeedd34ad60b541cb /interp
parent2e1aa5c15ad524cffd03c7979992af44ab2bb715 (diff)
parent6af420bb384af0acf94028fc44ef44fd5a6fd841 (diff)
Merge PR #8764: Add parsing of decimal constants (e.g., 1.02e+01)
Reviewed-by: Zimmi48 Reviewed-by: ejgallego Ack-by: gares Ack-by: herbelin Ack-by: ppedrot Ack-by: proux01
Diffstat (limited to 'interp')
-rw-r--r--interp/constrexpr.ml26
-rw-r--r--interp/constrexpr_ops.ml9
-rw-r--r--interp/constrextern.ml23
-rw-r--r--interp/constrintern.ml15
-rw-r--r--interp/interp.mllib1
-rw-r--r--interp/notation.ml117
-rw-r--r--interp/notation.mli17
-rw-r--r--interp/numTok.ml52
-rw-r--r--interp/numTok.mli18
9 files changed, 208 insertions, 70 deletions
diff --git a/interp/constrexpr.ml b/interp/constrexpr.ml
index 757d186c8b..7a14a4e708 100644
--- a/interp/constrexpr.ml
+++ b/interp/constrexpr.ml
@@ -48,16 +48,26 @@ type abstraction_kind = AbsLambda | AbsPi
type proj_flag = int option (** [Some n] = proj of the n-th visible argument *)
-(** Representation of integer literals that appear in Coq scripts.
- We now use raw strings of digits in base 10 (big-endian), and a separate
- sign flag. Note that this representation is not unique, due to possible
- multiple leading zeros, and -0 = +0 *)
-
-type sign = bool
-type raw_natural_number = string
+(** Representation of decimal literals that appear in Coq scripts.
+ We now use raw strings following the format defined by
+ [NumTok.t] and a separate sign flag.
+
+ Note that this representation is not unique, due to possible
+ multiple leading or trailing zeros, and -0 = +0, for instances.
+ The reason to keep the numeral exactly as it was parsed is that
+ specific notations can be declared for specific numerals
+ (e.g. [Notation "0" := False], or [Notation "00" := (nil,nil)], or
+ [Notation "2e1" := ...]). Those notations, which override the
+ generic interpretation as numeral, use the same representation of
+ numeral using the Numeral constructor. So the latter should be able
+ to record the form of the numeral which exactly matches the
+ notation. *)
+
+type sign = SPlus | SMinus
+type raw_numeral = NumTok.t
type prim_token =
- | Numeral of raw_natural_number * sign
+ | Numeral of sign * raw_numeral
| String of string
type instance_expr = Glob_term.glob_level list
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index 95a0039b0a..60610b92b8 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -50,13 +50,14 @@ let names_of_local_binders bl =
(**********************************************************************)
(* Functions on constr_expr *)
-(* Note: redundant Numeral representations such as -0 and +0 (or different
- numbers of leading zeros) are considered different here. *)
+(* Note: redundant Numeral representations, such as -0 and +0 (and others),
+ are considered different here. *)
let prim_token_eq t1 t2 = match t1, t2 with
-| Numeral (n1,s1), Numeral (n2,s2) -> String.equal n1 n2 && s1 == s2
+| Numeral (SPlus,n1), Numeral (SPlus,n2)
+| Numeral (SMinus,n1), Numeral (SMinus,n2) -> NumTok.equal n1 n2
| String s1, String s2 -> String.equal s1 s2
-| _ -> false
+| (Numeral ((SPlus|SMinus),_) | String _), _ -> false
let explicitation_eq ex1 ex2 = match ex1, ex2 with
| ExplByPos (i1, id1), ExplByPos (i2, id2) ->
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index c2afa097bb..24b1362e6d 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -318,29 +318,28 @@ let drop_implicits_in_patt cst nb_expl args =
let destPrim = function { CAst.v = CPrim t } -> Some t | _ -> None
let destPatPrim = function { CAst.v = CPatPrim t } -> Some t | _ -> None
-let is_number s =
- let rec aux i =
- Int.equal (String.length s) i ||
- match s.[i] with '0'..'9' -> aux (i+1) | _ -> false
- in aux 0
-
let is_zero s =
let rec aux i =
Int.equal (String.length s) i || (s.[i] == '0' && aux (i+1))
in aux 0
+let is_zero n = is_zero n.NumTok.int && is_zero n.NumTok.frac
let make_notation_gen loc ntn mknot mkprim destprim l bl =
match snd ntn,List.map destprim l with
(* Special case to avoid writing "- 3" for e.g. (Z.opp 3) *)
- | "- _", [Some (Numeral (p,true))] when not (is_zero p) ->
+ | "- _", [Some (Numeral (SPlus,p))] when not (is_zero p) ->
assert (bl=[]);
mknot (loc,ntn,([mknot (loc,(InConstrEntrySomeLevel,"( _ )"),l,[])]),[])
| _ ->
match decompose_notation_key ntn, l with
- | (InConstrEntrySomeLevel,[Terminal "-"; Terminal x]), [] when is_number x ->
- mkprim (loc, Numeral (x,false))
- | (InConstrEntrySomeLevel,[Terminal x]), [] when is_number x ->
- mkprim (loc, Numeral (x,true))
+ | (InConstrEntrySomeLevel,[Terminal "-"; Terminal x]), [] ->
+ begin match NumTok.of_string x with
+ | Some n -> mkprim (loc, Numeral (SMinus,n))
+ | None -> mknot (loc,ntn,l,bl) end
+ | (InConstrEntrySomeLevel,[Terminal x]), [] ->
+ begin match NumTok.of_string x with
+ | Some n -> mkprim (loc, Numeral (SPlus,n))
+ | None -> mknot (loc,ntn,l,bl) end
| _ -> mknot (loc,ntn,l,bl)
let make_notation loc ntn (terms,termlists,binders,binderlists as subst) =
@@ -969,7 +968,7 @@ let rec extern inctx (custom,scopes as allscopes) vars r =
CCast (sub_extern true scopes vars c,
map_cast_type (extern_typ scopes vars) c')
| GInt i ->
- CPrim(Numeral (Uint63.to_string i,true))
+ CPrim(Numeral (SPlus, NumTok.int (Uint63.to_string i)))
in insert_coercion coercion (CAst.make ?loc c)
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index 7a3e9881ea..59feb46dc1 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -1487,8 +1487,9 @@ let alias_of als = match als.alias_ids with
let is_zero s =
let rec aux i =
- Int.equal (String.length s) i || (s.[i] == '0' && aux (i+1))
+ Int.equal (String.length s) i || ((s.[i] == '0' || s.[i] == '_') && aux (i+1))
in aux 0
+let is_zero n = is_zero n.NumTok.int && is_zero n.NumTok.frac
let merge_subst s1 s2 = Id.Map.fold Id.Map.add s1 s2
@@ -1513,11 +1514,11 @@ let rec subst_pat_iterator y t = DAst.(map (function
| RCPatOr pl -> RCPatOr (List.map (subst_pat_iterator y t) pl)))
let is_non_zero c = match c with
-| { CAst.v = CPrim (Numeral (p, true)) } -> not (is_zero p)
+| { CAst.v = CPrim (Numeral (SPlus, p)) } -> not (is_zero p)
| _ -> false
let is_non_zero_pat c = match c with
-| { CAst.v = CPatPrim (Numeral (p, true)) } -> not (is_zero p)
+| { CAst.v = CPatPrim (Numeral (SPlus, p)) } -> not (is_zero p)
| _ -> false
let get_asymmetric_patterns = Goptions.declare_bool_option_and_ref
@@ -1628,8 +1629,8 @@ let drop_notations_pattern looked_for genv =
let (argscs1,_) = find_remaining_scopes expl_pl pl g in
DAst.make ?loc @@ RCPatCstr (g, List.map2 (in_pat_sc scopes) argscs1 expl_pl @ List.map (in_pat false scopes) pl, [])
| CPatNotation ((InConstrEntrySomeLevel,"- _"),([a],[]),[]) when is_non_zero_pat a ->
- let p = match a.CAst.v with CPatPrim (Numeral (p, _)) -> p | _ -> assert false in
- let pat, _df = Notation.interp_prim_token_cases_pattern_expr ?loc (ensure_kind false loc) (Numeral (p,false)) scopes in
+ let p = match a.CAst.v with CPatPrim (Numeral (_, p)) -> p | _ -> assert false in
+ let pat, _df = Notation.interp_prim_token_cases_pattern_expr ?loc (ensure_kind false loc) (Numeral (SMinus,p)) scopes in
rcp_of_glob scopes pat
| CPatNotation ((InConstrEntrySomeLevel,"( _ )"),([a],[]),[]) ->
in_pat top scopes a
@@ -1944,8 +1945,8 @@ let internalize globalenv env pattern_mode (_, ntnvars as lvar) c =
GLetIn (na.CAst.v, inc1, int,
intern (push_name_env ntnvars (impls_term_list inc1) env na) c2)
| CNotation ((InConstrEntrySomeLevel,"- _"), ([a],[],[],[])) when is_non_zero a ->
- let p = match a.CAst.v with CPrim (Numeral (p, _)) -> p | _ -> assert false in
- intern env (CAst.make ?loc @@ CPrim (Numeral (p,false)))
+ let p = match a.CAst.v with CPrim (Numeral (_, p)) -> p | _ -> assert false in
+ intern env (CAst.make ?loc @@ CPrim (Numeral (SMinus,p)))
| CNotation ((InConstrEntrySomeLevel,"( _ )"),([a],[],[],[])) -> intern env a
| CNotation (ntn,args) ->
intern_notation intern env ntnvars loc ntn args
diff --git a/interp/interp.mllib b/interp/interp.mllib
index 147eaf20dc..1262dbb181 100644
--- a/interp/interp.mllib
+++ b/interp/interp.mllib
@@ -1,3 +1,4 @@
+NumTok
Constrexpr
Tactypes
Stdarg
diff --git a/interp/notation.ml b/interp/notation.ml
index 2765661749..b9aca82cf4 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -476,7 +476,7 @@ let notation_constr_key = function (* Rem: NApp(NRef ref,[]) stands for @ref *)
(* Interpreting numbers (not in summary because functional objects) *)
type required_module = full_path * string list
-type rawnum = Constrexpr.raw_natural_number * Constrexpr.sign
+type rawnum = Constrexpr.sign * Constrexpr.raw_numeral
type prim_token_uid = string
@@ -499,15 +499,20 @@ module InnerPrimToken = struct
| StringInterp f, StringInterp f' -> f == f'
| _ -> false
- let ofNumeral n s =
- if s then Bigint.of_string n else Bigint.neg (Bigint.of_string n)
+ let ofNumeral s n =
+ let n = String.(concat "" (split_on_char '_' n)) in
+ match s with
+ | SPlus -> Bigint.of_string n
+ | SMinus -> Bigint.neg (Bigint.of_string n)
let do_interp ?loc interp primtok =
match primtok, interp with
- | Numeral (n,s), RawNumInterp interp -> interp ?loc (n,s)
- | Numeral (n,s), BigNumInterp interp -> interp ?loc (ofNumeral n s)
+ | Numeral (s,n), RawNumInterp interp -> interp ?loc (s,n)
+ | Numeral (s,{ NumTok.int = n; frac = ""; exp = "" }),
+ BigNumInterp interp -> interp ?loc (ofNumeral s n)
| String s, StringInterp interp -> interp ?loc s
- | _ -> raise Not_found
+ | (Numeral _ | String _),
+ (RawNumInterp _ | BigNumInterp _ | StringInterp _) -> raise Not_found
type uninterpreter =
| RawNumUninterp of (any_glob_constr -> rawnum option)
@@ -521,15 +526,17 @@ module InnerPrimToken = struct
| _ -> false
let mkNumeral n =
- if Bigint.is_pos_or_zero n then Numeral (Bigint.to_string n, true)
- else Numeral (Bigint.to_string (Bigint.neg n), false)
+ if Bigint.is_pos_or_zero n then
+ Numeral (SPlus,NumTok.int (Bigint.to_string n))
+ else
+ Numeral (SMinus,NumTok.int (Bigint.to_string (Bigint.neg n)))
let mkString = function
| None -> None
| Some s -> if Unicode.is_utf8 s then Some (String s) else None
let do_uninterp uninterp g = match uninterp with
- | RawNumUninterp u -> Option.map (fun (n,s) -> Numeral (n,s)) (u g)
+ | RawNumUninterp u -> Option.map (fun (s,n) -> Numeral (s,n)) (u g)
| BigNumUninterp u -> Option.map mkNumeral (u g)
| StringUninterp u -> mkString (u g)
@@ -559,8 +566,8 @@ exception PrimTokenNotationError of string * Environ.env * Evd.evar_map * prim_t
type numnot_option =
| Nop
- | Warning of raw_natural_number
- | Abstract of raw_natural_number
+ | Warning of string
+ | Abstract of string
type int_ty =
{ uint : Names.inductive;
@@ -570,11 +577,16 @@ type z_pos_ty =
{ z_ty : Names.inductive;
pos_ty : Names.inductive }
+type decimal_ty =
+ { int : int_ty;
+ decimal : Names.inductive }
+
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 *)
+ | Decimal of decimal_ty (* Coq.Init.Decimal.decimal + uint + int *)
type string_target_kind =
| ListByte
@@ -760,15 +772,29 @@ let coquint_of_rawnum uint str =
let nil = mkConstruct (uint,1) in
let rec do_chars s i acc =
if i < 0 then acc
- else
+ else if s.[i] == '_' then do_chars s (i-1) acc else
let dg = mkConstruct (uint, digit_of_char s.[i]) in
do_chars s (i-1) (mkApp(dg,[|acc|]))
in
do_chars str (String.length str - 1) nil
-let coqint_of_rawnum inds (str,sign) =
+let coqint_of_rawnum inds sign str =
let uint = coquint_of_rawnum inds.uint str in
- mkApp (mkConstruct (inds.int, if sign then 1 else 2), [|uint|])
+ let pos_neg = match sign with SPlus -> 1 | SMinus -> 2 in
+ mkApp (mkConstruct (inds.int, pos_neg), [|uint|])
+
+let coqdecimal_of_rawnum inds sign n =
+ let i, f, e = NumTok.(n.int, n.frac, n.exp) in
+ let i = coqint_of_rawnum inds.int sign i in
+ let f = coquint_of_rawnum inds.int.uint f in
+ if e = "" then mkApp (mkConstruct (inds.decimal, 1), [|i; f|]) (* Decimal *)
+ else
+ let sign, e = match e.[1] with
+ | '-' -> SMinus, String.sub e 2 (String.length e - 2)
+ | '+' -> SPlus, String.sub e 2 (String.length e - 2)
+ | _ -> SPlus, String.sub e 1 (String.length e - 1) in
+ let e = coqint_of_rawnum inds.int sign e in
+ mkApp (mkConstruct (inds.decimal, 2), [|i; f; e|]) (* DecimalExp *)
let rawnum_of_coquint c =
let rec of_uint_loop c buf =
@@ -788,17 +814,30 @@ let rawnum_of_coquint c =
(* To avoid ambiguities between Nil and (D0 Nil), we choose
to not display Nil alone as "0" *)
raise NotAValidPrimToken
- else Buffer.contents buf
+ else NumTok.int (Buffer.contents buf)
let rawnum_of_coqint c =
match Constr.kind c with
| App (c,[|c'|]) ->
(match Constr.kind c with
- | Construct ((_,1), _) (* Pos *) -> (rawnum_of_coquint c', true)
- | Construct ((_,2), _) (* Neg *) -> (rawnum_of_coquint c', false)
+ | Construct ((_,1), _) (* Pos *) -> (SPlus, rawnum_of_coquint c')
+ | Construct ((_,2), _) (* Neg *) -> (SMinus, rawnum_of_coquint c')
| _ -> raise NotAValidPrimToken)
| _ -> raise NotAValidPrimToken
+let rawnum_of_decimal c =
+ let of_ife i f e =
+ let sign, n = rawnum_of_coqint i in
+ let f =
+ try (rawnum_of_coquint f).NumTok.int with NotAValidPrimToken -> "" in
+ let e = match e with None -> "" | Some e -> match rawnum_of_coqint e with
+ | SPlus, e -> "e+" ^ e.NumTok.int
+ | SMinus, e -> "e-" ^ e.NumTok.int in
+ sign,{ n with NumTok.frac = f; exp = e } in
+ match Constr.kind c with
+ | App (_,[|i; f|]) -> of_ife i f None
+ | App (_,[|i; f; e|]) -> of_ife i f (Some e)
+ | _ -> raise NotAValidPrimToken
(***********************************************************************)
@@ -885,31 +924,42 @@ let bigint_of_int63 c =
| _ -> 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)
+ if Bigint.is_pos_or_zero n then
+ (SPlus, NumTok.int (Bigint.to_string n))
+ else
+ (SMinus, NumTok.int (Bigint.to_string (Bigint.neg n)))
-let raw2big (n,s) =
- if s then Bigint.of_string n else Bigint.neg (Bigint.of_string n)
+let raw2big s n = match s with
+ | SPlus -> Bigint.of_string n
+ | SMinus -> Bigint.neg (Bigint.of_string n)
let interp o ?loc n =
- begin match o.warning with
- | Warning threshold when snd n && rawnum_compare (fst n) threshold >= 0 ->
+ begin match o.warning, n with
+ | Warning threshold, (SPlus, { NumTok.int = n; frac = ""; exp = "" })
+ when rawnum_compare n threshold >= 0 ->
warn_large_num o.ty_name
| _ -> ()
end;
- let c = match fst o.to_kind with
- | Int int_ty -> coqint_of_rawnum int_ty 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)
+ let c = match fst o.to_kind, n with
+ | Int int_ty, (s, { NumTok.int = n; frac = ""; exp = "" }) ->
+ coqint_of_rawnum int_ty s n
+ | UInt uint_ty, (SPlus, { NumTok.int = n; frac = ""; exp = "" }) ->
+ coquint_of_rawnum uint_ty n
+ | Z z_pos_ty, (s, { NumTok.int = n; frac = ""; exp = "" }) ->
+ z_of_bigint z_pos_ty (raw2big s n)
+ | Int63, (s, { NumTok.int = n; frac = ""; exp = "" }) ->
+ interp_int63 ?loc (raw2big s n)
+ | (Int _ | UInt _ | Z _ | Int63), _ ->
+ no_such_prim_token "number" ?loc o.ty_name
+ | Decimal decimal_ty, (s,n) -> coqdecimal_of_rawnum decimal_ty s n
in
let env = Global.env () in
let sigma = Evd.from_env env in
let sigma,to_ty = Evd.fresh_global env sigma o.to_ty in
let to_ty = EConstr.Unsafe.to_constr to_ty in
match o.warning, snd o.to_kind with
- | Abstract threshold, Direct when rawnum_compare (fst n) threshold >= 0 ->
+ | Abstract threshold, Direct
+ when rawnum_compare (snd n).NumTok.int threshold >= 0 ->
warn_abstract_large_num (o.ty_name,o.to_ty);
glob_of_constr "numeral" ?loc env sigma (mkApp (to_ty,[|c|]))
| _ ->
@@ -922,9 +972,10 @@ let uninterp o n =
PrimTokenNotation.uninterp
begin function
| (Int _, c) -> rawnum_of_coqint c
- | (UInt _, c) -> (rawnum_of_coquint c, true)
+ | (UInt _, c) -> (SPlus, rawnum_of_coquint c)
| (Z _, c) -> big2raw (bigint_of_z c)
| (Int63, c) -> big2raw (bigint_of_int63 c)
+ | (Decimal _, c) -> rawnum_of_decimal c
end o n
end
@@ -1249,8 +1300,8 @@ let find_notation ntn sc =
(n.not_interp, n.not_location)
let notation_of_prim_token = function
- | Numeral (n,true) -> InConstrEntrySomeLevel, n
- | Numeral (n,false) -> InConstrEntrySomeLevel, "- "^n
+ | Numeral (SPlus,n) -> InConstrEntrySomeLevel, NumTok.to_string n
+ | Numeral (SMinus,n) -> InConstrEntrySomeLevel, "- "^NumTok.to_string n
| String _ -> raise Not_found
let find_prim_token check_allowed ?loc p sc =
diff --git a/interp/notation.mli b/interp/notation.mli
index 5dcc96dc29..57e2be16b9 100644
--- a/interp/notation.mli
+++ b/interp/notation.mli
@@ -70,14 +70,14 @@ val find_delimiters_scope : ?loc:Loc.t -> delimiters -> scope_name
(** {6 Declare and uses back and forth an interpretation of primitive token } *)
-(** A numeral interpreter is the pair of an interpreter for **integer**
+(** A numeral interpreter is the pair of an interpreter for **decimal**
numbers in terms and an optional interpreter in pattern, if
- negative numbers are not supported, the interpreter must fail with
- an appropriate error message *)
+ non integer or negative numbers are not supported, the interpreter
+ must fail with an appropriate error message *)
type notation_location = (DirPath.t * DirPath.t) * string
type required_module = full_path * string list
-type rawnum = Constrexpr.raw_natural_number * Constrexpr.sign
+type rawnum = Constrexpr.sign * Constrexpr.raw_numeral
(** The unique id string below will be used to refer to a particular
registered interpreter/uninterpreter of numeral or string notation.
@@ -112,8 +112,8 @@ exception PrimTokenNotationError of string * Environ.env * Evd.evar_map * prim_t
type numnot_option =
| Nop
- | Warning of raw_natural_number
- | Abstract of raw_natural_number
+ | Warning of string
+ | Abstract of string
type int_ty =
{ uint : Names.inductive;
@@ -123,11 +123,16 @@ type z_pos_ty =
{ z_ty : Names.inductive;
pos_ty : Names.inductive }
+type decimal_ty =
+ { int : int_ty;
+ decimal : Names.inductive }
+
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 *)
+ | Decimal of decimal_ty (* Coq.Init.Decimal.decimal + uint + int *)
type string_target_kind =
| ListByte
diff --git a/interp/numTok.ml b/interp/numTok.ml
new file mode 100644
index 0000000000..8f2004b889
--- /dev/null
+++ b/interp/numTok.ml
@@ -0,0 +1,52 @@
+type t = {
+ int : string;
+ frac : string;
+ exp : string
+}
+
+let equal n1 n2 =
+ String.(equal n1.int n2.int && equal n1.frac n2.frac && equal n1.exp n2.exp)
+
+let int s = { int = s; frac = ""; exp = "" }
+
+let to_string n = n.int ^ (if n.frac = "" then "" else "." ^ n.frac) ^ n.exp
+
+let parse =
+ let buff = ref (Bytes.create 80) in
+ let store len x =
+ let open Bytes in
+ if len >= length !buff then
+ buff := cat !buff (create (length !buff));
+ set !buff len x;
+ succ len in
+ let get_buff len = Bytes.sub_string !buff 0 len in
+ (* reads [0-9_]* *)
+ let rec number len s = match Stream.peek s with
+ | Some (('0'..'9' | '_') as c) -> Stream.junk s; number (store len c) s
+ | _ -> len in
+ fun s ->
+ let i = get_buff (number 0 s) in
+ let f =
+ match Stream.npeek 2 s with
+ | '.' :: (('0'..'9' | '_') as c) :: _ ->
+ Stream.junk s; Stream.junk s; get_buff (number (store 0 c) s)
+ | _ -> "" in
+ let e =
+ match (Stream.npeek 2 s) with
+ | (('e'|'E') as e) :: ('0'..'9' as c) :: _ ->
+ Stream.junk s; Stream.junk s; get_buff (number (store (store 0 e) c) s)
+ | (('e'|'E') as e) :: (('+'|'-') as sign) :: _ ->
+ begin match Stream.npeek 3 s with
+ | _ :: _ :: ('0'..'9' as c) :: _ ->
+ Stream.junk s; Stream.junk s; Stream.junk s;
+ get_buff (number (store (store (store 0 e) sign) c) s)
+ | _ -> ""
+ end
+ | _ -> "" in
+ { int = i; frac = f; exp = e }
+
+let of_string s =
+ if s = "" || s.[0] < '0' || s.[0] > '9' then None else
+ let strm = Stream.of_string (s ^ " ") in
+ let n = parse strm in
+ if Stream.count strm >= String.length s then Some n else None
diff --git a/interp/numTok.mli b/interp/numTok.mli
new file mode 100644
index 0000000000..0b6a877cbd
--- /dev/null
+++ b/interp/numTok.mli
@@ -0,0 +1,18 @@
+type t = {
+ int : string; (** \[0-9\]\[0-9_\]* *)
+ frac : string; (** empty or \[0-9_\]+ *)
+ exp : string (** empty or \[eE\]\[+-\]?\[0-9\]\[0-9_\]* *)
+}
+
+val equal : t -> t -> bool
+
+(** [int s] amounts to [\{ int = s; frac = ""; exp = "" \}] *)
+val int : string -> t
+
+val to_string : t -> string
+
+val of_string : string -> t option
+
+(** Precondition: the first char on the stream is a digit (\[0-9\]).
+ Precondition: at least two extra chars after the numeral to parse. *)
+val parse : char Stream.t -> t