aboutsummaryrefslogtreecommitdiffhomepage
path: root/engine/eConstr.ml
diff options
context:
space:
mode:
authorGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-12-04 15:54:28 +0100
committerGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2018-03-31 18:44:19 +0200
commit9f723f14e5342c1303646b5ea7bb5c0012a090ef (patch)
treed6a6a82ab8b73e975588a547eb15a5a2f83fd4c7 /engine/eConstr.ml
parent2d2d16430822f1768ce4f3c62ef0750b94e4747f (diff)
[econstr] Forbid calling `to_constr` in open terms.
We forbid calling `EConstr.to_constr` on terms that are not evar-free, as to progress towards enforcing the invariant that `Constr.t` is evar-free. [c.f. #6308] Due to compatibility constraints we provide an optional parameter to `to_constr`, `abort` which can be used to overcome this restriction until we fix all parts of the code. Now, grepping for `~abort:false` should return the questionable parts of the system. Not a lot of places had to be fixed, some comments: - problems with the interface due to `Evd/Constr` [`Evd.define` being the prime example] do seem real! - inductives also look bad with regards to `Constr/EConstr`. - code in plugins needs work. A notable user of this "feature" is `Obligations/Program` that seem to like to generate kernel-level entries with free evars, then to scan them and workaround this problem by generating constants.
Diffstat (limited to 'engine/eConstr.ml')
-rw-r--r--engine/eConstr.ml16
1 files changed, 11 insertions, 5 deletions
diff --git a/engine/eConstr.ml b/engine/eConstr.ml
index bd47a04f1..01b4fe851 100644
--- a/engine/eConstr.ml
+++ b/engine/eConstr.ml
@@ -40,7 +40,7 @@ val kind_of_type : Evd.evar_map -> t -> (t, t) Term.kind_of_type
val whd_evar : Evd.evar_map -> t -> t
val of_kind : (t, t, ESorts.t, EInstance.t) Constr.kind_of_term -> t
val of_constr : Constr.t -> t
-val to_constr : evar_map -> t -> Constr.t
+val to_constr : ?abort_on_undefined_evars:bool -> evar_map -> t -> Constr.t
val unsafe_to_constr : t -> Constr.t
val unsafe_eq : (t, Constr.t) eq
val of_named_decl : (Constr.t, Constr.types) Context.Named.Declaration.pt -> (t, t) Context.Named.Declaration.pt
@@ -110,11 +110,16 @@ let of_constr c = c
let unsafe_to_constr c = c
let unsafe_eq = Refl
-let rec to_constr sigma c = match Constr.kind c with
+let to_constr ?(abort_on_undefined_evars=true) sigma c =
+let rec to_constr c = match Constr.kind c with
| Evar ev ->
begin match safe_evar_value sigma ev with
- | Some c -> to_constr sigma c
- | None -> Constr.map (fun c -> to_constr sigma c) c
+ | Some c -> to_constr c
+ | None ->
+ if abort_on_undefined_evars then
+ anomaly ~label:"econstr" Pp.(str "grounding a non evar-free term")
+ else
+ Constr.map (fun c -> to_constr c) c
end
| Sort (Sorts.Type u) ->
let u' = Evd.normalize_universe sigma u in
@@ -128,7 +133,8 @@ let rec to_constr sigma c = match Constr.kind c with
| Construct (co, u) when not (Univ.Instance.is_empty u) ->
let u' = Evd.normalize_universe_instance sigma u in
if u' == u then c else mkConstructU (co, u')
-| _ -> Constr.map (fun c -> to_constr sigma c) c
+| _ -> Constr.map (fun c -> to_constr c) c
+in to_constr c
let of_named_decl d = d
let unsafe_to_named_decl d = d