diff options
author | Maxime Dénès <mail@maximedenes.fr> | 2017-03-21 14:29:56 +0100 |
---|---|---|
committer | Maxime Dénès <mail@maximedenes.fr> | 2017-03-21 17:15:27 +0100 |
commit | 09c856a3bcb7369244202ed94db247f7abe84dad (patch) | |
tree | c42d9900700d068d25d5707f5f395bd7805ca0ee /kernel | |
parent | cd87eac3757d8925ff4ba7dee85efadb195153a3 (diff) |
Do not typecheck twice the type of opaque constants.
I believe an unwanted shadowing was introduced by a4043608f704f0.
Diffstat (limited to 'kernel')
-rw-r--r-- | kernel/term_typing.ml | 41 |
1 files changed, 17 insertions, 24 deletions
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml index 749b5dbaf..2d5580347 100644 --- a/kernel/term_typing.ml +++ b/kernel/term_typing.ml @@ -22,24 +22,6 @@ open Entries open Typeops open Fast_typeops -let constrain_type env j poly subst = function - | `None -> - if not poly then (* Old-style polymorphism *) - make_polymorphic_if_constant_for_ind env j - else RegularArity (Vars.subst_univs_level_constr subst j.uj_type) - | `Some t -> - let tj = infer_type env t in - let _ = judge_of_cast env j DEFAULTcast tj in - assert (eq_constr t tj.utj_val); - RegularArity (Vars.subst_univs_level_constr subst t) - | `SomeWJ (t, tj) -> - let tj = infer_type env t in - let _ = judge_of_cast env j DEFAULTcast tj in - assert (eq_constr t tj.utj_val); - RegularArity (Vars.subst_univs_level_constr subst t) - -let map_option_typ = function None -> `None | Some x -> `Some x - (* Insertion of constants and parameters in environment. *) let mk_pure_proof c = (c, Univ.ContextSet.empty), [] @@ -193,15 +175,16 @@ let infer_declaration ~trust env kn dcl = let body, uctx, signatures = inline_side_effects env body uctx side_eff in let valid_signatures = check_signatures trust signatures in - let env' = push_context_set uctx env in + let env = push_context_set uctx env in let j = - let body,env',ectx = skip_trusted_seff valid_signatures body env' in - let j = infer env' body in + let body,env,ectx = skip_trusted_seff valid_signatures body env in + let j = infer env body in unzip ectx j in let j = hcons_j j in let subst = Univ.LMap.empty in - let _typ = constrain_type env' j c.const_entry_polymorphic subst - (`SomeWJ (typ,tyj)) in + let _ = judge_of_cast env j DEFAULTcast tyj in + assert (eq_constr typ tyj.utj_val); + let _typ = RegularArity (Vars.subst_univs_level_constr subst typ) in feedback_completion_typecheck feedback_id; j.uj_val, uctx) in let def = OpaqueDef (Opaqueproof.create proofterm) in @@ -221,7 +204,17 @@ let infer_declaration ~trust env kn dcl = let usubst, univs = Univ.abstract_universes abstract (Univ.ContextSet.to_context ctx) in let j = infer env body in - let typ = constrain_type env j c.const_entry_polymorphic usubst (map_option_typ typ) in + let typ = match typ with + | None -> + if not c.const_entry_polymorphic then (* Old-style polymorphism *) + make_polymorphic_if_constant_for_ind env j + else RegularArity (Vars.subst_univs_level_constr usubst j.uj_type) + | Some t -> + let tj = infer_type env t in + let _ = judge_of_cast env j DEFAULTcast tj in + assert (eq_constr t tj.utj_val); + RegularArity (Vars.subst_univs_level_constr usubst t) + in let def = hcons_constr (Vars.subst_univs_level_constr usubst j.uj_val) in let def = if opaque then OpaqueDef (Opaqueproof.create (Future.from_val (def, Univ.ContextSet.empty))) |