aboutsummaryrefslogtreecommitdiff
path: root/engine/evd.ml
diff options
context:
space:
mode:
authorHugo Herbelin2018-11-02 00:17:47 +0100
committerHugo Herbelin2019-05-13 18:23:58 +0200
commit6608f64f001f8f1a50b2dc41fefdf63c0b84b270 (patch)
tree8e320caff90f9fb0bf774ae038b5fb95df985150 /engine/evd.ml
parent6211fd6e067e781a160db8765dd87067428048f2 (diff)
Passing evar_map to evars_of_term rather than expecting the term to be evar-nf.
Diffstat (limited to 'engine/evd.ml')
-rw-r--r--engine/evd.ml54
1 files changed, 27 insertions, 27 deletions
diff --git a/engine/evd.ml b/engine/evd.ml
index d37b49e2dc..0f10a380d3 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -823,33 +823,6 @@ let loc_of_conv_pb evd (pbty,env,t1,t2) =
| Evar (evk2,_) -> fst (evar_source evk2 evd)
| _ -> None
-(** The following functions return the set of evars immediately
- contained in the object *)
-
-(* excluding defined evars *)
-
-let evars_of_term c =
- let rec evrec acc c =
- match kind c with
- | Evar (n, l) -> Evar.Set.add n (Array.fold_left evrec acc l)
- | _ -> Constr.fold evrec acc c
- in
- evrec Evar.Set.empty c
-
-let evars_of_named_context nc =
- Context.Named.fold_outside
- (NamedDecl.fold_constr (fun constr s -> Evar.Set.union s (evars_of_term constr)))
- nc
- ~init:Evar.Set.empty
-
-let evars_of_filtered_evar_info evi =
- Evar.Set.union (evars_of_term evi.evar_concl)
- (Evar.Set.union
- (match evi.evar_body with
- | Evar_empty -> Evar.Set.empty
- | Evar_defined b -> evars_of_term b)
- (evars_of_named_context (evar_filtered_context evi)))
-
(**********************************************************)
(* Sort variables *)
@@ -1404,3 +1377,30 @@ module MiniEConstr = struct
let to_rel_decl sigma d = Context.Rel.Declaration.map_constr (to_constr sigma) d
end
+
+(** The following functions return the set of evars immediately
+ contained in the object *)
+
+(* excluding defined evars *)
+
+let evars_of_term evd c =
+ let rec evrec acc c =
+ match MiniEConstr.kind evd c with
+ | Evar (n, l) -> Evar.Set.add n (Array.fold_left evrec acc l)
+ | _ -> Constr.fold evrec acc c
+ in
+ evrec Evar.Set.empty c
+
+let evars_of_named_context evd nc =
+ Context.Named.fold_outside
+ (NamedDecl.fold_constr (fun constr s -> Evar.Set.union s (evars_of_term evd constr)))
+ nc
+ ~init:Evar.Set.empty
+
+let evars_of_filtered_evar_info evd evi =
+ Evar.Set.union (evars_of_term evd evi.evar_concl)
+ (Evar.Set.union
+ (match evi.evar_body with
+ | Evar_empty -> Evar.Set.empty
+ | Evar_defined b -> evars_of_term evd b)
+ (evars_of_named_context evd (evar_filtered_context evi)))