aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--plugins/funind/glob_term_to_relation.ml13
-rw-r--r--plugins/funind/glob_termops.ml45
-rw-r--r--plugins/funind/glob_termops.mli7
3 files changed, 60 insertions, 5 deletions
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 68e097fe9..8c0b28ae8 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -1288,17 +1288,20 @@ let do_build_inductive
let t = EConstr.Unsafe.to_constr t in
evd,
Environ.push_named (LocalAssum (id,t))
- (* try *)
- (* Typing.e_type_of env evd (mkConstU c) *)
- (* with Not_found -> *)
- (* raise (UserError("do_build_inductive", str "Cannot handle partial fixpoint")) *)
env
)
funnames
(Array.of_list funconstants)
(evd,Global.env ())
in
- let resa = Array.map (build_entry_lc env funnames_as_set []) rta in
+ (* we solve and replace the implicits *)
+ let rta =
+ Array.mapi (fun i rt ->
+ let _,t = Typing.type_of env evd (EConstr.of_constr (mkConstU ((Array.of_list funconstants).(i)))) in
+ resolve_and_replace_implicits ~expected_type:(Pretyping.OfType t) env evd rt
+ ) rta
+ in
+ let resa = Array.map (build_entry_lc env funnames_as_set []) rta in
let env_with_graphs =
let rel_arity i funargs = (* Rebuilding arities (with parameters) *)
let rel_first_args :(Name.t * Glob_term.glob_constr * Glob_term.glob_constr option ) list =
diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml
index 0361e8cb1..416e196e8 100644
--- a/plugins/funind/glob_termops.ml
+++ b/plugins/funind/glob_termops.ml
@@ -707,3 +707,48 @@ let expand_as =
(loc,(idl,cpl, expand_as (List.fold_left add_as map cpl) rt))
in
expand_as Id.Map.empty
+
+
+
+
+(* [resolve_and_replace_implicits ?expected_type env sigma rt] solves implicits of [rt] w.r.t. [env] and [sigma] and then replace them by their solution
+ *)
+
+exception Found of Evd.evar_info
+let resolve_and_replace_implicits ?(flags=Pretyping.all_and_fail_flags) ?(expected_type=Pretyping.WithoutTypeConstraint) env sigma rt =
+ let open Evd in
+ let open Evar_kinds in
+ (* we first (pseudo) understand [rt] and get back the computed evar_map *)
+ (* FIXME : JF (30/03/2017) I'm not completely sure to have split understand as needed.
+If someone knows how to prevent solved existantial removal in understand, please do not hesitate to change the computation of [ctx] here *)
+ let ctx,_ = Pretyping.ise_pretype_gen flags env sigma Pretyping.empty_lvar expected_type rt in
+ let ctx, f = Evarutil.nf_evars_and_universes ctx in
+
+ (* then we map [rt] to replace the implicit holes by their values *)
+ let rec change rt =
+ match rt.CAst.v with
+ | GHole(ImplicitArg(grk,pk,bk),_,_) -> (* we only want to deal with implicit arguments *)
+ (
+ try (* we scan the new evar map to find the evar corresponding to this hole (by looking the source *)
+ Evd.fold (* to simulate an iter *)
+ (fun _ evi _ ->
+ match evi.evar_source with
+ | (loc_evi,ImplicitArg(gr_evi,p_evi,b_evi)) ->
+ if Globnames.eq_gr grk gr_evi && pk=p_evi && bk=b_evi && rt.CAst.loc = loc_evi
+ then raise (Found evi)
+ | _ -> ()
+ )
+ ctx
+ ();
+ (* the hole was not solved : we do nothing *)
+ rt
+ with Found evi -> (* we found the evar corresponding to this hole *)
+ match evi.evar_body with
+ | Evar_defined c ->
+ (* we just have to lift the solution in glob_term *)
+ Detyping.detype false [] env ctx (EConstr.of_constr (f c))
+ | Evar_empty -> rt (* the hole was not solved : we do nothing *)
+ )
+ | _ -> Glob_ops.map_glob_constr change rt
+ in
+ change rt
diff --git a/plugins/funind/glob_termops.mli b/plugins/funind/glob_termops.mli
index 25d79582f..99a258de9 100644
--- a/plugins/funind/glob_termops.mli
+++ b/plugins/funind/glob_termops.mli
@@ -119,3 +119,10 @@ val zeta_normalize : Glob_term.glob_constr -> Glob_term.glob_constr
val expand_as : glob_constr -> glob_constr
+
+
+(* [resolve_and_replace_implicits ?expected_type env sigma rt] solves implicits of [rt] w.r.t. [env] and [sigma] and then replace them by their solution
+ *)
+val resolve_and_replace_implicits :
+ ?flags:Pretyping.inference_flags ->
+ ?expected_type:Pretyping.typing_constraint -> Environ.env -> Evd.evar_map -> glob_constr -> glob_constr