aboutsummaryrefslogtreecommitdiffhomepage
path: root/vernac/lemmas.ml
diff options
context:
space:
mode:
authorGravatar Matthieu Sozeau <mattam@mattam.org>2017-07-27 14:54:41 +0200
committerGravatar Matthieu Sozeau <mattam@mattam.org>2017-09-19 10:28:03 +0200
commitf72a67569ec8cb9160d161699302b67919da5686 (patch)
treea86642e048c3ac571829e6b1eb6f3d53a34d3db0 /vernac/lemmas.ml
parentfc587e75d2d5d6e67365a9bc3a13ba5c86aba87b (diff)
Allow declaring universe constraints at definition level.
Introduce a "+" modifier for universe and constraint declarations to indicate that these can be extended in the final definition/proof. By default [Definition f] is equivalent to [Definition f@{+|+}], i.e universes can be introduced and constraints as well. For [f@{}] or [f@{i j}], the constraints can be extended, no universe introduced, to maintain compatibility with existing developments. Use [f@{i j | }] to indicate that no constraint (nor universe) can be introduced. These kind of definitions could benefit from asynchronous processing. Declarations of universe binders and constraints also works for monomorphic definitions.
Diffstat (limited to 'vernac/lemmas.ml')
-rw-r--r--vernac/lemmas.ml50
1 files changed, 25 insertions, 25 deletions
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 645320c60..d78cd14f6 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -210,11 +210,11 @@ let compute_proof_name locality = function
locality == Global && Nametab.exists_cci (Lib.make_path_except_section id)
then
user_err ?loc (pr_id id ++ str " already exists.");
- id, pl
+ id
| None ->
- next_global_ident_away default_thm_id (Proof_global.get_all_proof_names ()), None
+ next_global_ident_away default_thm_id (Proof_global.get_all_proof_names ())
-let save_remaining_recthms (locality,p,kind) norm ctx body opaq i ((id,pl),(t_i,(_,imps))) =
+let save_remaining_recthms (locality,p,kind) norm ctx binders body opaq i (id,(t_i,(_,imps))) =
let t_i = norm t_i in
match body with
| None ->
@@ -222,7 +222,7 @@ let save_remaining_recthms (locality,p,kind) norm ctx body opaq i ((id,pl),(t_i,
| Discharge ->
let impl = false in (* copy values from Vernacentries *)
let k = IsAssumption Conjectural in
- let c = SectionLocalAssum ((t_i,ctx),p,impl) in
+ let c = SectionLocalAssum ((t_i,Univ.ContextSet.of_context ctx),p,impl) in
let _ = declare_variable id (Lib.cwd(),c,k) in
(Discharge, VarRef id,imps)
| Local | Global ->
@@ -232,7 +232,6 @@ let save_remaining_recthms (locality,p,kind) norm ctx body opaq i ((id,pl),(t_i,
| Global -> false
| Discharge -> assert false
in
- let ctx = Univ.ContextSet.to_context ctx in
let decl = (ParameterEntry (None,p,(t_i,ctx),None), k) in
let kn = declare_constant id ~local decl in
(locality,ConstRef kn,imps))
@@ -250,12 +249,11 @@ let save_remaining_recthms (locality,p,kind) norm ctx body opaq i ((id,pl),(t_i,
match locality with
| Discharge ->
let const = definition_entry ~types:t_i ~opaque:opaq ~poly:p
- ~univs:(Univ.ContextSet.to_context ctx) body_i in
+ ~univs:ctx body_i in
let c = SectionLocalDef const in
let _ = declare_variable id (Lib.cwd(), c, k) in
(Discharge,VarRef id,imps)
| Local | Global ->
- let ctx = Univ.ContextSet.to_context ctx in
let local = match locality with
| Local -> true
| Global -> false
@@ -369,7 +367,7 @@ let start_proof_univs id ?pl kind sigma ?terminator ?sign c ?init_tac ?(compute_
let rec_tac_initializer finite guard thms snl =
if finite then
- match List.map (fun ((id,_),(t,_)) -> (id,EConstr.of_constr t)) thms with
+ match List.map (fun (id,(t,_)) -> (id,EConstr.of_constr t)) thms with
| (id,_)::l -> Tactics.mutual_cofix id l 0
| _ -> assert false
else
@@ -377,11 +375,11 @@ let rec_tac_initializer finite guard thms snl =
let nl = match snl with
| None -> List.map succ (List.map List.last guard)
| Some nl -> nl
- in match List.map2 (fun ((id,_),(t,_)) n -> (id,n, EConstr.of_constr t)) thms nl with
+ in match List.map2 (fun (id,(t,_)) n -> (id,n, EConstr.of_constr t)) thms nl with
| (id,n,_)::l -> Tactics.mutual_fix id n l 0
| _ -> assert false
-let start_proof_with_initialization kind ctx recguard thms snl hook =
+let start_proof_with_initialization kind ctx decl recguard thms snl hook =
let intro_tac (_, (_, (ids, _))) =
Tacticals.New.tclMAP (function
| Name id -> Tactics.intro_mustbe_force id
@@ -406,7 +404,7 @@ let start_proof_with_initialization kind ctx recguard thms snl hook =
(if Flags.is_auto_intros () then Some (intro_tac (List.hd thms)) else None), [] in
match thms with
| [] -> anomaly (Pp.str "No proof to start.")
- | ((id,pl),(t,(_,imps)))::other_thms ->
+ | (id,(t,(_,imps)))::other_thms ->
let hook ctx strength ref =
let ctx = match ctx with
| None -> Evd.empty_evar_universe_context
@@ -418,22 +416,24 @@ let start_proof_with_initialization kind ctx recguard thms snl hook =
let body,opaq = retrieve_first_recthm ctx ref in
let subst = Evd.evar_universe_context_subst ctx in
let norm c = Universes.subst_opt_univs_constr subst c in
- let ctx = UState.context_set (*FIXME*) ctx in
+ let binders, ctx = Evd.check_univ_decl (Evd.from_ctx ctx) decl in
let body = Option.map norm body in
- List.map_i (save_remaining_recthms kind norm ctx body opaq) 1 other_thms in
+ List.map_i (save_remaining_recthms kind norm ctx binders body opaq) 1 other_thms in
let thms_data = (strength,ref,imps)::other_thms_data in
List.iter (fun (strength,ref,imps) ->
maybe_declare_manual_implicits false ref imps;
call_hook (fun exn -> exn) hook strength ref) thms_data in
- start_proof_univs id ?pl kind ctx (EConstr.of_constr t) ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard
+ start_proof_univs id ~pl:decl kind ctx (EConstr.of_constr t) ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard
let start_proof_com ?inference_hook kind thms hook =
let env0 = Global.env () in
- let levels = Option.map snd (fst (List.hd thms)) in
- let evdref = ref (match levels with
- | None -> Evd.from_env env0
- | Some l -> Evd.from_ctx (Evd.make_evar_universe_context env0 l))
- in
+ let decl = fst (List.hd thms) in
+ let evd, decl =
+ match decl with
+ | None -> Evd.from_env env0, Univdecls.default_univ_decl
+ | Some decl ->
+ Univdecls.interp_univ_decl_opt env0 (snd decl) in
+ let evdref = ref evd in
let thms = List.map (fun (sopt,(bl,t)) ->
let impls, ((env, ctx), imps) = interp_context_evars env0 evdref bl in
let t', imps' = interp_type_evars_impls ~impls env evdref t in
@@ -449,16 +449,16 @@ let start_proof_com ?inference_hook kind thms hook =
let evd, nf = Evarutil.nf_evars_and_universes !evdref in
let thms = List.map (fun (n, (t, info)) -> (n, (nf t, info))) thms in
let () =
- match levels with
- | None -> ()
- | Some l -> ignore (Evd.universe_context evd ?names:l)
+ if not decl.Misctypes.univdecl_extensible_instance then
+ ignore (Evd.universe_context evd ~names:decl.Misctypes.univdecl_instance)
+ else ()
in
let evd =
if pi2 kind then evd
else (* We fix the variables to ensure they won't be lowered to Set *)
Evd.fix_undefined_variables evd
in
- start_proof_with_initialization kind evd recguard thms snl hook
+ start_proof_with_initialization kind evd decl recguard thms snl hook
(* Saving a proof *)
@@ -507,9 +507,9 @@ let save_proof ?proof = function
let ids_def = Environ.global_vars_set env pproof in
Some (Environ.keep_hyps env (Idset.union ids_typ ids_def))
| _ -> None in
- let names = Proof_global.get_universe_binders () in
+ let decl = Proof_global.get_universe_decl () in
let evd = Evd.from_ctx universes in
- let binders, ctx = Evd.universe_context ?names evd in
+ let binders, ctx = Evd.check_univ_decl evd decl in
Admitted(id,k,(sec_vars, pi2 k, (typ, ctx), None),
(universes, Some binders))
in