aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2016-09-24 16:37:04 +0200
committerGravatar Hugo Herbelin <Hugo.Herbelin@inria.fr>2016-10-20 20:22:07 +0200
commita07f67f6f1deba8b14672c618c003ec345d7970a (patch)
treec9f33707c6c7dc3fae71d1e0a7e35e5686751a5d
parent317ae3b327d201530730ed2cce5f44e8763814d4 (diff)
A fix for #5097 (status of evars refined by "clear" in ltac: closed wrt evars).
If an existing evar was cleared in pretyping (typically while processing "ltac:"), it created an evar considered as new. Updating them instead along the "cleared" flag. If correct, I suspect similar treatment should be done for refining along "change", "rename" and "move".
-rw-r--r--pretyping/evarutil.ml22
-rw-r--r--pretyping/evarutil.mli6
-rw-r--r--pretyping/pretyping.ml31
-rw-r--r--proofs/proofview.ml36
-rw-r--r--test-suite/bugs/closed/5097.v7
5 files changed, 59 insertions, 43 deletions
diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml
index 759e0e4d6..9a9c946ae 100644
--- a/pretyping/evarutil.ml
+++ b/pretyping/evarutil.ml
@@ -600,6 +600,28 @@ let gather_dependent_evars evm l =
(* /spiwack *)
+(** [advance sigma g] returns [Some g'] if [g'] is undefined and is
+ the current avatar of [g] (for instance [g] was changed by [clear]
+ into [g']). It returns [None] if [g] has been (partially)
+ solved. *)
+(* spiwack: [advance] is probably performance critical, and the good
+ behaviour of its definition may depend sensitively to the actual
+ definition of [Evd.find]. Currently, [Evd.find] starts looking for
+ a value in the heap of undefined variable, which is small. Hence in
+ the most common case, where [advance] is applied to an unsolved
+ goal ([advance] is used to figure if a side effect has modified the
+ goal) it terminates quickly. *)
+let rec advance sigma evk =
+ let evi = Evd.find sigma evk in
+ match evi.evar_body with
+ | Evar_empty -> Some evk
+ | Evar_defined v ->
+ if Option.default false (Store.get evi.evar_extra cleared) then
+ let (evk,_) = Term.destEvar v in
+ advance sigma evk
+ else
+ None
+
(** The following functions return the set of undefined evars
contained in the object, the defined evars being traversed.
This is roughly a combination of the previous functions and
diff --git a/pretyping/evarutil.mli b/pretyping/evarutil.mli
index f68651a74..b60daae6d 100644
--- a/pretyping/evarutil.mli
+++ b/pretyping/evarutil.mli
@@ -123,6 +123,12 @@ val evar_absorb_arguments : env -> evar_map -> existential -> constr list ->
its (partial) definition. *)
val gather_dependent_evars : evar_map -> evar list -> (Evar.Set.t option) Evar.Map.t
+(** [advance sigma g] returns [Some g'] if [g'] is undefined and is
+ the current avatar of [g] (for instance [g] was changed by [clear]
+ into [g']). It returns [None] if [g] has been (partially)
+ solved. *)
+val advance : evar_map -> evar -> evar option
+
(** The following functions return the set of undefined evars
contained in the object, the defined evars being traversed.
This is roughly a combination of the previous functions and
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 2e164e540..c4ea79f95 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -183,17 +183,24 @@ type inference_flags = {
expand_evars : bool
}
-let frozen_holes (sigma, sigma') =
- (); fun ev -> Evar.Map.mem ev (Evd.undefined_map sigma)
-
-let pending_holes (sigma, sigma') =
- let fold evk _ accu =
- if not (Evd.mem sigma evk) then Evar.Set.add evk accu else accu
- in
- Evd.fold_undefined fold sigma' Evar.Set.empty
+(* Compute the set of still-undefined initial evars up to restriction
+ (e.g. clearing) and the set of yet-unsolved evars freshly created
+ in the extension [sigma'] of [sigma] (excluding the restrictions of
+ the undefined evars of [sigma] to be freshly created evars of
+ [sigma']). Otherwise said, we partition the undefined evars of
+ [sigma'] into those already in [sigma] or deriving from an evar in
+ [sigma] by restriction, and the evars properly created in [sigma'] *)
+
+let frozen_and_pending_holes (sigma, sigma') =
+ let add_derivative_of evk evi acc =
+ match advance sigma' evk with None -> acc | Some evk' -> Evar.Set.add evk' acc in
+ let frozen = Evd.fold_undefined add_derivative_of sigma Evar.Set.empty in
+ let fold evk _ accu = if not (Evar.Set.mem evk frozen) then Evar.Set.add evk accu else accu in
+ let pending = Evd.fold_undefined fold sigma' Evar.Set.empty in
+ (frozen,pending)
let apply_typeclasses env evdref frozen fail_evar =
- let filter_frozen = frozen in
+ let filter_frozen evk = Evar.Set.mem evk frozen in
evdref := Typeclasses.resolve_typeclasses
~filter:(if Flags.is_program_mode ()
then (fun evk evi -> Typeclasses.no_goals_or_obligations evk evi && not (filter_frozen evk))
@@ -244,8 +251,7 @@ let check_evars_are_solved env current_sigma frozen pending =
(* Try typeclasses, hooks, unification heuristics ... *)
let solve_remaining_evars flags env current_sigma pending =
- let frozen = frozen_holes pending in
- let pending = pending_holes pending in
+ let frozen,pending = frozen_and_pending_holes pending in
let evdref = ref current_sigma in
if flags.use_typeclasses then apply_typeclasses env evdref frozen false;
if Option.has_some flags.use_hook then
@@ -255,8 +261,7 @@ let solve_remaining_evars flags env current_sigma pending =
!evdref
let check_evars_are_solved env current_sigma pending =
- let frozen = frozen_holes pending in
- let pending = pending_holes pending in
+ let frozen,pending = frozen_and_pending_holes pending in
check_evars_are_solved env current_sigma frozen pending
let process_inference_flags flags env initial_sigma (sigma,c) =
diff --git a/proofs/proofview.ml b/proofs/proofview.ml
index 2fc404235..46a370d53 100644
--- a/proofs/proofview.ml
+++ b/proofs/proofview.ml
@@ -146,33 +146,9 @@ let focus i j sp =
let (new_comb, context) = focus_sublist i j sp.comb in
( { sp with comb = new_comb } , context )
-
-(** [advance sigma g] returns [Some g'] if [g'] is undefined and is
- the current avatar of [g] (for instance [g] was changed by [clear]
- into [g']). It returns [None] if [g] has been (partially)
- solved. *)
-(* spiwack: [advance] is probably performance critical, and the good
- behaviour of its definition may depend sensitively to the actual
- definition of [Evd.find]. Currently, [Evd.find] starts looking for
- a value in the heap of undefined variable, which is small. Hence in
- the most common case, where [advance] is applied to an unsolved
- goal ([advance] is used to figure if a side effect has modified the
- goal) it terminates quickly. *)
-let rec advance sigma g =
- let open Evd in
- let evi = Evd.find sigma g in
- match evi.evar_body with
- | Evar_empty -> Some g
- | Evar_defined v ->
- if Option.default false (Store.get evi.evar_extra Evarutil.cleared) then
- let (e,_) = Term.destEvar v in
- advance sigma e
- else
- None
-
(** [undefined defs l] is the list of goals in [l] which are still
unsolved (after advancing cleared goals). *)
-let undefined defs l = CList.map_filter (advance defs) l
+let undefined defs l = CList.map_filter (Evarutil.advance defs) l
(** Unfocuses a proofview with respect to a context. *)
let unfocus c sp =
@@ -429,7 +405,7 @@ let iter_goal i =
Comb.get >>= fun initial ->
Proof.List.fold_left begin fun (subgoals as cur) goal ->
Solution.get >>= fun step ->
- match advance step goal with
+ match Evarutil.advance step goal with
| None -> return cur
| Some goal ->
Comb.set [goal] >>
@@ -453,7 +429,7 @@ let fold_left2_goal i s l =
in
Proof.List.fold_left2 err begin fun ((r,subgoals) as cur) goal a ->
Solution.get >>= fun step ->
- match advance step goal with
+ match Evarutil.advance step goal with
| None -> return cur
| Some goal ->
Comb.set [goal] >>
@@ -497,7 +473,7 @@ let tclDISPATCHGEN0 join tacs =
let open Proof in
Pv.get >>= function
| { comb=[goal] ; solution } ->
- begin match advance solution goal with
+ begin match Evarutil.advance solution goal with
| None -> tclUNIT (join [])
| Some _ -> Proof.map (fun res -> join [res]) tac
end
@@ -1012,7 +988,7 @@ module Goal = struct
Pv.get >>= fun step ->
let sigma = step.solution in
let map goal =
- match advance sigma goal with
+ match Evarutil.advance sigma goal with
| None -> None (** ppedrot: Is this check really necessary? *)
| Some goal ->
let gl =
@@ -1026,7 +1002,7 @@ module Goal = struct
let unsolved { self=self } =
tclEVARMAP >>= fun sigma ->
- tclUNIT (not (Option.is_empty (advance sigma self)))
+ tclUNIT (not (Option.is_empty (Evarutil.advance sigma self)))
(* compatibility *)
let goal { self=self } = self
diff --git a/test-suite/bugs/closed/5097.v b/test-suite/bugs/closed/5097.v
new file mode 100644
index 000000000..37b239cf6
--- /dev/null
+++ b/test-suite/bugs/closed/5097.v
@@ -0,0 +1,7 @@
+(* Tracing existing evars along the weakening rule ("clear") *)
+Goal forall y, exists x, x=0->x=y.
+intros.
+eexists ?[x].
+intros.
+let x:=constr:(ltac:(clear y; exact 0)) in idtac x.
+Abort.