aboutsummaryrefslogtreecommitdiffhomepage
path: root/tactics/tacticals.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/tacticals.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/tacticals.mli')
-rw-r--r--tactics/tacticals.mli83
1 files changed, 82 insertions, 1 deletions
diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli
index 13eaaff5c..58aff5fdb 100644
--- a/tactics/tacticals.mli
+++ b/tactics/tacticals.mli
@@ -52,7 +52,6 @@ val tclAT_LEAST_ONCE : tactic -> tactic
val tclFAIL : int -> std_ppcmds -> tactic
val tclFAIL_lazy : int -> std_ppcmds Lazy.t -> tactic
val tclDO : int -> tactic -> tactic
-val tclTIMEOUT : int -> tactic -> tactic
val tclWEAK_PROGRESS : tactic -> tactic
val tclPROGRESS : tactic -> tactic
val tclSHOWHYPS : tactic -> tactic
@@ -173,3 +172,85 @@ val simple_elimination_then :
val elim_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic
val case_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic
+
+(** Tacticals defined directly in term of Proofview *)
+module New : sig
+ open Proofview
+
+
+ val tclTHEN : unit tactic -> unit tactic -> unit tactic
+ (* [tclFAIL n msg] fails with [msg] as an error message at level [n]
+ (meaning that it will jump over [n] error catching tacticals FROM
+ THIS MODULE. *)
+ val tclFAIL : int -> Pp.std_ppcmds -> 'a tactic
+ val tclORELSE0 : 'a tactic -> 'a tactic -> 'a tactic
+ val tclORELSE : 'a tactic -> 'a tactic -> 'a tactic
+
+ (** [tclTHENS3PARTS tac1 [|t1 ; ... ; tn|] tac2 [|t'1 ; ... ; t'm|]
+ gls] applies the tactic [tac1] to [gls] then, applies [t1], ...,
+ [tn] to the first [n] resulting subgoals, [t'1], ..., [t'm] to the
+ last [m] subgoals and [tac2] to the rest of the subgoals in the
+ middle. Raises an error if the number of resulting subgoals is
+ strictly less than [n+m] *)
+ val tclTHENS3PARTS : unit tactic -> unit tactic array -> unit tactic -> unit tactic array -> unit tactic
+ val tclTHENSFIRSTn : unit tactic -> unit tactic array -> unit tactic -> unit tactic
+ val tclTHENFIRSTn : unit tactic -> unit tactic array -> unit tactic
+ (** [tclTHENFIRST tac1 tac2 gls] applies the tactic [tac1] to [gls]
+ and [tac2] to the first resulting subgoal *)
+ val tclTHENFIRST : unit tactic -> unit tactic -> unit tactic
+ val tclTHENLASTn : unit tactic -> unit tactic array -> unit tactic
+ val tclTHENLAST : unit tactic -> unit tactic -> unit tactic
+ (* [tclTHENS t l = t <*> tclDISPATCH l] *)
+ val tclTHENS : unit tactic -> unit tactic list -> unit tactic
+ (* [tclTHENLIST [t1;…;tn]] is [t1<*>…<*>tn] *)
+ val tclTHENLIST : unit tactic list -> unit tactic
+
+ (** [tclMAP f [x1..xn]] builds [(f x1);(f x2);...(f xn)] *)
+ val tclMAP : ('a -> unit tactic) -> 'a list -> unit tactic
+
+ val tclTRY : unit tactic -> unit tactic
+ val tclFIRST : 'a tactic list -> 'a tactic
+ val tclFIRST_PROGRESS_ON : ('a -> unit tactic) -> 'a list -> unit tactic
+ val tclIFTHENELSE : unit tactic -> unit tactic -> unit tactic -> unit tactic
+ val tclIFTHENSVELSE : unit tactic -> unit tactic array -> unit tactic -> unit tactic
+ val tclIFTHENTRYELSEMUST : unit tactic -> unit tactic -> unit tactic
+
+ val tclDO : int -> unit tactic -> unit tactic
+ val tclREPEAT : unit tactic -> unit tactic
+ (* Repeat on the first subgoal (no failure if no more subgoal) *)
+ val tclREPEAT_MAIN : unit tactic -> unit tactic
+ val tclCOMPLETE : 'a tactic -> 'a tactic
+ val tclSOLVE : unit tactic list -> unit tactic
+ val tclWITHHOLES : bool -> ('a -> unit tactic) -> Evd.evar_map -> 'a -> unit tactic
+
+
+ val ifOnHyp : (identifier * types -> bool) ->
+ (identifier -> unit Proofview.tactic) -> (identifier -> unit Proofview.tactic) ->
+ identifier -> unit Proofview.tactic
+
+ val onNthHypId : int -> (identifier -> unit tactic) -> unit tactic
+ val onLastHypId : (identifier -> unit tactic) -> unit tactic
+ val onLastHyp : (constr -> unit tactic) -> unit tactic
+ val onLastDecl : (named_declaration -> unit tactic) -> unit tactic
+
+ val tryAllHyps : (identifier -> unit tactic) -> unit tactic
+ val tryAllHypsAndConcl : (identifier option -> unit tactic) -> unit tactic
+ val onClause : (identifier option -> unit tactic) -> clause -> unit tactic
+
+ val elimination_then :
+ (branch_args -> unit Proofview.tactic) ->
+ (arg_bindings * arg_bindings) -> constr -> unit Proofview.tactic
+
+ val case_then_using :
+ intro_pattern_expr located option -> (branch_args -> unit Proofview.tactic) ->
+ constr option -> (arg_bindings * arg_bindings) ->
+ inductive -> clausenv -> unit Proofview.tactic
+
+ val case_nodep_then_using :
+ intro_pattern_expr located option -> (branch_args -> unit Proofview.tactic) ->
+ constr option -> (arg_bindings * arg_bindings) ->
+ inductive -> clausenv -> unit Proofview.tactic
+
+ val elim_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic
+ val case_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic
+end