aboutsummaryrefslogtreecommitdiff
path: root/plugins
diff options
context:
space:
mode:
authorPierre-Marie Pédrot2020-08-25 16:26:01 +0200
committerPierre-Marie Pédrot2020-08-27 17:08:15 +0200
commit4109b8945103850c3e063e15c13aac9104a55846 (patch)
tree2aaaa81d128829230002b1e081a216d2614e4edc /plugins
parenta87c09c13028502ea86a553724a4131c5246145a (diff)
Remove a call to the old refiner in ssr.
Diffstat (limited to 'plugins')
-rw-r--r--plugins/ssr/ssrelim.ml11
1 files changed, 8 insertions, 3 deletions
diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml
index 1e182b52fa..582c45cde1 100644
--- a/plugins/ssr/ssrelim.ml
+++ b/plugins/ssr/ssrelim.ml
@@ -478,11 +478,16 @@ let revtoptac n0 =
Proofview.Goal.enter begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let concl = Proofview.Goal.concl gl in
+ let env = Proofview.Goal.env gl in
let n = nb_prod sigma concl - n0 in
let dc, cl = EConstr.decompose_prod_n_assum sigma n concl in
- let dc' = dc @ [Context.Rel.Declaration.LocalAssum(make_annot (Name rev_id) Sorts.Relevant, EConstr.it_mkProd_or_LetIn cl (List.rev dc))] in
- let f = EConstr.it_mkLambda_or_LetIn (mkEtaApp (EConstr.mkRel (n + 1)) (-n) 1) dc' in
- Logic.refiner ~check:true EConstr.Unsafe.(to_constr (EConstr.mkApp (f, [|Evarutil.mk_new_meta ()|])))
+ let ty = EConstr.it_mkProd_or_LetIn cl (List.rev dc) in
+ let dc' = dc @ [Context.Rel.Declaration.LocalAssum(make_annot (Name rev_id) Sorts.Relevant, ty)] in
+ Refine.refine ~typecheck:true begin fun sigma ->
+ let f = EConstr.it_mkLambda_or_LetIn (mkEtaApp (EConstr.mkRel (n + 1)) (-n) 1) dc' in
+ let sigma, ev = Evarutil.new_evar env sigma ty in
+ sigma, (EConstr.mkApp (f, [|ev|]))
+ end
end
let nothing_to_inject =