aboutsummaryrefslogtreecommitdiff
path: root/vernac/search.mli
diff options
context:
space:
mode:
authorHugo Herbelin2020-05-11 17:41:58 +0200
committerThéo Zimmermann2020-05-15 18:05:09 +0200
commit96333c1bf27f33c15e2475c11c7bfefe87d3a6e1 (patch)
tree2d1545c2d416846b3e0c52b8537a437582a805c6 /vernac/search.mli
parenta28c4b025439a747354c2432223a3b30912ba182 (diff)
Search: new clauses for searching head, conclusion, kind...
- new clauses "hyp:", "concl:", "headhyp:" and "headconcl:" to restrict match to an hypothesis or the conclusion, possibly only at the head (like SearchHead in this latter case) - new clause "is:" to search by kind of object (for some list of kinds) - support for any combination of negations, disjunctions and conjunctions, using a syntax close to that of intropatterns.
Diffstat (limited to 'vernac/search.mli')
-rw-r--r--vernac/search.mli17
1 files changed, 13 insertions, 4 deletions
diff --git a/vernac/search.mli b/vernac/search.mli
index 5eeb77a6d0..3dc437d85a 100644
--- a/vernac/search.mli
+++ b/vernac/search.mli
@@ -12,15 +12,24 @@ open Names
open Constr
open Environ
open Pattern
+open Vernacexpr
(** {6 Search facilities. } *)
type glob_search_item =
- | GlobSearchSubPattern of constr_pattern
+ | GlobSearchSubPattern of glob_search_where * bool * constr_pattern
| GlobSearchString of string
+ | GlobSearchKind of Decls.logical_kind
+ | GlobSearchFilter of (GlobRef.t -> bool)
-type filter_function = GlobRef.t -> env -> constr -> bool
-type display_function = GlobRef.t -> env -> constr -> unit
+type glob_search_request =
+ | GlobSearchLiteral of glob_search_item
+ | GlobSearchDisjConj of (bool * glob_search_request) list list
+
+type filter_function =
+ GlobRef.t -> Decls.logical_kind option -> env -> constr -> bool
+type display_function =
+ GlobRef.t -> Decls.logical_kind option -> env -> constr -> unit
(** {6 Generic filter functions} *)
@@ -44,7 +53,7 @@ val search_rewrite : ?pstate:Declare.Proof.t -> int option -> constr_pattern ->
-> display_function -> unit
val search_pattern : ?pstate:Declare.Proof.t -> int option -> constr_pattern -> DirPath.t list * bool
-> display_function -> unit
-val search : ?pstate:Declare.Proof.t -> int option -> (bool * glob_search_item) list
+val search : ?pstate:Declare.Proof.t -> int option -> (bool * glob_search_request) list
-> DirPath.t list * bool -> display_function -> unit
type search_constraint =