aboutsummaryrefslogtreecommitdiffhomepage
path: root/tactics/auto.mli
diff options
context:
space:
mode:
authorGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2003-04-07 17:19:41 +0000
committerGravatar herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7>2003-04-07 17:19:41 +0000
commit928287134ab9dd23258c395589f8633e422e939f (patch)
tree192971793635b1057b78004b14df4ad5dfac9561 /tactics/auto.mli
parentc4ef643444acecdb4c08a74f37b9006ae060c5af (diff)
Globalisation des noms de tactiques dans les définitions de tactiques
pour compatibilité avec les modules. Globalisation partielle des invocations de tactiques hors définitions (partielle car noms des Intros/Assert/Inversion/... non connus). git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@3857 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'tactics/auto.mli')
-rw-r--r--tactics/auto.mli15
1 files changed, 11 insertions, 4 deletions
diff --git a/tactics/auto.mli b/tactics/auto.mli
index 4cd017e5f..5fe5ebb86 100644
--- a/tactics/auto.mli
+++ b/tactics/auto.mli
@@ -28,7 +28,7 @@ type auto_tactic =
| Give_exact of constr
| Res_pf_THEN_trivial_fail of constr * unit clausenv (* Hint Immediate *)
| Unfold_nth of global_reference (* Hint Unfold *)
- | Extern of Tacexpr.raw_tactic_expr (* Hint Extern *)
+ | Extern of Tacexpr.glob_tactic_expr (* Hint Extern *)
open Rawterm
@@ -114,11 +114,18 @@ val make_resolve_hyp :
(* [make_extern name pri pattern tactic_expr] *)
val make_extern :
- identifier -> int -> constr_pattern -> Tacexpr.raw_tactic_expr
+ identifier -> int -> constr_pattern -> Tacexpr.glob_tactic_expr
-> constr_label * pri_auto_tactic
val set_extern_interp :
- ((int * constr) list -> Tacexpr.raw_tactic_expr -> tactic) -> unit
+ ((int * constr) list -> Tacexpr.glob_tactic_expr -> tactic) -> unit
+
+val set_extern_intern_tac :
+ (int list -> Tacexpr.raw_tactic_expr -> Tacexpr.glob_tactic_expr) -> unit
+
+val set_extern_subst_tactic :
+ (Names.substitution -> Tacexpr.glob_tactic_expr -> Tacexpr.glob_tactic_expr)
+ -> unit
(* Create a Hint database from the pairs (name, constr).
Useful to take the current goal hypotheses as hints *)
@@ -137,7 +144,7 @@ val unify_resolve : (constr * unit clausenv) -> tactic
[Pattern.somatches], then replace [?1] [?2] metavars in tacast by the
right values to build a tactic *)
-val conclPattern : constr -> constr_pattern -> Tacexpr.raw_tactic_expr -> tactic
+val conclPattern : constr -> constr_pattern -> Tacexpr.glob_tactic_expr -> tactic
(* The Auto tactic *)