aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--proofs/refiner.ml7
-rw-r--r--proofs/refiner.mli4
-rw-r--r--tactics/extratactics.ml43
-rw-r--r--tactics/tacticals.ml1
-rw-r--r--tactics/tacticals.mli1
5 files changed, 11 insertions, 5 deletions
diff --git a/proofs/refiner.ml b/proofs/refiner.ml
index 76fe0129f..7010153ba 100644
--- a/proofs/refiner.ml
+++ b/proofs/refiner.ml
@@ -223,9 +223,7 @@ let refiner = function
ref = Some(Daimon,[])})
-let local_Constraints gl = refiner (Prim Change_evars) gl
-
-let norm_evar_tac = local_Constraints
+let norm_evar_tac gl = refiner (Prim Change_evars) gl
let norm_evar_proof sigma pf =
let nf_subgoal i sgl =
@@ -312,6 +310,9 @@ let idtac_valid = function
(* [goal_goal_list : goal sigma -> goal list sigma] *)
let goal_goal_list gls = {it=[gls.it];sigma=gls.sigma}
+(* forces propagation of evar constraints *)
+let tclNORMEVAR = norm_evar_tac
+
(* identity tactic without any message *)
let tclIDTAC gls = (goal_goal_list gls, idtac_valid)
diff --git a/proofs/refiner.mli b/proofs/refiner.mli
index 530a72a37..ba092e89e 100644
--- a/proofs/refiner.mli
+++ b/proofs/refiner.mli
@@ -47,7 +47,6 @@ val frontier : transformation_tactic
val list_pf : proof_tree -> goal list
val unTAC : tactic -> goal sigma -> proof_tree sigma
-val local_Constraints : tactic
(* [frontier_map f n p] applies f on the n-th open subgoal of p and
rebuilds proof-tree.
@@ -61,6 +60,9 @@ val frontier_mapi :
(*s Tacticals. *)
+(* [tclNORMEVAR] forces propagation of evar constraints *)
+val tclNORMEVAR : tactic
+
(* [tclIDTAC] is the identity tactic without message printing*)
val tclIDTAC : tactic
val tclIDTAC_MESSAGE : Pp.std_ppcmds -> tactic
diff --git a/tactics/extratactics.ml4 b/tactics/extratactics.ml4
index 3ddb4188b..ef6ac60f4 100644
--- a/tactics/extratactics.ml4
+++ b/tactics/extratactics.ml4
@@ -290,16 +290,17 @@ TACTIC EXTEND evar
END
open Tacexpr
+open Tacticals
TACTIC EXTEND instantiate
[ "instantiate" "(" integer(i) ":=" raw(c) ")" hloc(hl) ] ->
[instantiate i c hl ]
+| [ "instantiate" ] -> [ tclNORMEVAR ]
END
(** Nijmegen "step" tactic for setoid rewriting *)
-open Tacticals
open Tactics
open Tactics
open Libnames
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index f13897d1f..e5093a947 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -37,6 +37,7 @@ open Tacexpr
(* Tacticals re-exported from the Refiner module.*)
(*************************************************)
+let tclNORMEVAR = tclNORMEVAR
let tclIDTAC = tclIDTAC
let tclIDTAC_MESSAGE = tclIDTAC_MESSAGE
let tclORELSE = tclORELSE
diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli
index de6ce5964..c84218210 100644
--- a/tactics/tacticals.mli
+++ b/tactics/tacticals.mli
@@ -24,6 +24,7 @@ open Tacexpr
(* Tacticals i.e. functions from tactics to tactics. *)
+val tclNORMEVAR : tactic
val tclIDTAC : tactic
val tclIDTAC_MESSAGE : std_ppcmds -> tactic
val tclORELSE : tactic -> tactic -> tactic