aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--printing/printer.ml3
-rw-r--r--proofs/logic.ml14
-rw-r--r--proofs/logic.mli3
-rw-r--r--proofs/proof_type.mli1
-rw-r--r--proofs/tacmach.ml4
-rw-r--r--proofs/tacmach.mli1
-rw-r--r--tactics/tactics.ml13
7 files changed, 20 insertions, 19 deletions
diff --git a/printing/printer.ml b/printing/printer.ml
index 28fd92659e..6d54a5b3d5 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -733,9 +733,6 @@ let pr_prim_rule = function
str(if Termops.occur_meta c then "refine " else "exact ") ++
Constrextern.with_meta_as_hole pr_constr c
- | Move (id1,id2) ->
- (str"move " ++ pr_id id1 ++ Miscprint.pr_move_location pr_id id2)
-
(* Backwards compatibility *)
let prterm = pr_lconstr
diff --git a/proofs/logic.ml b/proofs/logic.ml
index aa0b9bac6f..e4c833627a 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -276,6 +276,11 @@ let move_hyp toleft (left,declfrom,right) hto =
List.fold_left (fun sign d -> push_named_context_val d sign)
right left
+let move_hyp_in_named_context hfrom hto sign =
+ let (left,right,declfrom,toleft) =
+ split_sign hfrom hto (named_context_of_val sign) in
+ move_hyp toleft (left,declfrom,right) hto
+
(**********************************************************************)
@@ -549,12 +554,3 @@ let prim_refiner r sigma goal =
let sgl = List.rev sgl in
let sigma = Goal.V82.partial_solution sigma goal oterm in
(sgl, sigma)
-
- | Move (hfrom, hto) ->
- let (left,right,declfrom,toleft) =
- split_sign hfrom hto (named_context_of_val sign) in
- let hyps' =
- move_hyp toleft (left,declfrom,right) hto in
- let (gl,ev,sigma) = mk_goal hyps' cl in
- let sigma = Goal.V82.partial_solution_to sigma goal gl ev in
- ([gl], sigma)
diff --git a/proofs/logic.mli b/proofs/logic.mli
index 2764d28c02..0dba9ef1ee 100644
--- a/proofs/logic.mli
+++ b/proofs/logic.mli
@@ -56,3 +56,6 @@ val catchable_exception : exn -> bool
val convert_hyp : bool -> Environ.named_context_val -> evar_map ->
Context.Named.Declaration.t -> Environ.named_context_val
+
+val move_hyp_in_named_context : Id.t -> Id.t Misctypes.move_location ->
+ Environ.named_context_val -> Environ.named_context_val
diff --git a/proofs/proof_type.mli b/proofs/proof_type.mli
index f7798a0edb..c120796220 100644
--- a/proofs/proof_type.mli
+++ b/proofs/proof_type.mli
@@ -22,7 +22,6 @@ open Misctypes
type prim_rule =
| Cut of bool * bool * Id.t * types
| Refine of constr
- | Move of Id.t * Id.t move_location
(** Nowadays, the only rules we'll consider are the primitive rules *)
diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml
index 50984c48e0..b9330ff007 100644
--- a/proofs/tacmach.ml
+++ b/proofs/tacmach.ml
@@ -121,15 +121,11 @@ let internal_cut_rev_no_check replace id t gl =
let refine_no_check c gl =
refiner (Refine c) gl
-let move_hyp_no_check id1 id2 gl =
- refiner (Move (id1,id2)) gl
-
(* Versions with consistency checks *)
let internal_cut b d t = with_check (internal_cut_no_check b d t)
let internal_cut_rev b d t = with_check (internal_cut_rev_no_check b d t)
let refine c = with_check (refine_no_check c)
-let move_hyp id id' = with_check (move_hyp_no_check id id')
(* Pretty-printers *)
diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli
index 100ed1522e..727efcf6dc 100644
--- a/proofs/tacmach.mli
+++ b/proofs/tacmach.mli
@@ -92,7 +92,6 @@ val refine_no_check : constr -> tactic
val internal_cut : bool -> Id.t -> types -> tactic
val internal_cut_rev : bool -> Id.t -> types -> tactic
val refine : constr -> tactic
-val move_hyp : Id.t -> Id.t move_location -> tactic
(** {6 Pretty-printing functions (debug only). } *)
val pr_gls : goal sigma -> Pp.std_ppcmds
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 8dfb526325..ce77df69f4 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -313,7 +313,18 @@ let apply_clear_request clear_flag dft c =
else Tacticals.New.tclIDTAC
(* Moving hypotheses *)
-let move_hyp id dest = Proofview.V82.tactic (Tacmach.move_hyp id dest)
+let move_hyp id dest =
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let env = Proofview.Goal.env gl in
+ let ty = Proofview.Goal.raw_concl gl in
+ let store = Proofview.Goal.extra gl in
+ let sign = named_context_val env in
+ let sign' = move_hyp_in_named_context id dest sign in
+ let env = reset_with_named_context sign' env in
+ Refine.refine ~unsafe:true { run = begin fun sigma ->
+ Evarutil.new_evar env sigma ~principal:true ~store ty
+ end }
+ end }
(* Renaming hypotheses *)
let rename_hyp repl =