Skip to content

Commit

Permalink
Merge PR #12341: [search] [ssr] Emit deprecated message when calling …
Browse files Browse the repository at this point in the history
…search from ssreflect

Reviewed-by: Zimmi48
Ack-by: gares
  • Loading branch information
Zimmi48 committed May 18, 2020
2 parents a891dd5 + 57c023f commit b7c14a8
Show file tree
Hide file tree
Showing 4 changed files with 39 additions and 1 deletion.
29 changes: 29 additions & 0 deletions plugins/ssr/ssrvernac.mlg
Original file line number Diff line number Diff line change
Expand Up @@ -300,6 +300,35 @@ VERNAC COMMAND EXTEND HintView CLASSIFIED AS SIDEFF
Ssrview.AdaptorDb.declare k hints }
END

(** Search compatibility *)

{

let warn_search_moved_enabled = ref true
let warn_search_moved = CWarnings.create ~name:"ssr-search-moved"
~category:"deprecated" ~default:CWarnings.Enabled
(fun () ->
(Pp.strbrk
"SSReflect's Search command has been moved to the \
ssrsearch module; please Require that module if you \
still want to use SSReflect's Search command"))

open G_vernac
}

GRAMMAR EXTEND Gram
GLOBAL: query_command;

query_command:
[ [ IDENT "Search"; s = search_query; l = search_queries; "." ->
{ let (sl,m) = l in
if !warn_search_moved_enabled then warn_search_moved ();
fun g ->
Vernacexpr.VernacSearch (Vernacexpr.Search (s::sl),g, m) }
] ]
;
END

(** Keyword compatibility fixes. *)

(* Coq v8.1 notation uses "by" and "of" quasi-keywords, i.e., reserved *)
Expand Down
2 changes: 2 additions & 0 deletions plugins/ssr/ssrvernac.mli
Original file line number Diff line number Diff line change
Expand Up @@ -9,3 +9,5 @@
(************************************************************************)

(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)

val warn_search_moved_enabled : bool ref
4 changes: 4 additions & 0 deletions plugins/ssrsearch/g_search.mlg
Original file line number Diff line number Diff line change
Expand Up @@ -299,6 +299,10 @@ let ssrdisplaysearch gr env t =
let pr_res = pr_global gr ++ str ":" ++ spc () ++ pr_lconstr_env env Evd.empty t in
Feedback.msg_notice (hov 2 pr_res ++ fnl ())

(* Remove the warning entirely when this plugin is loaded. *)
let _ =
Ssreflect_plugin.Ssrvernac.warn_search_moved_enabled := false

let deprecated_search =
CWarnings.create
~name:"deprecated-ssr-search"
Expand Down
5 changes: 4 additions & 1 deletion vernac/g_vernac.mlg
Original file line number Diff line number Diff line change
Expand Up @@ -35,6 +35,9 @@ open Attributes

let query_command = Entry.create "vernac:query_command"

let search_query = Entry.create "vernac:search_query"
let search_queries = Entry.create "vernac:search_queries"

let subprf = Entry.create "vernac:subprf"

let quoted_attributes = Entry.create "vernac:quoted_attributes"
Expand Down Expand Up @@ -819,7 +822,7 @@ GRAMMAR EXTEND Gram
END

GRAMMAR EXTEND Gram
GLOBAL: command query_command class_rawexpr gallina_ext;
GLOBAL: command query_command class_rawexpr gallina_ext search_query search_queries;

gallina_ext:
[ [ IDENT "Export"; "Set"; table = option_table; v = option_setting ->
Expand Down

0 comments on commit b7c14a8

Please sign in to comment.