summaryrefslogtreecommitdiff
path: root/pretyping/evarsolve.ml
diff options
context:
space:
mode:
Diffstat (limited to 'pretyping/evarsolve.ml')
-rw-r--r--pretyping/evarsolve.ml261
1 files changed, 147 insertions, 114 deletions
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
index 3bf6f376..4fd03084 100644
--- a/pretyping/evarsolve.ml
+++ b/pretyping/evarsolve.ml
@@ -7,11 +7,10 @@
(************************************************************************)
open Util
-open Errors
+open CErrors
open Names
open Term
open Vars
-open Context
open Environ
open Termops
open Evd
@@ -20,6 +19,7 @@ open Retyping
open Reductionops
open Evarutil
open Pretype_errors
+open Sigma.Notations
let normalize_evar evd ev =
match kind_of_term (whd_evar evd (mkEvar ev)) with
@@ -42,28 +42,39 @@ let get_polymorphic_positions f =
templ.template_param_levels)
| _ -> assert false
-let refresh_level evd s =
- match Evd.is_sort_variable evd s with
- | None -> true
- | Some l -> not (Evd.is_flexible_level evd l)
-
-let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) pbty env evd t =
+let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
+ pbty env evd t =
let evdref = ref evd in
let modified = ref false in
- let rec refresh status dir t =
- match kind_of_term t with
- | Sort (Type u as s) when
- (match Univ.universe_level u with
- | None -> true
- | Some l -> not onlyalg && refresh_level evd s) ->
+ (* direction: true for fresh universes lower than the existing ones *)
+ let refresh_sort status ~direction s =
let s' = evd_comb0 (new_sort_variable status) evdref in
let evd =
- if dir then set_leq_sort env !evdref s' s
+ if direction then set_leq_sort env !evdref s' s
else set_leq_sort env !evdref s s'
in
- modified := true; evdref := evd; mkSort s'
+ modified := true; evdref := evd; mkSort s'
+ in
+ let rec refresh ~onlyalg status ~direction t =
+ match kind_of_term t with
+ | Sort (Type u as s) ->
+ (match Univ.universe_level u with
+ | None -> refresh_sort status ~direction s
+ | Some l ->
+ (match Evd.universe_rigidity evd l with
+ | UnivRigid ->
+ if not onlyalg then refresh_sort status ~direction s
+ else t
+ | UnivFlexible alg ->
+ if onlyalg && alg then
+ (evdref := Evd.make_flexible_variable !evdref false l; t)
+ else t))
+ | Sort (Prop Pos as s) when refreshset && not direction ->
+ (* Cannot make a universe "lower" than "Set",
+ only refreshing when we want higher universes. *)
+ refresh_sort status ~direction s
| Prod (na,u,v) ->
- mkProd (na,u,refresh status dir v)
+ mkProd (na, u, refresh ~onlyalg status ~direction v)
| _ -> t
(** Refresh the types of evars under template polymorphic references *)
and refresh_term_evars onevars top t =
@@ -76,11 +87,11 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) pbty env evd t =
Array.iter (refresh_term_evars onevars false) args
| Evar (ev, a) when onevars ->
let evi = Evd.find !evdref ev in
- let ty' = refresh univ_flexible true evi.evar_concl in
+ let ty' = refresh ~onlyalg univ_flexible ~direction:true evi.evar_concl in
if !modified then
evdref := Evd.add !evdref ev {evi with evar_concl = ty'}
else ()
- | _ -> iter_constr (refresh_term_evars onevars false) t
+ | _ -> Constr.iter (refresh_term_evars onevars false) t
and refresh_polymorphic_positions args pos =
let rec aux i = function
| Some l :: ls ->
@@ -96,9 +107,11 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) pbty env evd t =
in
let t' =
if isArity t then
- (match pbty with
- | None -> t
- | Some dir -> refresh status dir t)
+ match pbty with
+ | None ->
+ (* No cumulativity needed, but we still need to refresh the algebraics *)
+ refresh ~onlyalg:true univ_flexible ~direction:false t
+ | Some direction -> refresh ~onlyalg status ~direction t
else (refresh_term_evars false true t; t)
in
if !modified then !evdref, t' else !evdref, t
@@ -140,7 +153,7 @@ let recheck_applications conv_algo env evdref t =
let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref x) args in
let rec aux i ty =
if i < Array.length argsty then
- match kind_of_term (whd_betadeltaiota env !evdref ty) with
+ match kind_of_term (whd_all env !evdref ty) with
| Prod (na, dom, codom) ->
(match conv_algo env !evdref Reduction.CUMUL argsty.(i) dom with
| Success evd -> evdref := evd;
@@ -163,7 +176,8 @@ type 'a update =
| UpdateWith of 'a
| NoUpdate
-let inst_of_vars sign = Array.map_of_list (fun (id,_,_) -> mkVar id) sign
+open Context.Named.Declaration
+let inst_of_vars sign = Array.map_of_list (mkVar % get_id) sign
let restrict_evar_key evd evk filter candidates =
match filter, candidates with
@@ -181,7 +195,9 @@ let restrict_evar_key evd evk filter candidates =
let candidates = match candidates with
| NoUpdate -> evi.evar_candidates
| UpdateWith c -> Some c in
- restrict_evar evd evk filter candidates
+ let sigma = Sigma.Unsafe.of_evar_map evd in
+ let Sigma (evk, sigma, _) = restrict_evar sigma evk filter candidates in
+ (Sigma.to_evar_map sigma, evk)
end
(* Restrict an applied evar and returns its restriction in the same context *)
@@ -206,32 +222,32 @@ let restrict_instance evd evk filter argsv =
let evi = Evd.find evd evk in
Filter.filter_array (Filter.compose (evar_filter evi) filter) argsv
+open Context.Rel.Declaration
let noccur_evar env evd evk c =
let cache = ref Int.Set.empty (* cache for let-ins *) in
- let rec occur_rec (k, env as acc) c =
+ let rec occur_rec check_types (k, env as acc) c =
match kind_of_term c with
| Evar (evk',args' as ev') ->
(match safe_evar_value evd ev' with
- | Some c -> occur_rec acc c
+ | Some c -> occur_rec check_types acc c
| None ->
if Evar.equal evk evk' then raise Occur
- else Array.iter (occur_rec acc) args')
+ else (if check_types then
+ occur_rec false acc (existential_type evd ev');
+ Array.iter (occur_rec check_types acc) args'))
| Rel i when i > k ->
- if not (Int.Set.mem (i-k) !cache) then
- (match pi2 (Environ.lookup_rel i env) with
- | None -> ()
- | Some b -> cache := Int.Set.add (i-k) !cache; occur_rec acc (lift i b))
- | Proj (p,c) ->
- let c =
- try Retyping.expand_projection env evd p c []
- with Retyping.RetypeError _ ->
- (* Can happen when called from w_unify which doesn't assign evars/metas
- eagerly enough *) c
- in occur_rec acc c
+ if not (Int.Set.mem (i-k) !cache) then
+ let decl = Environ.lookup_rel i env in
+ if check_types then
+ (cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (get_type decl)));
+ (match decl with
+ | LocalAssum _ -> ()
+ | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i b))
+ | Proj (p,c) -> occur_rec true acc c
| _ -> iter_constr_with_full_binders (fun rd (k,env) -> (succ k, push_rel rd env))
- occur_rec acc c
+ (occur_rec check_types) acc c
in
- try occur_rec (0,env) c; true with Occur -> false
+ try occur_rec false (0,env) c; true with Occur -> false
(***************************************)
(* Managing chains of local definitons *)
@@ -242,9 +258,11 @@ let noccur_evar env evd evk c =
variable in its family of aliased variables *)
let compute_var_aliases sign =
- List.fold_right (fun (id,b,c) aliases ->
- match b with
- | Some t ->
+ let open Context.Named.Declaration in
+ List.fold_right (fun decl aliases ->
+ let id = get_id decl in
+ match decl with
+ | LocalDef (_,t,_) ->
(match kind_of_term t with
| Var id' ->
let aliases_of_id =
@@ -252,27 +270,30 @@ let compute_var_aliases sign =
Id.Map.add id (aliases_of_id@[t]) aliases
| _ ->
Id.Map.add id [t] aliases)
- | None -> aliases)
+ | LocalAssum _ -> aliases)
sign Id.Map.empty
let compute_rel_aliases var_aliases rels =
- snd (List.fold_right (fun (_,b,u) (n,aliases) ->
- (n-1,
- match b with
- | Some t ->
- (match kind_of_term t with
- | Var id' ->
- let aliases_of_n =
- try Id.Map.find id' var_aliases with Not_found -> [] in
- Int.Map.add n (aliases_of_n@[t]) aliases
- | Rel p ->
- let aliases_of_n =
- try Int.Map.find (p+n) aliases with Not_found -> [] in
- Int.Map.add n (aliases_of_n@[mkRel (p+n)]) aliases
- | _ ->
- Int.Map.add n [lift n (mkCast(t,DEFAULTcast,u))] aliases)
- | None -> aliases))
- rels (List.length rels,Int.Map.empty))
+ snd (List.fold_right
+ (fun decl (n,aliases) ->
+ (n-1,
+ match decl with
+ | LocalDef (_,t,u) ->
+ (match kind_of_term t with
+ | Var id' ->
+ let aliases_of_n =
+ try Id.Map.find id' var_aliases with Not_found -> [] in
+ Int.Map.add n (aliases_of_n@[t]) aliases
+ | Rel p ->
+ let aliases_of_n =
+ try Int.Map.find (p+n) aliases with Not_found -> [] in
+ Int.Map.add n (aliases_of_n@[mkRel (p+n)]) aliases
+ | _ ->
+ Int.Map.add n [lift n (mkCast(t,DEFAULTcast,u))] aliases)
+ | LocalAssum _ -> aliases)
+ )
+ rels
+ (List.length rels,Int.Map.empty))
let make_alias_map env =
(* We compute the chain of aliases for each var and rel *)
@@ -306,13 +327,13 @@ let normalize_alias aliases x =
let normalize_alias_var var_aliases id =
destVar (normalize_alias (var_aliases,Int.Map.empty) (mkVar id))
-let extend_alias (_,b,_) (var_aliases,rel_aliases) =
+let extend_alias decl (var_aliases,rel_aliases) =
let rel_aliases =
Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (lift 1) l))
rel_aliases Int.Map.empty in
let rel_aliases =
- match b with
- | Some t ->
+ match decl with
+ | LocalDef(_,t,_) ->
(match kind_of_term t with
| Var id' ->
let aliases_of_binder =
@@ -324,7 +345,7 @@ let extend_alias (_,b,_) (var_aliases,rel_aliases) =
Int.Map.add 1 (aliases_of_binder@[mkRel (p+1)]) rel_aliases
| _ ->
Int.Map.add 1 [lift 1 t] rel_aliases)
- | None -> rel_aliases in
+ | LocalAssum _ -> rel_aliases in
(var_aliases, rel_aliases)
let expand_alias_once aliases x =
@@ -430,16 +451,17 @@ let get_actual_deps aliases l t =
| Rel n -> Int.Set.mem n fv_rels
| _ -> assert false) l
+open Context.Named.Declaration
let remove_instance_local_defs evd evk args =
let evi = Evd.find evd evk in
let len = Array.length args in
let rec aux sign i = match sign with
| [] ->
let () = assert (i = len) in []
- | (_, None, _) :: sign ->
+ | LocalAssum _ :: sign ->
let () = assert (i < len) in
(Array.unsafe_get args i) :: aux sign (succ i)
- | (_, Some _, _) :: sign ->
+ | LocalDef _ :: sign ->
aux sign (succ i)
in
aux (evar_filtered_context evi) 0
@@ -501,7 +523,8 @@ let solve_pattern_eqn env l c =
match kind_of_term a with
(* Rem: if [a] links to a let-in, do as if it were an assumption *)
| Rel n ->
- let d = map_rel_declaration (lift n) (lookup_rel n env) in
+ let open Context.Rel.Declaration in
+ let d = map_constr (lift n) (lookup_rel n env) in
mkLambda_or_LetIn d c'
| Var id ->
let d = lookup_named id env in mkNamedLambda_or_LetIn d c'
@@ -509,7 +532,7 @@ let solve_pattern_eqn env l c =
l c in
(* Warning: we may miss some opportunity to eta-reduce more since c'
is not in normal form *)
- whd_eta c'
+ shrink_eta c'
(*****************************************)
(* Refining/solving unification problems *)
@@ -530,9 +553,9 @@ let make_projectable_subst aliases sigma evi args =
let evar_aliases = compute_var_aliases sign in
let (_,full_subst,cstr_subst) =
List.fold_right
- (fun (id,b,c) (args,all,cstrs) ->
- match b,args with
- | None, a::rest ->
+ (fun decl (args,all,cstrs) ->
+ match decl,args with
+ | LocalAssum (id,c), a::rest ->
let a = whd_evar sigma a in
let cstrs =
let a',args = decompose_app_vect a in
@@ -542,7 +565,7 @@ let make_projectable_subst aliases sigma evi args =
Constrmap.add (fst cstr) ((args,id)::l) cstrs
| _ -> cstrs in
(rest,Id.Map.add id [a,normalize_alias_opt aliases a,id] all,cstrs)
- | Some c, a::rest ->
+ | LocalDef (id,c,_), a::rest ->
let a = whd_evar sigma a in
(match kind_of_term c with
| Var id' ->
@@ -571,7 +594,9 @@ let make_projectable_subst aliases sigma evi args =
*)
let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_sign sign filter inst_in_env =
- let evd,evar_in_env = new_evar_instance sign evd ty_t_in_sign ~filter ~src inst_in_env in
+ let evd = Sigma.Unsafe.of_evar_map evd in
+ let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd ty_t_in_sign ~filter ~src inst_in_env in
+ let evd = Sigma.to_evar_map evd in
let t_in_env = whd_evar evd t_in_env in
let evd = define_fun env evd None (destEvar evar_in_env) t_in_env in
let ctxt = named_context_of_val sign in
@@ -596,16 +621,24 @@ let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_si
* substitution u1..uq.
*)
+exception MorePreciseOccurCheckNeeeded
+
let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
+ if Evd.is_defined evd evk1 then
+ (* Some circularity somewhere (see e.g. #3209) *)
+ raise MorePreciseOccurCheckNeeeded;
+ let (evk1,args1) = destEvar (whd_evar evd (mkEvar (evk1,args1))) in
let evi1 = Evd.find_undefined evd evk1 in
let env1,rel_sign = env_rel_context_chop k env in
let sign1 = evar_hyps evi1 in
let filter1 = evar_filter evi1 in
let src = subterm_source evk1 evi1.evar_source in
- let ids1 = List.map pi1 (named_context_of_val sign1) in
+ let ids1 = List.map get_id (named_context_of_val sign1) in
let inst_in_sign = List.map mkVar (Filter.filter_list filter1 ids1) in
+ let open Context.Rel.Declaration in
let (sign2,filter2,inst2_in_env,inst2_in_sign,_,evd,_) =
- List.fold_right (fun (na,b,t_in_env as d) (sign,filter,inst_in_env,inst_in_sign,env,evd,avoid) ->
+ List.fold_right (fun d (sign,filter,inst_in_env,inst_in_sign,env,evd,avoid) ->
+ let LocalAssum (na,t_in_env) | LocalDef (na,_,t_in_env) = d in
let id = next_name_away na avoid in
let evd,t_in_sign =
let s = Retyping.get_sort_of env evd t_in_env in
@@ -613,13 +646,13 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
~status:univ_flexible (Some false) env evd (mkSort s) in
define_evar_from_virtual_equation define_fun env evd src t_in_env
ty_t_in_sign sign filter inst_in_env in
- let evd,b_in_sign = match b with
- | None -> evd,None
- | Some b ->
+ let evd,b_in_sign = match d with
+ | LocalAssum _ -> evd,None
+ | LocalDef (_,b,_) ->
let evd,b = define_evar_from_virtual_equation define_fun env evd src b
t_in_sign sign filter inst_in_env in
evd,Some b in
- (push_named_context_val (id,b_in_sign,t_in_sign) sign, Filter.extend 1 filter,
+ (push_named_context_val (Context.Named.Declaration.of_tuple (id,b_in_sign,t_in_sign)) sign, Filter.extend 1 filter,
(mkRel 1)::(List.map (lift 1) inst_in_env),
(mkRel 1)::(List.map (lift 1) inst_in_sign),
push_rel d env,evd,id::avoid))
@@ -632,8 +665,10 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
~status:univ_flexible (Some false) env evd (mkSort s) in
define_evar_from_virtual_equation define_fun env evd src ty_in_env
ty_t_in_sign sign2 filter2 inst2_in_env in
- let evd,ev2_in_sign =
+ let evd = Sigma.Unsafe.of_evar_map evd in
+ let Sigma (ev2_in_sign, evd, _) =
new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src inst2_in_sign in
+ let evd = Sigma.to_evar_map evd in
let ev2_in_env = (fst (destEvar ev2_in_sign), Array.of_list inst2_in_env) in
(evd, ev2_in_sign, ev2_in_env)
@@ -757,9 +792,10 @@ let project_with_effects aliases sigma effects t subst =
effects := p :: !effects;
c
+open Context.Named.Declaration
let rec find_solution_type evarenv = function
- | (id,ProjectVar)::l -> pi3 (lookup_named id evarenv)
- | [id,ProjectEvar _] -> (* bugged *) pi3 (lookup_named id evarenv)
+ | (id,ProjectVar)::l -> get_type (lookup_named id evarenv)
+ | [id,ProjectEvar _] -> (* bugged *) get_type (lookup_named id evarenv)
| (id,ProjectEvar _)::l -> find_solution_type evarenv l
| [] -> assert false
@@ -779,7 +815,7 @@ let rec do_projection_effects define_fun env ty evd = function
let evd = Evd.define evk (mkVar id) evd in
(* TODO: simplify constraints involving evk *)
let evd = do_projection_effects define_fun env ty evd p in
- let ty = whd_betadeltaiota env evd (Lazy.force ty) in
+ let ty = whd_all env evd (Lazy.force ty) in
if not (isSort ty) then
(* Don't try to instantiate if a sort because if evar_concl is an
evar it may commit to a univ level which is not the right
@@ -893,7 +929,7 @@ let invert_invertible_arg fullenv evd aliases k (evk,argsv) args' =
*)
let set_of_evctx l =
- List.fold_left (fun s (id,_,_) -> Id.Set.add id s) Id.Set.empty l
+ List.fold_left (fun s decl -> Id.Set.add (get_id decl) s) Id.Set.empty l
let filter_effective_candidates evi filter candidates =
match filter with
@@ -925,7 +961,13 @@ let closure_of_filter evd evk = function
| Some filter ->
let evi = Evd.find_undefined evd evk in
let vars = collect_vars (Evarutil.nf_evar evd (evar_concl evi)) in
- let test b (id,c,_) = b || Idset.mem id vars || match c with None -> false | Some c -> not (isRel c || isVar c) in
+ let test b decl = b || Idset.mem (get_id decl) vars ||
+ match decl with
+ | LocalAssum _ ->
+ false
+ | LocalDef (_,c,_) ->
+ not (isRel c || isVar c)
+ in
let newfilter = Filter.map_along test filter (evar_context evi) in
(* Now ensure that restriction is at least what is was originally *)
let newfilter = Option.cata (Filter.map_along (&&) newfilter) newfilter (Filter.repr (evar_filter evi)) in
@@ -1007,21 +1049,6 @@ let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs =
* Note: argument f is the function used to instantiate evars.
*)
-let are_canonical_instances args1 args2 env =
- let n1 = Array.length args1 in
- let n2 = Array.length args2 in
- let rec aux n = function
- | (id,_,c)::sign
- when n < n1 && isVarId id args1.(n) && isVarId id args2.(n) ->
- aux (n+1) sign
- | [] ->
- let rec aux2 n =
- Int.equal n n1 ||
- (isRelN (n1-n) args1.(n) && isRelN (n1-n) args2.(n) && aux2 (n+1))
- in aux2 n
- | _ -> false in
- Int.equal n1 n2 && aux 0 (named_context env)
-
let filter_compatible_candidates conv_algo env evd evi args rhs c =
let c' = instantiate_evar_array evi c args in
match conv_algo env evd Reduction.CONV rhs c' with
@@ -1296,7 +1323,7 @@ let occur_evar_upto_types sigma n c =
seen := Evar.Set.add sp !seen;
Option.iter occur_rec (existential_opt_value sigma e);
occur_rec (existential_type sigma e))
- | _ -> iter_constr occur_rec c
+ | _ -> Constr.iter occur_rec c
in
try occur_rec c; false with Occur -> true
@@ -1381,15 +1408,16 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
let t = whd_evar !evdref t in
match kind_of_term t with
| Rel i when i>k ->
- (match pi2 (Environ.lookup_rel (i-k) env') with
- | None -> project_variable (mkRel (i-k))
- | Some b ->
+ let open Context.Rel.Declaration in
+ (match Environ.lookup_rel (i-k) env' with
+ | LocalAssum _ -> project_variable (mkRel (i-k))
+ | LocalDef (_,b,_) ->
try project_variable (mkRel (i-k))
with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i b))
| Var id ->
- (match pi2 (Environ.lookup_named id env') with
- | None -> project_variable t
- | Some b ->
+ (match Environ.lookup_named id env' with
+ | LocalAssum _ -> project_variable t
+ | LocalDef (_,b,_) ->
try project_variable t
with NotInvertibleUsingOurAlgorithm _ -> imitate envk b)
| LetIn (na,b,u,c) ->
@@ -1450,7 +1478,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
imitate envk t in
t::l
- with e when Errors.noncritical e -> l in
+ with e when CErrors.noncritical e -> l in
(match candidates with
| [x] -> x
| _ ->
@@ -1469,7 +1497,8 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
let names = ref Idset.empty in
let rec is_id_subst ctxt s =
match ctxt, s with
- | ((id, _, _) :: ctxt'), (c :: s') ->
+ | (decl :: ctxt'), (c :: s') ->
+ let id = get_id decl in
names := Idset.add id !names;
isVarId id c && is_id_subst ctxt' s'
| [], [] -> true
@@ -1538,11 +1567,13 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
postpone_non_unique_projection env evd pbty ev sols rhs
| NotEnoughInformationEvarEvar t ->
add_conv_oriented_pb (pbty,env,mkEvar ev,t) evd
+ | MorePreciseOccurCheckNeeeded ->
+ add_conv_oriented_pb (pbty,env,mkEvar ev,rhs) evd
| NotInvertibleUsingOurAlgorithm _ | MetaOccurInBodyInternal as e ->
raise e
| OccurCheckIn (evd,rhs) ->
(* last chance: rhs actually reduces to ev *)
- let c = whd_betadeltaiota env evd rhs in
+ let c = whd_all env evd rhs in
match kind_of_term c with
| Evar (evk',argsv2) when Evar.equal evk evk' ->
solve_refl (fun env sigma pb c c' -> is_fconv pb env sigma c c')
@@ -1580,7 +1611,7 @@ let status_changed lev (pbty,_,t1,t2) =
(try Evar.Set.mem (head_evar t1) lev with NoHeadEvar -> false) ||
(try Evar.Set.mem (head_evar t2) lev with NoHeadEvar -> false)
-let reconsider_conv_pbs conv_algo evd =
+let reconsider_unif_constraints conv_algo evd =
let (evd,pbs) = extract_changed_conv_pbs evd status_changed in
List.fold_left
(fun p (pbty,env,t1,t2 as x) ->
@@ -1593,6 +1624,8 @@ let reconsider_conv_pbs conv_algo evd =
(Success evd)
pbs
+let reconsider_conv_pbs = reconsider_unif_constraints
+
(* Tries to solve problem t1 = t2.
* Precondition: t1 is an uninstantiated evar
* Returns an optional list of evars that were instantiated, or None
@@ -1603,7 +1636,7 @@ let solve_simple_eqn conv_algo ?(choose=false) env evd (pbty,(evk1,args1 as ev1)
try
let t2 = whd_betaiota evd t2 in (* includes whd_evar *)
let evd = evar_define conv_algo ~choose env evd pbty ev1 t2 in
- reconsider_conv_pbs conv_algo evd
+ reconsider_unif_constraints conv_algo evd
with
| NotInvertibleUsingOurAlgorithm t ->
UnifFailure (evd,NotClean (ev1,env,t))