From 8ab4f8f76958d2603858ad3e4073be61ad38d113 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 5 Apr 2017 23:17:56 +0200 Subject: evd: restrict_evar with candidates, can raise NoCandidatesLeft When restricting an evar with candidates, raise an exception if this restriction would leave the evar without candidates, i.e. unsolvable. - evarutil: mark restricted evars as "cleared" They would otherwise escape being catched by the [advance] function of clenv, and result in dangling evars not being registered to the shelf. - engine: restrict_evar marks it cleared, update the future goals We make the new evar a future goal and remove the old one. If we did nothing, [unshelve tac] would work correctly as it uses [Proofview.advance] to find the shelved goals, going through the cleared evar. But [Unshelve] would fail as it expects only undefined evars on the shelf and throws away the defined ones. --- engine/evarutil.ml | 30 +++++++++++++++++++++++++----- engine/evarutil.mli | 12 +++++++++--- engine/evd.ml | 3 +++ engine/evd.mli | 4 +++- 4 files changed, 40 insertions(+), 9 deletions(-) (limited to 'engine') diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 82be4791f..cdc2897c8 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -426,10 +426,6 @@ let push_rel_context_to_named_context ?hypnaming env sigma typ = let default_source = Loc.tag @@ Evar_kinds.InternalHole -let restrict_evar evd evk filter ?src candidates = - let evd, evk' = Evd.restrict evk filter ?candidates ?src evd in - Evd.declare_future_goal evk' evd, evk' - let new_pure_evar_full evd evi = let (evd, evk) = Evd.new_evar evd evi in let evd = Evd.declare_future_goal evk evd in @@ -532,11 +528,33 @@ let generalize_evar_over_rels sigma (ev,args) = type clear_dependency_error = | OccurHypInSimpleClause of Id.t option | EvarTypingBreak of existential +| NoCandidatesLeft of Evar.t exception ClearDependencyError of Id.t * clear_dependency_error * GlobRef.t option exception Depends of Id.t +let set_of_evctx l = + List.fold_left (fun s decl -> Id.Set.add (NamedDecl.get_id decl) s) Id.Set.empty l + +let filter_effective_candidates evd evi filter candidates = + let ids = set_of_evctx (Filter.filter_list filter (evar_context evi)) in + List.filter (fun a -> Id.Set.subset (collect_vars evd a) ids) candidates + +let restrict_evar evd evk filter ?src candidates = + let evar_info = Evd.find_undefined evd evk in + let candidates = Option.map (filter_effective_candidates evd evar_info filter) candidates in + match candidates with + | Some [] -> raise (ClearDependencyError (*FIXME*)(Id.of_string "blah", (NoCandidatesLeft evk), None)) + | _ -> + let evd, evk' = Evd.restrict evk filter ?candidates ?src evd in + (** Mark new evar as future goal, removing previous one, + circumventing Proofview.advance but making Proof.run_tactic catch these. *) + let future_goals = Evd.save_future_goals evd in + let future_goals = Evd.filter_future_goals (fun evk' -> not (Evar.equal evk evk')) future_goals in + let evd = Evd.restore_future_goals evd future_goals in + (Evd.declare_future_goal evk' evd, evk') + let rec check_and_clear_in_constr env evdref err ids global c = (* returns a new constr where all the evars have been 'cleaned' (ie the hypotheses ids have been removed from the contexts of @@ -606,7 +624,9 @@ let rec check_and_clear_in_constr env evdref err ids global c = let origfilter = Evd.evar_filter evi in let filter = Evd.Filter.apply_subfilter origfilter filter in let evd = !evdref in - let (evd,_) = restrict_evar evd evk filter None in + let candidates = Evd.evar_candidates evi in + let candidates = Option.map (List.map EConstr.of_constr) candidates in + let (evd,_) = restrict_evar evd evk filter candidates in evdref := evd; Evd.existential_value0 !evdref ev diff --git a/engine/evarutil.mli b/engine/evarutil.mli index c17f3d168..1003c4feb 100644 --- a/engine/evarutil.mli +++ b/engine/evarutil.mli @@ -65,9 +65,6 @@ val new_type_evar : val new_Type : ?rigid:rigid -> env -> evar_map -> evar_map * constr -val restrict_evar : evar_map -> Evar.t -> Filter.t -> - ?src:Evar_kinds.t Loc.located -> constr list option -> evar_map * Evar.t - (** Polymorphic constants *) val new_global : evar_map -> GlobRef.t -> evar_map * constr @@ -223,9 +220,18 @@ raise OccurHypInSimpleClause if the removal breaks dependencies *) type clear_dependency_error = | OccurHypInSimpleClause of Id.t option | EvarTypingBreak of Constr.existential +| NoCandidatesLeft of Evar.t exception ClearDependencyError of Id.t * clear_dependency_error * GlobRef.t option +(** Restrict an undefined evar according to a (sub)filter and candidates. + The evar will be defined if there is only one candidate left, +@raise ClearDependencyError NoCandidatesLeft if the filter turns the candidates + into an empty list. *) + +val restrict_evar : evar_map -> Evar.t -> Filter.t -> + ?src:Evar_kinds.t Loc.located -> constr list option -> evar_map * Evar.t + val clear_hyps_in_evi : env -> evar_map -> named_context_val -> types -> Id.Set.t -> evar_map * named_context_val * types diff --git a/engine/evd.ml b/engine/evd.ml index 0c9c3a29b..7e0c397b2 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -171,6 +171,8 @@ let evar_context evi = named_context_of_val evi.evar_hyps let evar_filtered_context evi = Filter.filter_list (evar_filter evi) (evar_context evi) +let evar_candidates evi = evi.evar_candidates + let evar_hyps evi = evi.evar_hyps let evar_filtered_hyps evi = match Filter.repr (evar_filter evi) with @@ -620,6 +622,7 @@ let merge_universe_context evd uctx' = let set_universe_context evd uctx' = { evd with universes = uctx' } +(* TODO: make unique *) let add_conv_pb ?(tail=false) pb d = (** MS: we have duplicates here, why? *) if tail then {d with conv_pbs = d.conv_pbs @ [pb]} diff --git a/engine/evd.mli b/engine/evd.mli index c40e925d8..83fd00065 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -113,6 +113,7 @@ val evar_filtered_context : evar_info -> (econstr, etypes) Context.Named.pt val evar_hyps : evar_info -> named_context_val val evar_filtered_hyps : evar_info -> named_context_val val evar_body : evar_info -> evar_body +val evar_candidates : evar_info -> constr list option val evar_filter : evar_info -> Filter.t val evar_env : evar_info -> env val evar_filtered_env : evar_info -> env @@ -243,7 +244,8 @@ val evars_reset_evd : ?with_conv_pbs:bool -> ?with_univs:bool -> val restrict : Evar.t-> Filter.t -> ?candidates:econstr list -> ?src:Evar_kinds.t located -> evar_map -> evar_map * Evar.t (** Restrict an undefined evar into a new evar by filtering context and - possibly limiting the instances to a set of candidates *) + possibly limiting the instances to a set of candidates (candidates + are filtered according to the filter) *) val is_restricted_evar : evar_info -> Evar.t option (** Tell if an evar comes from restriction of another evar, and if yes, which *) -- cgit v1.2.3