From 1b3c25792c0598ae0cc44e9db1e7bde3f5789638 Mon Sep 17 00:00:00 2001 From: Hugo Herbelin Date: Sat, 31 Mar 2018 15:03:27 +0200 Subject: Splitting primitive numeral parser/printer for positive, N, Z into three files. --- plugins/syntax/n_syntax.ml | 81 +++++++++++++++++ plugins/syntax/n_syntax_plugin.mlpack | 1 + plugins/syntax/positive_syntax.ml | 101 +++++++++++++++++++++ plugins/syntax/positive_syntax_plugin.mlpack | 1 + plugins/syntax/z_syntax.ml | 128 +-------------------------- 5 files changed, 186 insertions(+), 126 deletions(-) create mode 100644 plugins/syntax/n_syntax.ml create mode 100644 plugins/syntax/n_syntax_plugin.mlpack create mode 100644 plugins/syntax/positive_syntax.ml create mode 100644 plugins/syntax/positive_syntax_plugin.mlpack (limited to 'plugins') diff --git a/plugins/syntax/n_syntax.ml b/plugins/syntax/n_syntax.ml new file mode 100644 index 000000000..0e202be47 --- /dev/null +++ b/plugins/syntax/n_syntax.ml @@ -0,0 +1,81 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* bignat_of_pos a + | GRef (a,_) when GlobRef.equal a glob_N0 -> Bigint.zero + | _ -> raise Non_closed_number + ) n + +let uninterp_n (AnyGlobConstr p) = + try Some (bignat_of_n p) + with Non_closed_number -> None + +(************************************************************************) +(* Declaring interpreters and uninterpreters for N *) + +let _ = Notation.declare_numeral_interpreter "N_scope" + (n_path,binnums) + n_of_int + ([DAst.make @@ GRef (glob_N0, None); + DAst.make @@ GRef (glob_Npos, None)], + uninterp_n, + true) diff --git a/plugins/syntax/n_syntax_plugin.mlpack b/plugins/syntax/n_syntax_plugin.mlpack new file mode 100644 index 000000000..4c56645f0 --- /dev/null +++ b/plugins/syntax/n_syntax_plugin.mlpack @@ -0,0 +1 @@ +N_syntax diff --git a/plugins/syntax/positive_syntax.ml b/plugins/syntax/positive_syntax.ml new file mode 100644 index 000000000..0c82e4744 --- /dev/null +++ b/plugins/syntax/positive_syntax.ml @@ -0,0 +1,101 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* DAst.make ?loc @@ GApp (ref_xO,[pos_of q]) + | (q,true) when not (Bigint.equal q zero) -> DAst.make ?loc @@ GApp (ref_xI,[pos_of q]) + | (q,true) -> ref_xH + in + pos_of x + +let error_non_positive ?loc = + user_err ?loc ~hdr:"interp_positive" + (str "Only strictly positive numbers in type \"positive\".") + +let interp_positive ?loc n = + if is_strictly_pos n then pos_of_bignat ?loc n + else error_non_positive ?loc + +(**********************************************************************) +(* Printing positive via scopes *) +(**********************************************************************) + +let is_gr c gr = match DAst.get c with +| GRef (r, _) -> GlobRef.equal r gr +| _ -> false + +let rec bignat_of_pos x = DAst.with_val (function + | GApp (r ,[a]) when is_gr r glob_xO -> mult_2(bignat_of_pos a) + | GApp (r ,[a]) when is_gr r glob_xI -> add_1(mult_2(bignat_of_pos a)) + | GRef (a, _) when GlobRef.equal a glob_xH -> Bigint.one + | _ -> raise Non_closed_number + ) x + +let uninterp_positive (AnyGlobConstr p) = + try + Some (bignat_of_pos p) + with Non_closed_number -> + None + +(************************************************************************) +(* Declaring interpreters and uninterpreters for positive *) +(************************************************************************) + +let _ = Notation.declare_numeral_interpreter "positive_scope" + (positive_path,binnums) + interp_positive + ([DAst.make @@ GRef (glob_xI, None); + DAst.make @@ GRef (glob_xO, None); + DAst.make @@ GRef (glob_xH, None)], + uninterp_positive, + true) diff --git a/plugins/syntax/positive_syntax_plugin.mlpack b/plugins/syntax/positive_syntax_plugin.mlpack new file mode 100644 index 000000000..ac8f3c425 --- /dev/null +++ b/plugins/syntax/positive_syntax_plugin.mlpack @@ -0,0 +1 @@ +Positive_syntax diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml index 09fe8bf70..2534162e3 100644 --- a/plugins/syntax/z_syntax.ml +++ b/plugins/syntax/z_syntax.ml @@ -8,20 +8,16 @@ (* * (see LICENSE file for the text of the license) *) (************************************************************************) -open Pp -open CErrors -open Util open Names open Bigint +open Positive_syntax_plugin.Positive_syntax (* Poor's man DECLARE PLUGIN *) let __coq_plugin_name = "z_syntax_plugin" let () = Mltop.add_known_module __coq_plugin_name -exception Non_closed_number - (**********************************************************************) -(* Parsing positive via scopes *) +(* Parsing Z via scopes *) (**********************************************************************) open Globnames @@ -32,129 +28,9 @@ let binnums = ["Coq";"Numbers";"BinNums"] 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 positive_path = make_path binnums "positive" - (* TODO: temporary hack *) let make_kn dir id = Globnames.encode_mind dir id -let positive_kn = make_kn (make_dir binnums) (Id.of_string "positive") -let glob_positive = IndRef (positive_kn,0) -let path_of_xI = ((positive_kn,0),1) -let path_of_xO = ((positive_kn,0),2) -let path_of_xH = ((positive_kn,0),3) -let glob_xI = ConstructRef path_of_xI -let glob_xO = ConstructRef path_of_xO -let glob_xH = ConstructRef path_of_xH - -let pos_of_bignat ?loc x = - let ref_xI = DAst.make ?loc @@ GRef (glob_xI, None) in - let ref_xH = DAst.make ?loc @@ GRef (glob_xH, None) in - let ref_xO = DAst.make ?loc @@ GRef (glob_xO, None) in - let rec pos_of x = - match div2_with_rest x with - | (q,false) -> DAst.make ?loc @@ GApp (ref_xO,[pos_of q]) - | (q,true) when not (Bigint.equal q zero) -> DAst.make ?loc @@ GApp (ref_xI,[pos_of q]) - | (q,true) -> ref_xH - in - pos_of x - -let error_non_positive ?loc = - user_err ?loc ~hdr:"interp_positive" - (str "Only strictly positive numbers in type \"positive\".") - -let interp_positive ?loc n = - if is_strictly_pos n then pos_of_bignat ?loc n - else error_non_positive ?loc - -(**********************************************************************) -(* Printing positive via scopes *) -(**********************************************************************) - -let is_gr c gr = match DAst.get c with -| GRef (r, _) -> GlobRef.equal r gr -| _ -> false - -let rec bignat_of_pos x = DAst.with_val (function - | GApp (r ,[a]) when is_gr r glob_xO -> mult_2(bignat_of_pos a) - | GApp (r ,[a]) when is_gr r glob_xI -> add_1(mult_2(bignat_of_pos a)) - | GRef (a, _) when GlobRef.equal a glob_xH -> Bigint.one - | _ -> raise Non_closed_number - ) x - -let uninterp_positive (AnyGlobConstr p) = - try - Some (bignat_of_pos p) - with Non_closed_number -> - None - -(************************************************************************) -(* Declaring interpreters and uninterpreters for positive *) -(************************************************************************) - -let _ = Notation.declare_numeral_interpreter "positive_scope" - (positive_path,binnums) - interp_positive - ([DAst.make @@ GRef (glob_xI, None); - DAst.make @@ GRef (glob_xO, None); - DAst.make @@ GRef (glob_xH, None)], - uninterp_positive, - true) - -(**********************************************************************) -(* Parsing N via scopes *) -(**********************************************************************) - -let n_kn = make_kn (make_dir binnums) (Id.of_string "N") -let glob_n = IndRef (n_kn,0) -let path_of_N0 = ((n_kn,0),1) -let path_of_Npos = ((n_kn,0),2) -let glob_N0 = ConstructRef path_of_N0 -let glob_Npos = ConstructRef path_of_Npos - -let n_path = make_path binnums "N" - -let n_of_binnat ?loc pos_or_neg n = DAst.make ?loc @@ - if not (Bigint.equal n zero) then - GApp(DAst.make @@ GRef (glob_Npos,None), [pos_of_bignat ?loc n]) - else - GRef(glob_N0, None) - -let error_negative ?loc = - user_err ?loc ~hdr:"interp_N" (str "No negative numbers in type \"N\".") - -let n_of_int ?loc n = - if is_pos_or_zero n then n_of_binnat ?loc true n - else error_negative ?loc - -(**********************************************************************) -(* Printing N via scopes *) -(**********************************************************************) - -let bignat_of_n n = DAst.with_val (function - | GApp (r, [a]) when is_gr r glob_Npos -> bignat_of_pos a - | GRef (a,_) when GlobRef.equal a glob_N0 -> Bigint.zero - | _ -> raise Non_closed_number - ) n - -let uninterp_n (AnyGlobConstr p) = - try Some (bignat_of_n p) - with Non_closed_number -> None - -(************************************************************************) -(* Declaring interpreters and uninterpreters for N *) - -let _ = Notation.declare_numeral_interpreter "N_scope" - (n_path,binnums) - n_of_int - ([DAst.make @@ GRef (glob_N0, None); - DAst.make @@ GRef (glob_Npos, None)], - uninterp_n, - true) - -(**********************************************************************) -(* Parsing Z via scopes *) -(**********************************************************************) - let z_path = make_path binnums "Z" let z_kn = make_kn (make_dir binnums) (Id.of_string "Z") let glob_z = IndRef (z_kn,0) -- cgit v1.2.3