aboutsummaryrefslogtreecommitdiffhomepage
path: root/tactics/eauto.mli
diff options
context:
space:
mode:
authorGravatar msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7>2008-02-08 17:03:37 +0000
committerGravatar msozeau <msozeau@85f007b7-540e-0410-9357-904b9bb8a0f7>2008-02-08 17:03:37 +0000
commit6703700903619004f05ad56293b7ec0a2e672d3a (patch)
tree7686794722387220929994965c01dc6642d5e8e0 /tactics/eauto.mli
parent7e324da8bd211f01593952ac51bd309e80c7546a (diff)
Change implementation of resolution for typeclasses to use a customized
eauto instead of an arbitrary tactic. Export more from eauto to allow easier debugging. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@10534 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'tactics/eauto.mli')
-rw-r--r--tactics/eauto.mli40
1 files changed, 39 insertions, 1 deletions
diff --git a/tactics/eauto.mli b/tactics/eauto.mli
index 34c4cab78..940648c2e 100644
--- a/tactics/eauto.mli
+++ b/tactics/eauto.mli
@@ -14,6 +14,7 @@ open Auto
open Topconstr
open Evd
open Environ
+open Explore
(*i*)
val rawwit_hintbases : hint_db_name list option raw_abstract_argument_type
@@ -26,10 +27,47 @@ val registered_e_assumption : tactic
val e_give_exact_constr : constr -> tactic
+type search_state = {
+ depth : int; (*r depth of search before failing *)
+ tacres : goal list sigma * validation;
+ last_tactic : Pp.std_ppcmds;
+ dblist : Auto.Hint_db.t list;
+ localdb : Auto.Hint_db.t list }
+
+module SearchProblem : sig
+ type state = search_state
+
+ val filter_tactics : Proof_type.goal list Tacmach.sigma * (Proof_type.proof_tree list -> 'a) ->
+ (Tacmach.tactic * Pp.std_ppcmds) list ->
+ ((Proof_type.goal list Tacmach.sigma * (Proof_type.proof_tree list -> 'a)) *
+ Pp.std_ppcmds)
+ list
+
+ val compare : search_state -> search_state -> int
+
+ val branching : state -> state list
+ val success : state -> bool
+ val pp : state -> unit
+end
+
+module Search : sig
+ val depth_first : search_state -> search_state
+ val debug_depth_first : search_state -> search_state
+
+ val breadth_first : search_state -> search_state
+ val debug_breadth_first : search_state -> search_state
+end
+
val full_eauto_gls : bool -> bool * int -> constr list -> goal list sigma * validation ->
goal list sigma * validation
-val resolve_all_evars : env -> (constr -> bool) -> evar_defs -> evar_defs
+val resolve_all_evars : bool -> bool * int -> env -> (evar -> evar_info -> bool) -> evar_defs ->
+ evar_defs
+
+val valid : Evd.evar_map ->
+ (Evd.evar -> Evd.evar_info -> bool) ->
+ Evd.evar_map ref -> Proof_type.proof_tree list -> 'a
+
val gen_eauto : bool -> bool * int -> constr list ->
hint_db_name list option -> tactic