aboutsummaryrefslogtreecommitdiff
path: root/tactics/autorewrite.ml
diff options
context:
space:
mode:
authorEmilio Jesus Gallego Arias2019-02-13 05:37:09 +0100
committerEmilio Jesus Gallego Arias2019-03-27 23:56:18 +0100
commitc0cff3a7ebb79d1142090108c56e9aa64c3b481d (patch)
treeb316f32c8c5d879324031dd17ca317cb29ce4b1f /tactics/autorewrite.ml
parent178672504f1c92b162c2575b14034cc7b698b6a4 (diff)
[geninterp] Track polymorphic status in tactic interpretation.
Diffstat (limited to 'tactics/autorewrite.ml')
-rw-r--r--tactics/autorewrite.ml8
1 files changed, 6 insertions, 2 deletions
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index d9c0a26f91..51708670f5 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -99,11 +99,15 @@ let one_base general_rewrite_maybe_in tac_main bas =
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma)
(general_rewrite_maybe_in dir c' tc)
end in
- let lrul = List.map (fun h ->
+ let open Proofview.Notations in
+ Proofview.tclProofInfo [@ocaml.warning "-3"] >>= fun (_name, poly) ->
+ let lrul = List.map (fun h ->
let tac = match h.rew_tac with
| None -> Proofview.tclUNIT ()
| Some (Genarg.GenArg (Genarg.Glbwit wit, tac)) ->
- let ist = { Geninterp.lfun = Id.Map.empty; extra = Geninterp.TacStore.empty } in
+ let ist = { Geninterp.lfun = Id.Map.empty
+ ; poly
+ ; extra = Geninterp.TacStore.empty } in
Ftactic.run (Geninterp.interp wit ist tac) (fun _ -> Proofview.tclUNIT ())
in
(h.rew_ctx,h.rew_lemma,h.rew_l2r,tac)) lrul in