aboutsummaryrefslogtreecommitdiffhomepage
path: root/tactics/equality.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/equality.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/equality.mli')
-rw-r--r--tactics/equality.mli68
1 files changed, 34 insertions, 34 deletions
diff --git a/tactics/equality.mli b/tactics/equality.mli
index 70480d26e..3bb204caf 100644
--- a/tactics/equality.mli
+++ b/tactics/equality.mli
@@ -40,66 +40,66 @@ type conditions =
val general_rewrite_bindings :
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) -> constr with_bindings -> evars_flag -> tactic
+ ?tac:(unit Proofview.tactic * conditions) -> constr with_bindings -> evars_flag -> unit Proofview.tactic
val general_rewrite :
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) -> constr -> tactic
+ ?tac:(unit Proofview.tactic * conditions) -> constr -> unit Proofview.tactic
(* Equivalent to [general_rewrite l2r] *)
-val rewriteLR : ?tac:(tactic * conditions) -> constr -> tactic
-val rewriteRL : ?tac:(tactic * conditions) -> constr -> tactic
+val rewriteLR : ?tac:(unit Proofview.tactic * conditions) -> constr -> unit Proofview.tactic
+val rewriteRL : ?tac:(unit Proofview.tactic * conditions) -> constr -> unit Proofview.tactic
(* Warning: old [general_rewrite_in] is now [general_rewrite_bindings_in] *)
val general_rewrite_clause :
(Id.t option -> orientation ->
- occurrences -> constr with_bindings -> new_goals:constr list -> tactic) Hook.t
+ occurrences -> constr with_bindings -> new_goals:constr list -> unit Proofview.tactic) Hook.t
val is_applied_rewrite_relation : (env -> evar_map -> rel_context -> constr -> constr option) Hook.t
val general_rewrite_ebindings_clause : Id.t option ->
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) -> constr with_bindings -> evars_flag -> tactic
+ ?tac:(unit Proofview.tactic * conditions) -> constr with_bindings -> evars_flag -> unit Proofview.tactic
val general_rewrite_bindings_in :
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) ->
- Id.t -> constr with_bindings -> evars_flag -> tactic
+ ?tac:(unit Proofview.tactic * conditions) ->
+ Id.t -> constr with_bindings -> evars_flag -> unit Proofview.tactic
val general_rewrite_in :
orientation -> occurrences -> freeze_evars_flag -> dep_proof_flag ->
- ?tac:(tactic * conditions) -> Id.t -> constr -> evars_flag -> tactic
+ ?tac:(unit Proofview.tactic * conditions) -> Id.t -> constr -> evars_flag -> unit Proofview.tactic
val general_multi_rewrite :
- orientation -> evars_flag -> ?tac:(tactic * conditions) -> constr with_bindings -> clause -> tactic
+ orientation -> evars_flag -> ?tac:(unit Proofview.tactic * conditions) -> constr with_bindings -> clause -> unit Proofview.tactic
type delayed_open_constr_with_bindings =
env -> evar_map -> evar_map * constr with_bindings
val general_multi_multi_rewrite :
evars_flag -> (bool * multi * delayed_open_constr_with_bindings) list ->
- clause -> (tactic * conditions) option -> tactic
-
-val replace_in_clause_maybe_by : constr -> constr -> clause -> tactic option -> tactic
-val replace : constr -> constr -> tactic
-val replace_in : Id.t -> constr -> constr -> tactic
-val replace_by : constr -> constr -> tactic -> tactic
-val replace_in_by : Id.t -> constr -> constr -> tactic -> tactic
-
-val discr : evars_flag -> constr with_bindings -> tactic
-val discrConcl : tactic
-val discrClause : evars_flag -> clause -> tactic
-val discrHyp : Id.t -> tactic
-val discrEverywhere : evars_flag -> tactic
+ clause -> (unit Proofview.tactic * conditions) option -> unit Proofview.tactic
+
+val replace_in_clause_maybe_by : constr -> constr -> clause -> unit Proofview.tactic option -> unit Proofview.tactic
+val replace : constr -> constr -> unit Proofview.tactic
+val replace_in : Id.t -> constr -> constr -> unit Proofview.tactic
+val replace_by : constr -> constr -> unit Proofview.tactic -> unit Proofview.tactic
+val replace_in_by : Id.t -> constr -> constr -> unit Proofview.tactic -> unit Proofview.tactic
+
+val discr : evars_flag -> constr with_bindings -> unit Proofview.tactic
+val discrConcl : unit Proofview.tactic
+val discrClause : evars_flag -> clause -> unit Proofview.tactic
+val discrHyp : Id.t -> unit Proofview.tactic
+val discrEverywhere : evars_flag -> unit Proofview.tactic
val discr_tac : evars_flag ->
- constr with_bindings induction_arg option -> tactic
+ constr with_bindings induction_arg option -> unit Proofview.tactic
val inj : intro_pattern_expr Loc.located list option -> evars_flag ->
- constr with_bindings -> tactic
+ constr with_bindings -> unit Proofview.tactic
val injClause : intro_pattern_expr Loc.located list option -> evars_flag ->
- constr with_bindings induction_arg option -> tactic
-val injHyp : Id.t -> tactic
-val injConcl : tactic
+ constr with_bindings induction_arg option -> unit Proofview.tactic
+val injHyp : Id.t -> unit Proofview.tactic
+val injConcl : unit Proofview.tactic
-val dEq : evars_flag -> constr with_bindings induction_arg option -> tactic
-val dEqThen : evars_flag -> (constr -> int -> tactic) -> constr with_bindings induction_arg option -> tactic
+val dEq : evars_flag -> constr with_bindings induction_arg option -> unit Proofview.tactic
+val dEqThen : evars_flag -> (constr -> int -> unit Proofview.tactic) -> constr with_bindings induction_arg option -> unit Proofview.tactic
val make_iterated_tuple :
env -> evar_map -> constr -> (constr * types) -> constr * constr * constr
@@ -126,15 +126,15 @@ type subst_tactic_flags = {
only_leibniz : bool;
rewrite_dependent_proof : bool
}
-val subst_gen : bool -> Id.t list -> tactic
-val subst : Id.t list -> tactic
-val subst_all : ?flags:subst_tactic_flags -> tactic
+val subst_gen : bool -> Id.t list -> unit Proofview.tactic
+val subst : Id.t list -> unit Proofview.tactic
+val subst_all : ?flags:subst_tactic_flags -> unit Proofview.tactic
(* Replace term *)
(* [replace_multi_term dir_opt c cl]
perfoms replacement of [c] by the first value found in context
(according to [dir] if given to get the rewrite direction) in the clause [cl]
*)
-val replace_multi_term : bool option -> constr -> clause -> tactic
+val replace_multi_term : bool option -> constr -> clause -> unit Proofview.tactic
val set_eq_dec_scheme_kind : mutual scheme_kind -> unit