diff options
-rw-r--r-- | Makefile.build | 5 | ||||
-rw-r--r-- | Makefile.common | 4 | ||||
-rw-r--r-- | dev/doc/changes.md | 6 | ||||
-rw-r--r-- | plugins/ssr/ssrparser.ml4 | 1 | ||||
-rw-r--r-- | plugins/ssrmatching/g_ssrmatching.ml4 | 101 | ||||
-rw-r--r-- | plugins/ssrmatching/g_ssrmatching.mli | 17 | ||||
-rw-r--r-- | plugins/ssrmatching/ssrmatching.ml (renamed from plugins/ssrmatching/ssrmatching.ml4) | 105 | ||||
-rw-r--r-- | plugins/ssrmatching/ssrmatching.mli | 34 | ||||
-rw-r--r-- | plugins/ssrmatching/ssrmatching_plugin.mlpack | 1 | ||||
-rw-r--r-- | plugins/syntax/n_syntax.ml | 81 | ||||
-rw-r--r-- | plugins/syntax/n_syntax_plugin.mlpack | 1 | ||||
-rw-r--r-- | plugins/syntax/positive_syntax.ml | 101 | ||||
-rw-r--r-- | plugins/syntax/positive_syntax_plugin.mlpack | 1 | ||||
-rw-r--r-- | plugins/syntax/z_syntax.ml | 128 | ||||
-rw-r--r-- | test-suite/bugs/closed/7712.v | 4 | ||||
-rw-r--r-- | test-suite/output/Notations3.out | 5 | ||||
-rw-r--r-- | test-suite/output/Notations3.v | 14 | ||||
-rw-r--r-- | theories/Numbers/BinNums.v | 2 | ||||
-rw-r--r-- | vernac/metasyntax.ml | 30 |
19 files changed, 406 insertions, 235 deletions
diff --git a/Makefile.build b/Makefile.build index b85418243..ed251af77 100644 --- a/Makefile.build +++ b/Makefile.build @@ -390,9 +390,10 @@ grammar/%.cmi: grammar/%.mli coqbinaries: $(TOPBIN) $(CHICKEN) $(CSDPCERT) $(FAKEIDE) coqbyte: $(TOPBYTE) $(CHICKENBYTE) -# Special rule for coqtop +# Special rule for coqtop, we imitate `ocamlopt` can delete the target +# to avoid #7666 $(COQTOPEXE): $(TOPBIN:.opt=.$(BEST)) - cp $< $@ + rm -f $@ && cp $< $@ bin/%.opt$(EXE): topbin/%_bin.ml $(LINKCMX) $(LIBCOQRUN) $(SHOW)'COQMKTOP -o $@' diff --git a/Makefile.common b/Makefile.common index 5b1def40a..94b965a7f 100644 --- a/Makefile.common +++ b/Makefile.common @@ -143,8 +143,8 @@ BTAUTOCMO:=plugins/btauto/btauto_plugin.cmo RTAUTOCMO:=plugins/rtauto/rtauto_plugin.cmo NATSYNTAXCMO:=plugins/syntax/nat_syntax_plugin.cmo OTHERSYNTAXCMO:=$(addprefix plugins/syntax/, \ - z_syntax_plugin.cmo \ - r_syntax_plugin.cmo \ + positive_syntax_plugin.cmo n_syntax_plugin.cmo \ + z_syntax_plugin.cmo r_syntax_plugin.cmo \ int31_syntax_plugin.cmo \ ascii_syntax_plugin.cmo \ string_syntax_plugin.cmo ) diff --git a/dev/doc/changes.md b/dev/doc/changes.md index cbb95625f..e6d741159 100644 --- a/dev/doc/changes.md +++ b/dev/doc/changes.md @@ -82,6 +82,12 @@ Vernacular commands `tactics/`. In all cases adapting is a matter of changing the module name. +Primitive number parsers + +- For better modularity, the primitive parsers for positive, N and Z + have been split over three files (plugins/syntax/positive_syntax.ml, + plugins/syntax/n_syntax.ml, plugins/syntax/z_syntax.ml). + ### Unit testing The test suite now allows writing unit tests against OCaml code in the Coq diff --git a/plugins/ssr/ssrparser.ml4 b/plugins/ssr/ssrparser.ml4 index 347a1e4e2..6b183dab1 100644 --- a/plugins/ssr/ssrparser.ml4 +++ b/plugins/ssr/ssrparser.ml4 @@ -962,6 +962,7 @@ END (* the default simpl and unfold tactics would erase blindly. *) open Ssrmatching_plugin.Ssrmatching +open Ssrmatching_plugin.G_ssrmatching let pr_wgen = function | (clr, Some((id,k),None)) -> spc() ++ pr_clear mt clr ++ str k ++ pr_hoi id diff --git a/plugins/ssrmatching/g_ssrmatching.ml4 b/plugins/ssrmatching/g_ssrmatching.ml4 new file mode 100644 index 000000000..746c368aa --- /dev/null +++ b/plugins/ssrmatching/g_ssrmatching.ml4 @@ -0,0 +1,101 @@ +(************************************************************************) +(* * 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 Ltac_plugin +open Genarg +open Pcoq +open Pcoq.Constr +open Ssrmatching +open Ssrmatching.Internal + +(* Defining grammar rules with "xx" in it automatically declares keywords too, + * we thus save the lexer to restore it at the end of the file *) +let frozen_lexer = CLexer.get_keyword_state () ;; + +DECLARE PLUGIN "ssrmatching_plugin" + +let pr_rpattern _ _ _ = pr_rpattern + +ARGUMENT EXTEND rpattern + TYPED AS rpatternty + PRINTED BY pr_rpattern + INTERPRETED BY interp_rpattern + GLOBALIZED BY glob_rpattern + SUBSTITUTED BY subst_rpattern + | [ lconstr(c) ] -> [ mk_rpattern (T (mk_lterm c None)) ] + | [ "in" lconstr(c) ] -> [ mk_rpattern (In_T (mk_lterm c None)) ] + | [ lconstr(x) "in" lconstr(c) ] -> + [ mk_rpattern (X_In_T (mk_lterm x None, mk_lterm c None)) ] + | [ "in" lconstr(x) "in" lconstr(c) ] -> + [ mk_rpattern (In_X_In_T (mk_lterm x None, mk_lterm c None)) ] + | [ lconstr(e) "in" lconstr(x) "in" lconstr(c) ] -> + [ mk_rpattern (E_In_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) ] + | [ lconstr(e) "as" lconstr(x) "in" lconstr(c) ] -> + [ mk_rpattern (E_As_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None)) ] +END + +let pr_ssrterm _ _ _ = pr_ssrterm + +ARGUMENT EXTEND cpattern + PRINTED BY pr_ssrterm + INTERPRETED BY interp_ssrterm + GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm + RAW_PRINTED BY pr_ssrterm + GLOB_PRINTED BY pr_ssrterm +| [ "Qed" constr(c) ] -> [ mk_lterm c None ] +END + +let input_ssrtermkind strm = match Util.stream_nth 0 strm with + | Tok.KEYWORD "(" -> '(' + | Tok.KEYWORD "@" -> '@' + | _ -> ' ' +let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind + +GEXTEND Gram + GLOBAL: cpattern; + cpattern: [[ k = ssrtermkind; c = constr -> + let pattern = mk_term k c None in + if loc_of_cpattern pattern <> Some !@loc && k = '(' + then mk_term 'x' c None + else pattern ]]; +END + +ARGUMENT EXTEND lcpattern + TYPED AS cpattern + PRINTED BY pr_ssrterm + INTERPRETED BY interp_ssrterm + GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm + RAW_PRINTED BY pr_ssrterm + GLOB_PRINTED BY pr_ssrterm +| [ "Qed" lconstr(c) ] -> [ mk_lterm c None ] +END + +GEXTEND Gram + GLOBAL: lcpattern; + lcpattern: [[ k = ssrtermkind; c = lconstr -> + let pattern = mk_term k c None in + if loc_of_cpattern pattern <> Some !@loc && k = '(' + then mk_term 'x' c None + else pattern ]]; +END + +ARGUMENT EXTEND ssrpatternarg TYPED AS rpattern PRINTED BY pr_rpattern +| [ rpattern(pat) ] -> [ pat ] +END + +TACTIC EXTEND ssrinstoftpat +| [ "ssrinstancesoftpat" cpattern(arg) ] -> [ Proofview.V82.tactic (ssrinstancesof arg) ] +END + +(* We wipe out all the keywords generated by the grammar rules we defined. *) +(* The user is supposed to Require Import ssreflect or Require ssreflect *) +(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) +(* consequence the extended ssreflect grammar. *) +let () = CLexer.set_keyword_state frozen_lexer ;; diff --git a/plugins/ssrmatching/g_ssrmatching.mli b/plugins/ssrmatching/g_ssrmatching.mli new file mode 100644 index 000000000..bb5161a0e --- /dev/null +++ b/plugins/ssrmatching/g_ssrmatching.mli @@ -0,0 +1,17 @@ +(* (c) Copyright 2006-2015 Microsoft Corporation and Inria. *) +(* Distributed under the terms of CeCILL-B. *) + +open Genarg +open Ssrmatching + +(** CS cpattern: (f _), (X in t), (t in X in t), (t as X in t) *) +val cpattern : cpattern Pcoq.Gram.entry +val wit_cpattern : cpattern uniform_genarg_type + +(** OS cpattern: f _, (X in t), (t in X in t), (t as X in t) *) +val lcpattern : cpattern Pcoq.Gram.entry +val wit_lcpattern : cpattern uniform_genarg_type + +(** OS rpattern: f _, in t, X in t, in X in t, t in X in t, t as X in t *) +val rpattern : rpattern Pcoq.Gram.entry +val wit_rpattern : rpattern uniform_genarg_type diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml index 9d9b1b2e8..05eda14e9 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml @@ -10,10 +10,6 @@ (* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *) -(* Defining grammar rules with "xx" in it automatically declares keywords too, - * we thus save the lexer to restore it at the end of the file *) -let frozen_lexer = CLexer.get_keyword_state () ;; - open Ltac_plugin open Names open Pp @@ -22,8 +18,6 @@ open Stdarg open Term module CoqConstr = Constr open CoqConstr -open Pcoq -open Pcoq.Constr open Vars open Libnames open Tactics @@ -46,8 +40,6 @@ open Evar_kinds open Constrexpr open Constrexpr_ops -DECLARE PLUGIN "ssrmatching_plugin" - let errorstrm = CErrors.user_err ~hdr:"ssrmatching" let loc_error loc msg = CErrors.user_err ?loc ~hdr:msg (str msg) let ppnl = Feedback.msg_info @@ -907,7 +899,6 @@ let pr_pattern_aux pr_constr = function let pp_pattern (sigma, p) = pr_pattern_aux (fun t -> pr_constr_pat (EConstr.Unsafe.to_constr (pi3 (nf_open_term sigma sigma (EConstr.of_constr t))))) p let pr_cpattern = pr_term -let pr_rpattern _ _ _ = pr_pattern let wit_rpatternty = add_genarg "rpatternty" pr_pattern @@ -987,27 +978,7 @@ let interp_rpattern s = function | E_As_X_In_T(e,x,t) -> E_As_X_In_T (interp_ssrterm s e,interp_ssrterm s x,interp_ssrterm s t) -let interp_rpattern ist gl t = Tacmach.project gl, interp_rpattern ist t - -ARGUMENT EXTEND rpattern - TYPED AS rpatternty - PRINTED BY pr_rpattern - INTERPRETED BY interp_rpattern - GLOBALIZED BY glob_rpattern - SUBSTITUTED BY subst_rpattern - | [ lconstr(c) ] -> [ T (mk_lterm c None) ] - | [ "in" lconstr(c) ] -> [ In_T (mk_lterm c None) ] - | [ lconstr(x) "in" lconstr(c) ] -> - [ X_In_T (mk_lterm x None, mk_lterm c None) ] - | [ "in" lconstr(x) "in" lconstr(c) ] -> - [ In_X_In_T (mk_lterm x None, mk_lterm c None) ] - | [ lconstr(e) "in" lconstr(x) "in" lconstr(c) ] -> - [ E_In_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None) ] - | [ lconstr(e) "as" lconstr(x) "in" lconstr(c) ] -> - [ E_As_X_In_T (mk_lterm e None, mk_lterm x None, mk_lterm c None) ] -END - - +let interp_rpattern0 ist gl t = Tacmach.project gl, interp_rpattern ist t type cpattern = char * glob_constr_and_expr * Geninterp.interp_sign option let tag_of_cpattern = pi1 @@ -1051,52 +1022,9 @@ let interp_wit wit ist gl x = let interp_open_constr ist gl gc = interp_wit wit_open_constr ist gl gc let pf_intern_term gl (_, c, ist) = glob_constr ist (pf_env gl) (project gl) c -let pr_ssrterm _ _ _ = pr_term -let input_ssrtermkind strm = match stream_nth 0 strm with - | Tok.KEYWORD "(" -> '(' - | Tok.KEYWORD "@" -> '@' - | _ -> ' ' -let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind let interp_ssrterm ist gl t = Tacmach.project gl, interp_ssrterm ist t -ARGUMENT EXTEND cpattern - PRINTED BY pr_ssrterm - INTERPRETED BY interp_ssrterm - GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm - RAW_PRINTED BY pr_ssrterm - GLOB_PRINTED BY pr_ssrterm -| [ "Qed" constr(c) ] -> [ mk_lterm c None ] -END - -GEXTEND Gram - GLOBAL: cpattern; - cpattern: [[ k = ssrtermkind; c = constr -> - let pattern = mk_term k c None in - if loc_ofCG pattern <> Some !@loc && k = '(' - then mk_term 'x' c None - else pattern ]]; -END - -ARGUMENT EXTEND lcpattern - TYPED AS cpattern - PRINTED BY pr_ssrterm - INTERPRETED BY interp_ssrterm - GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm - RAW_PRINTED BY pr_ssrterm - GLOB_PRINTED BY pr_ssrterm -| [ "Qed" lconstr(c) ] -> [ mk_lterm c None ] -END - -GEXTEND Gram - GLOBAL: lcpattern; - lcpattern: [[ k = ssrtermkind; c = lconstr -> - let pattern = mk_term k c None in - if loc_ofCG pattern <> Some !@loc && k = '(' - then mk_term 'x' c None - else pattern ]]; -END - let interp_term gl = function | (_, c, Some ist) -> on_snd EConstr.Unsafe.to_constr (interp_open_constr ist gl c) @@ -1416,10 +1344,6 @@ let is_wildcard ((_, (l, r), _) : cpattern) : bool = match DAst.get l, r with (* "ssrpattern" *) -ARGUMENT EXTEND ssrpatternarg TYPED AS rpattern PRINTED BY pr_rpattern -| [ rpattern(pat) ] -> [ pat ] -END - let pr_rpattern = pr_pattern let pf_merge_uc uc gl = @@ -1428,6 +1352,9 @@ let pf_merge_uc uc gl = let pf_unsafe_merge_uc uc gl = re_sig (sig_it gl) (Evd.set_universe_context (project gl) uc) +(** All the pattern types reuse the same dynamic toplevel tag *) +let wit_ssrpatternarg = wit_rpatternty + let interp_rpattern = interp_rpattern ~wit_ssrpatternarg let ssrpatterntac _ist arg gl = @@ -1479,14 +1406,20 @@ let ssrinstancesof arg gl = done; raise NoMatch with NoMatch -> ppnl (str"END INSTANCES"); tclIDTAC gl -TACTIC EXTEND ssrinstoftpat -| [ "ssrinstancesoftpat" cpattern(arg) ] -> [ Proofview.V82.tactic (ssrinstancesof arg) ] -END - -(* We wipe out all the keywords generated by the grammar rules we defined. *) -(* The user is supposed to Require Import ssreflect or Require ssreflect *) -(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) -(* consequence the extended ssreflect grammar. *) -let () = CLexer.set_keyword_state frozen_lexer ;; +module Internal = +struct + let wit_rpatternty = wit_rpatternty + let glob_rpattern = glob_rpattern + let subst_rpattern = subst_rpattern + let interp_rpattern = interp_rpattern0 + let pr_rpattern = pr_rpattern + let mk_rpattern x = x + let mk_lterm = mk_lterm + let mk_term = mk_term + let glob_cpattern = glob_cpattern + let subst_ssrterm = subst_ssrterm + let interp_ssrterm = interp_ssrterm + let pr_ssrterm = pr_term +end (* vim: set filetype=ocaml foldmethod=marker: *) diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index c55081e0f..f478d48ea 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -2,7 +2,6 @@ (* Distributed under the terms of CeCILL-B. *) open Goal -open Genarg open Environ open Evd open Constr @@ -19,24 +18,12 @@ open Tacexpr type cpattern val pr_cpattern : cpattern -> Pp.t -(** CS cpattern: (f _), (X in t), (t in X in t), (t as X in t) *) -val cpattern : cpattern Pcoq.Gram.entry -val wit_cpattern : cpattern uniform_genarg_type - -(** OS cpattern: f _, (X in t), (t in X in t), (t as X in t) *) -val lcpattern : cpattern Pcoq.Gram.entry -val wit_lcpattern : cpattern uniform_genarg_type - (** The type of rewrite patterns, the patterns of the [rewrite] tactic. These patterns also include patterns that identify all the subterms of a context (i.e. "in" prefix) *) type rpattern val pr_rpattern : rpattern -> Pp.t -(** OS rpattern: f _, in t, X in t, in X in t, t in X in t, t as X in t *) -val rpattern : rpattern Pcoq.Gram.entry -val wit_rpattern : rpattern uniform_genarg_type - (** Pattern interpretation and matching *) exception NoMatch @@ -238,4 +225,25 @@ val debug : bool -> unit * "Unset SsrMatchingProfiling" to get timings *) val profile : bool -> unit +val ssrinstancesof : cpattern -> Tacmach.tactic + +(** Functions used for grammar extensions. Do not use. *) + +module Internal : +sig + val wit_rpatternty : (rpattern, rpattern, rpattern) Genarg.genarg_type + val glob_rpattern : Genintern.glob_sign -> rpattern -> rpattern + val subst_rpattern : Mod_subst.substitution -> rpattern -> rpattern + val interp_rpattern : Geninterp.interp_sign -> Proof_type.goal Evd.sigma -> rpattern -> Evd.evar_map * rpattern + val pr_rpattern : rpattern -> Pp.t + val mk_rpattern : (cpattern, cpattern) ssrpattern -> rpattern + val mk_lterm : Constrexpr.constr_expr -> Geninterp.interp_sign option -> cpattern + val mk_term : char -> Constrexpr.constr_expr -> Geninterp.interp_sign option -> cpattern + + val glob_cpattern : Genintern.glob_sign -> cpattern -> cpattern + val subst_ssrterm : Mod_subst.substitution -> cpattern -> cpattern + val interp_ssrterm : Geninterp.interp_sign -> Proof_type.goal Evd.sigma -> cpattern -> Evd.evar_map * cpattern + val pr_ssrterm : cpattern -> Pp.t +end + (* eof *) diff --git a/plugins/ssrmatching/ssrmatching_plugin.mlpack b/plugins/ssrmatching/ssrmatching_plugin.mlpack index 5fb1f1567..02c75f14e 100644 --- a/plugins/ssrmatching/ssrmatching_plugin.mlpack +++ b/plugins/ssrmatching/ssrmatching_plugin.mlpack @@ -1 +1,2 @@ Ssrmatching +G_ssrmatching 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 *) +(* <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 Pp +open CErrors +open Names +open Bigint +open Positive_syntax_plugin.Positive_syntax + +(* Poor's man DECLARE PLUGIN *) +let __coq_plugin_name = "n_syntax_plugin" +let () = Mltop.add_known_module __coq_plugin_name + +(**********************************************************************) +(* Parsing N via scopes *) +(**********************************************************************) + +open Globnames +open Glob_term + +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) + +(* TODO: temporary hack *) +let make_kn dir id = Globnames.encode_mind dir id + +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) 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 *) +(* <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 Pp +open CErrors +open Util +open Names +open Bigint + +(* Poor's man DECLARE PLUGIN *) +let __coq_plugin_name = "positive_syntax_plugin" +let () = Mltop.add_known_module __coq_plugin_name + +exception Non_closed_number + +(**********************************************************************) +(* Parsing positive via scopes *) +(**********************************************************************) + +open Globnames +open Glob_term + +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) 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) diff --git a/test-suite/bugs/closed/7712.v b/test-suite/bugs/closed/7712.v new file mode 100644 index 000000000..a4e9697fa --- /dev/null +++ b/test-suite/bugs/closed/7712.v @@ -0,0 +1,4 @@ +(* This used to raise an anomaly *) + +Fail Reserved Notation "'[tele_arg' x ';' .. ';' z ]" + (format "[tele_arg '[hv' x .. z ']' ]"). diff --git a/test-suite/output/Notations3.out b/test-suite/output/Notations3.out index 996af5927..5ab616160 100644 --- a/test-suite/output/Notations3.out +++ b/test-suite/output/Notations3.out @@ -241,3 +241,8 @@ Notation Notation "( x , y , .. , z )" := pair .. (pair x y) .. z : core_scope (default interpretation) +1 subgoal + + ============================ + ##@% + ^^^ diff --git a/test-suite/output/Notations3.v b/test-suite/output/Notations3.v index 3cf0c913f..876aaa394 100644 --- a/test-suite/output/Notations3.v +++ b/test-suite/output/Notations3.v @@ -385,3 +385,17 @@ Module LocateNotations. Locate "exists". Locate "( _ , _ , .. , _ )". End LocateNotations. + +Module Issue7731. + +Axiom (P : nat -> Prop). +Parameter (X : nat). +Notation "## @ E ^^^" := (P E) (at level 20, E at level 1, format "'[ ' ## '/' @ E '/' ^^^ ']'"). +Notation "%" := X. + +Set Printing Width 7. +Goal ## @ % ^^^. +Show. +Abort. + +End Issue7731. diff --git a/theories/Numbers/BinNums.v b/theories/Numbers/BinNums.v index f8b3d9e1d..d5eb4f268 100644 --- a/theories/Numbers/BinNums.v +++ b/theories/Numbers/BinNums.v @@ -12,6 +12,8 @@ Set Implicit Arguments. +Declare ML Module "positive_syntax_plugin". +Declare ML Module "n_syntax_plugin". Declare ML Module "z_syntax_plugin". (** [positive] is a datatype representing the strictly positive integers diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml index da14358ef..240147c8d 100644 --- a/vernac/metasyntax.ml +++ b/vernac/metasyntax.ml @@ -487,6 +487,15 @@ and check_no_ldots_in_box = function let error_not_same ?loc () = user_err ?loc Pp.(str "The format is not the same on the right- and left-hand sides of the special token \"..\".") +let find_prod_list_loc sfmt fmt = + (* [fmt] is some [UnpTerminal x :: sfmt @ UnpTerminal ".." :: sfmt @ UnpTerminal y :: rest] *) + if List.is_empty sfmt then + (* No separators; we highlight the sequence "x .." *) + Loc.merge_opt (fst (List.hd fmt)) (fst (List.hd (List.tl fmt))) + else + (* A separator; we highlight the separating sequence *) + Loc.merge_opt (fst (List.hd sfmt)) (fst (List.last sfmt)) + let skip_var_in_recursive_format = function | (_,UnpTerminal s) :: sl (* skip first var *) when not (List.for_all (fun c -> c = " ") (String.explode s)) -> (* To do, though not so important: check that the names match @@ -496,6 +505,8 @@ let skip_var_in_recursive_format = function | [] -> assert false let read_recursive_format sl fmt = + (* Turn [[UnpTerminal s :: some-list @ UnpTerminal ".." :: same-some-list @ UnpTerminal s' :: rest] *) + (* into [(some-list,rest)] *) let get_head fmt = let sl = skip_var_in_recursive_format fmt in try split_format_at_ldots [] sl with Exit -> error_not_same ?loc:(fst (List.last (if sl = [] then fmt else sl))) () in @@ -528,10 +539,10 @@ let hunks_of_format (from,(vars,typs)) symfmt = let i = index_id m vars in let typ = List.nth typs (i-1) in let _,prec = precedence_of_entry_type from typ in - let slfmt,fmt = read_recursive_format sl fmt in - let sl, slfmt = aux (sl,slfmt) in - if not (List.is_empty sl) then error_format ?loc:(fst (List.last fmt)) (); - let symbs, l = aux (symbs,fmt) in + let loc_slfmt,rfmt = read_recursive_format sl fmt in + let sl, slfmt = aux (sl,loc_slfmt) in + if not (List.is_empty sl) then error_format ?loc:(find_prod_list_loc loc_slfmt fmt) (); + let symbs, l = aux (symbs,rfmt) in let hunk = match typ with | ETConstr _ -> UnpListMetaVar (i,prec,slfmt) | ETBinder isopen -> @@ -1312,8 +1323,15 @@ let make_pa_rule level (typs,symbols) ntn need_squash = let make_pp_rule level (typs,symbols) fmt = match fmt with - | None -> [UnpBox (PpHOVB 0, make_hunks typs symbols level)] - | Some fmt -> hunks_of_format (level, List.split typs) (symbols, parse_format fmt) + | None -> + let hunks = make_hunks typs symbols level in + if List.exists (function _,(UnpCut (PpBrk _) | UnpListMetaVar _) -> true | _ -> false) hunks then + [UnpBox (PpHOVB 0,hunks)] + else + (* Optimization to work around what seems an ocaml Format bug (see Mantis #7804/#7807) *) + List.map snd hunks (* drop locations which are dummy *) + | Some fmt -> + hunks_of_format (level, List.split typs) (symbols, parse_format fmt) (* let make_syntax_rules i_typs (ntn,prec,need_squash) sy_data fmt extra onlyprint compat = *) let make_syntax_rules (sd : SynData.syn_data) = let open SynData in |