aboutsummaryrefslogtreecommitdiffhomepage
path: root/tactics/auto.mli
diff options
context:
space:
mode:
authorGravatar aspiwack <aspiwack@85f007b7-540e-0410-9357-904b9bb8a0f7>2013-11-02 15:34:01 +0000
committerGravatar aspiwack <aspiwack@85f007b7-540e-0410-9357-904b9bb8a0f7>2013-11-02 15:34:01 +0000
commit260965dcf60d793ba01110ace8945cf51ef6531f (patch)
treed07323383e16bb5a63492e2721cf0502ba931716 /tactics/auto.mli
parent328279514e65f47a689e2d23f132c43c86870c05 (diff)
Makes the new Proofview.tactic the basic type of Ltac.
On the compilation of Coq, we can see an increase of ~20% compile time on my completely non-scientific tests. Hopefully this can be fixed. There are a lot of low hanging fruits, but this is an iso-functionality commit. With a few exceptions which were not necessary for the compilation of the theories: - The declarative mode is not yet ported - The timeout tactical is currently deactivated because it needs some subtle I/O. The framework is ready to handle it, but I haven't done it yet. - For much the same reason, the ltac debugger is unplugged. It will be more difficult, but will eventually be back. A few comments: I occasionnally used a coercion from [unit Proofview.tactic] to the old [Prooftype.tactic]. It should work smoothely, but loses any backtracking information: the coerced tactics has at most one success. - It is used in autorewrite (it shouldn't be a problem there). Autorewrite's code is fairly old and tricky - It is used in eauto, mostly for "Hint Extern". It may be an issue as time goes as we might want to have various success in a "Hint Extern". But it would require a heavy port of eauto.ml4 - It is used in typeclass eauto, but with a little help from Matthieu, it should be easy to port the whole thing to the new tactic engine, actually simplifying the code. - It is used in fourier. I believe it to be inocuous. - It is used in firstorder and congruence. I think it's ok. Their code is somewhat intricate and I'm not sure they would be easy to actually port. - It is used heavily in Function. And honestly, I have no idea whether it can do harm or not. Updates: (11 June 2013) Pierre-Marie Pédrot contributed the rebase over his new stream based architecture for Ltac matching (r16533), which avoid painfully and expensively working around the exception-throwing control flow of the previous API. (11 October 2013) Rebasing over recent commits (somewhere in r16721-r16730) rendered a major bug in my implementation of Tacticals.New.tclREPEAT_MAIN apparent. It caused Field_theory.v to loop. The bug made rewrite !lemma, rewrite ?lemma and autorewrite incorrect (tclREPEAT_MAIN was essentially tclREPEAT, causing rewrites to be tried in the side-conditions of conditional rewrites as well). The new implementation makes Coq faster, but it is pretty much impossible to tell if it is significant at all. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16967 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'tactics/auto.mli')
-rw-r--r--tactics/auto.mli28
1 files changed, 14 insertions, 14 deletions
diff --git a/tactics/auto.mli b/tactics/auto.mli
index b8860f097..7ce351f43 100644
--- a/tactics/auto.mli
+++ b/tactics/auto.mli
@@ -175,7 +175,7 @@ val make_extern :
-> hint_entry
val extern_interp :
- (patvar_map -> Tacexpr.glob_tactic_expr -> tactic) Hook.t
+ (patvar_map -> Tacexpr.glob_tactic_expr -> unit Proofview.tactic) Hook.t
val extern_intern_tac :
(patvar list -> Tacexpr.raw_tactic_expr -> Tacexpr.glob_tactic_expr) Hook.t
@@ -205,7 +205,7 @@ val unify_resolve : Unification.unify_flags -> (constr * clausenv) -> tactic
[Pattern.somatches], then replace [?1] [?2] metavars in tacast by the
right values to build a tactic *)
-val conclPattern : constr -> constr_pattern option -> Tacexpr.glob_tactic_expr -> tactic
+val conclPattern : constr -> constr_pattern option -> Tacexpr.glob_tactic_expr -> unit Proofview.tactic
(** The Auto tactic *)
@@ -215,47 +215,47 @@ val conclPattern : constr -> constr_pattern option -> Tacexpr.glob_tactic_expr -
val make_db_list : hint_db_name list -> hint_db list
val auto : ?debug:Tacexpr.debug ->
- int -> open_constr list -> hint_db_name list -> tactic
+ int -> open_constr list -> hint_db_name list -> unit Proofview.tactic
(** Auto with more delta. *)
val new_auto : ?debug:Tacexpr.debug ->
- int -> open_constr list -> hint_db_name list -> tactic
+ int -> open_constr list -> hint_db_name list -> unit Proofview.tactic
(** auto with default search depth and with the hint database "core" *)
-val default_auto : tactic
+val default_auto : unit Proofview.tactic
(** auto with all hint databases except the "v62" compatibility database *)
val full_auto : ?debug:Tacexpr.debug ->
- int -> open_constr list -> tactic
+ int -> open_constr list -> unit Proofview.tactic
(** auto with all hint databases except the "v62" compatibility database
and doing delta *)
val new_full_auto : ?debug:Tacexpr.debug ->
- int -> open_constr list -> tactic
+ int -> open_constr list -> unit Proofview.tactic
(** auto with default search depth and with all hint databases
except the "v62" compatibility database *)
-val default_full_auto : tactic
+val default_full_auto : unit Proofview.tactic
(** The generic form of auto (second arg [None] means all bases) *)
val gen_auto : ?debug:Tacexpr.debug ->
- int option -> open_constr list -> hint_db_name list option -> tactic
+ int option -> open_constr list -> hint_db_name list option -> unit Proofview.tactic
(** The hidden version of auto *)
val h_auto : ?debug:Tacexpr.debug ->
- int option -> open_constr list -> hint_db_name list option -> tactic
+ int option -> open_constr list -> hint_db_name list option -> unit Proofview.tactic
(** Trivial *)
val trivial : ?debug:Tacexpr.debug ->
- open_constr list -> hint_db_name list -> tactic
+ open_constr list -> hint_db_name list -> unit Proofview.tactic
val gen_trivial : ?debug:Tacexpr.debug ->
- open_constr list -> hint_db_name list option -> tactic
+ open_constr list -> hint_db_name list option -> unit Proofview.tactic
val full_trivial : ?debug:Tacexpr.debug ->
- open_constr list -> tactic
+ open_constr list -> unit Proofview.tactic
val h_trivial : ?debug:Tacexpr.debug ->
- open_constr list -> hint_db_name list option -> tactic
+ open_constr list -> hint_db_name list option -> unit Proofview.tactic
val pr_autotactic : 'a auto_tactic -> Pp.std_ppcmds