From d437078a4ca82f7ca6d19be5ee9452359724f9a0 Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Fri, 15 Sep 2017 15:46:30 +0200 Subject: Use Maps and ids for universe binders Before sometimes there were lists and strings. --- vernac/classes.ml | 4 ++-- vernac/command.ml | 2 +- vernac/obligations.ml | 4 ++-- vernac/record.ml | 2 +- 4 files changed, 6 insertions(+), 6 deletions(-) (limited to 'vernac') diff --git a/vernac/classes.ml b/vernac/classes.ml index 20cb43b24..ed37bab6d 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -411,14 +411,14 @@ let context poly l = let decl = (Discharge, poly, Definitional) in let nstatus = match b with | None -> - pi3 (Command.declare_assumption false decl (t, !uctx) [] [] impl + pi3 (Command.declare_assumption false decl (t, !uctx) Universes.empty_binders [] impl Vernacexpr.NoInline (Loc.tag id)) | Some b -> let ctx = Univ.ContextSet.to_context !uctx in let decl = (Discharge, poly, Definition) in let entry = Declare.definition_entry ~poly ~univs:ctx ~types:t b in let hook = Lemmas.mk_hook (fun _ gr -> gr) in - let _ = DeclareDef.declare_definition id decl entry [] [] hook in + let _ = DeclareDef.declare_definition id decl entry Universes.empty_binders [] hook in Lib.sections_are_opened () || Lib.is_modtype_strict () in status && nstatus diff --git a/vernac/command.ml b/vernac/command.ml index 257c003b5..5d83de070 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -258,7 +258,7 @@ let do_assumptions_unbound_univs (_, poly, _ as kind) nl l = let l = List.map (on_pi2 nf_evar) l in pi2 (List.fold_left (fun (subst,status,ctx) ((is_coe,idl),t,imps) -> let t = replace_vars subst t in - let (refs,status') = declare_assumptions idl is_coe kind (t,ctx) [] imps false nl in + let (refs,status') = declare_assumptions idl is_coe kind (t,ctx) Universes.empty_binders imps false nl in let subst' = List.map2 (fun (_,id) (c,u) -> (id,Universes.constr_of_global_univ (c,u))) idl refs diff --git a/vernac/obligations.ml b/vernac/obligations.ml index a44de66e9..590332d08 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -484,7 +484,7 @@ let declare_definition prg = let () = progmap_remove prg in let cst = DeclareDef.declare_definition prg.prg_name - prg.prg_kind ce [] prg.prg_implicits + prg.prg_kind ce Universes.empty_binders prg.prg_implicits (Lemmas.mk_hook (fun l r -> Lemmas.call_hook fix_exn prg.prg_hook l r prg.prg_ctx; r)) in Universes.register_universe_binders cst pl; @@ -554,7 +554,7 @@ let declare_mutual_definition l = (* Declare the recursive definitions *) let ctx = Evd.evar_context_universe_context first.prg_ctx in let fix_exn = Hook.get get_fix_exn () in - let kns = List.map4 (DeclareDef.declare_fix ~opaque (local, poly, kind) [] ctx) + let kns = List.map4 (DeclareDef.declare_fix ~opaque (local, poly, kind) Universes.empty_binders ctx) fixnames fixdecls fixtypes fiximps in (* Declare notations *) List.iter (Metasyntax.add_notation_interpretation (Global.env())) first.prg_notations; diff --git a/vernac/record.ml b/vernac/record.ml index f09b57048..6ffbd1584 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -426,7 +426,7 @@ let declare_structure finite univs id idbuild paramimpls params arity template else mie in - let kn = Command.declare_mutual_inductive_with_eliminations mie [] [(paramimpls,[])] in + let kn = Command.declare_mutual_inductive_with_eliminations mie Universes.empty_binders [(paramimpls,[])] in let rsp = (kn,0) in (* This is ind path of idstruc *) let cstr = (rsp,1) in let fields = -- cgit v1.2.3 From a5feb9687819c5e7ef0db6e7b74d0e236a296674 Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Fri, 15 Sep 2017 21:01:57 +0200 Subject: Separate checking univ_decls and obtaining universe binder names. --- API/API.mli | 3 +- engine/evd.ml | 2 ++ engine/evd.mli | 7 ++-- engine/uState.ml | 8 +++-- engine/uState.mli | 6 ++-- plugins/funind/functional_principles_types.ml | 2 +- plugins/funind/recdef.ml | 2 +- plugins/ltac/rewrite.ml | 2 +- plugins/setoid_ring/newring.ml | 2 +- printing/printer.ml | 7 ++++ printing/printer.mli | 1 + proofs/proof_global.ml | 10 +++--- tactics/leminv.ml | 4 +-- vernac/class.ml | 2 +- vernac/classes.ml | 8 ++--- vernac/command.ml | 48 +++++++++++++++------------ vernac/indschemes.ml | 2 +- vernac/lemmas.ml | 12 +++---- vernac/obligations.ml | 15 ++++----- vernac/record.ml | 5 +-- vernac/vernacentries.ml | 6 ++-- 21 files changed, 87 insertions(+), 67 deletions(-) (limited to 'vernac') diff --git a/API/API.mli b/API/API.mli index 63c675221..889050d90 100644 --- a/API/API.mli +++ b/API/API.mli @@ -2882,7 +2882,8 @@ sig val evar_ident : evar -> evar_map -> Names.Id.t option val extract_all_conv_pbs : evar_map -> evar_map * evar_constraint list val universe_context : names:(Names.Id.t Loc.located) list -> extensible:bool -> evar_map -> - Universes.universe_binders * Univ.UContext.t + Univ.UContext.t + val universe_binders : evar_map -> Universes.universe_binders val nf_constraints : evar_map -> evar_map val from_ctx : UState.t -> evar_map diff --git a/engine/evd.ml b/engine/evd.ml index ca1196b94..67ae73ceb 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -936,6 +936,8 @@ let universe_of_name evd s = UState.universe_of_name evd.universes s let add_universe_name evd s l = { evd with universes = UState.add_universe_name evd.universes s l } +let universe_binders evd = UState.universe_binders evd.universes + let universes evd = UState.ugraph evd.universes let update_sigma_env evd env = diff --git a/engine/evd.mli b/engine/evd.mli index 36c399e98..b6157202d 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -509,7 +509,8 @@ val restrict_universe_context : evar_map -> Univ.LSet.t -> evar_map val universe_of_name : evar_map -> Id.t -> Univ.Level.t val add_universe_name : evar_map -> Id.t -> Univ.Level.t -> evar_map -val add_constraints_context : UState.t -> +val universe_binders : evar_map -> Universes.universe_binders +val add_constraints_context : UState.t -> Univ.constraints -> UState.t @@ -552,12 +553,12 @@ val check_leq : evar_map -> Univ.Universe.t -> Univ.Universe.t -> bool val evar_universe_context : evar_map -> UState.t val universe_context_set : evar_map -> Univ.ContextSet.t val universe_context : names:(Id.t located) list -> extensible:bool -> evar_map -> - Universes.universe_binders * Univ.UContext.t + Univ.UContext.t val universe_subst : evar_map -> Universes.universe_opt_subst val universes : evar_map -> UGraph.t val check_univ_decl : evar_map -> UState.universe_decl -> - Universes.universe_binders * Univ.UContext.t + Univ.UContext.t val merge_universe_context : evar_map -> UState.t -> evar_map val set_universe_context : evar_map -> UState.t -> evar_map diff --git a/engine/uState.ml b/engine/uState.ml index 498d73fd3..282acb3d6 100644 --- a/engine/uState.ml +++ b/engine/uState.ml @@ -114,6 +114,8 @@ let of_binders b = in { ctx with uctx_names = b, rmap } +let universe_binders ctx = fst ctx.uctx_names + let instantiate_variable l b v = try v := Univ.LMap.update l (Some b) !v with Not_found -> assert false @@ -291,7 +293,7 @@ let universe_context ~names ~extensible uctx = let inst = Univ.Instance.of_array inst in let ctx = Univ.UContext.make (inst, Univ.ContextSet.constraints uctx.uctx_local) in - fst uctx.uctx_names, ctx + ctx let check_implication uctx cstrs ctx = let gr = initial_graph uctx in @@ -303,14 +305,14 @@ let check_implication uctx cstrs ctx = let check_univ_decl uctx decl = let open Misctypes in - let pl, ctx = universe_context + let ctx = universe_context ~names:decl.univdecl_instance ~extensible:decl.univdecl_extensible_instance uctx in if not decl.univdecl_extensible_constraints then check_implication uctx decl.univdecl_constraints ctx; - pl, ctx + ctx let restrict ctx vars = let uctx' = Univops.restrict_universe_context ctx.uctx_local vars in diff --git a/engine/uState.mli b/engine/uState.mli index 7f2357a68..38af08f91 100644 --- a/engine/uState.mli +++ b/engine/uState.mli @@ -32,6 +32,8 @@ val of_context_set : Univ.ContextSet.t -> t val of_binders : Universes.universe_binders -> t +val universe_binders : t -> Universes.universe_binders + (** {5 Projections} *) val context_set : t -> Univ.ContextSet.t @@ -135,12 +137,12 @@ val normalize : t -> t Also return the association list of universe names and universes (including those not in [names]). *) val universe_context : names:(Id.t Loc.located) list -> extensible:bool -> t -> - Universes.universe_binders * Univ.UContext.t + Univ.UContext.t type universe_decl = (Names.Id.t Loc.located list, Univ.Constraint.t) Misctypes.gen_universe_decl -val check_univ_decl : t -> universe_decl -> Universes.universe_binders * Univ.UContext.t +val check_univ_decl : t -> universe_decl -> Univ.UContext.t (** {5 TODO: Document me} *) diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 996e2b6af..d5bff57af 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -348,7 +348,7 @@ let generate_functional_principle (evd: Evd.evar_map ref) let evd',value = change_property_sort evd' s new_principle_type new_princ_name in let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr value)) in (* Pp.msgnl (str "new principle := " ++ pr_lconstr value); *) - let univs = (snd (Evd.universe_context ~names:[] ~extensible:true evd')) in + let univs = Evd.universe_context ~names:[] ~extensible:true evd' in let ce = Declare.definition_entry ~poly:(Flags.is_universe_polymorphism ()) ~univs value in ignore( Declare.declare_constant diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 3089ec470..63896bb56 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -1556,7 +1556,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num let functional_id = add_suffix function_name "_F" in let term_id = add_suffix function_name "_terminate" in let functional_ref = - let ctx = (snd (Evd.universe_context ~names:[] ~extensible:true evm)) in + let ctx = Evd.universe_context ~names:[] ~extensible:true evm in declare_fun functional_id (IsDefinition Decl_kinds.Definition) ~ctx res in (* Refresh the global universes, now including those of _F *) diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 14b0742a7..2eb660219 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -1884,7 +1884,7 @@ let declare_projection n instance_id r = in it_mkProd_or_LetIn ccl ctx in let typ = it_mkProd_or_LetIn typ ctx in - let pl, ctx = Evd.universe_context ~names:[] ~extensible:true sigma in + let ctx = Evd.universe_context ~names:[] ~extensible:true sigma in let typ = EConstr.to_constr sigma typ in let term = EConstr.to_constr sigma term in let cst = diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 1c3bdb958..a37b35b3d 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -220,7 +220,7 @@ let exec_tactic env evd n f args = let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic_ist ist (ltac_call f (args@[getter]))) gl in let evd, nf = Evarutil.nf_evars_and_universes (Refiner.project gls) in let nf c = nf (constr_of c) in - Array.map nf !tactic_res, snd (Evd.universe_context ~names:[] ~extensible:true evd) + Array.map nf !tactic_res, Evd.universe_context ~names:[] ~extensible:true evd let stdlib_modules = [["Coq";"Setoids";"Setoid"]; diff --git a/printing/printer.ml b/printing/printer.ml index d7bb0460d..4d5cfcba3 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -256,6 +256,13 @@ let safe_pr_constr t = let (sigma, env) = Pfedit.get_current_context () in safe_pr_constr_env env sigma t +let pr_universe_ctx_set sigma c = + if !Detyping.print_universes && not (Univ.ContextSet.is_empty c) then + fnl()++pr_in_comment (fun c -> v 0 + (Univ.pr_universe_context_set (Termops.pr_evd_level sigma) c)) c + else + mt() + let pr_universe_ctx sigma c = if !Detyping.print_universes && not (Univ.UContext.is_empty c) then fnl()++pr_in_comment (fun c -> v 0 diff --git a/printing/printer.mli b/printing/printer.mli index e014baa2c..a3d0eaac2 100644 --- a/printing/printer.mli +++ b/printing/printer.mli @@ -121,6 +121,7 @@ val pr_polymorphic : bool -> Pp.t val pr_cumulative : bool -> bool -> Pp.t val pr_universe_instance : evar_map -> Univ.UContext.t -> Pp.t val pr_universe_ctx : evar_map -> Univ.UContext.t -> Pp.t +val pr_universe_ctx_set : evar_map -> Univ.ContextSet.t -> Pp.t val pr_cumulativity_info : evar_map -> Univ.CumulativityInfo.t -> Pp.t (** Printing global references using names as short as possible *) diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index fdc9a236b..9faff3211 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -331,8 +331,8 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now let initial_euctx = Proof.initial_euctx proof in let fpl, univs = Future.split2 fpl in let universes = if poly || now then Future.force univs else initial_euctx in - let binders, univctx = Evd.check_univ_decl (Evd.from_ctx universes) universe_decl in - let binders = if poly then Some binders else None in + let univctx = Evd.check_univ_decl (Evd.from_ctx universes) universe_decl in + let binders = if poly then Some (UState.universe_binders universes) else None in (* Because of dependent subgoals at the beginning of proofs, we could have existential variables in the initial types of goals, we need to normalise them for the kernel. *) @@ -360,7 +360,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now the body. So we keep the two sets distinct. *) let used_univs = Univ.LSet.union used_univs_body used_univs_typ in let ctx_body = UState.restrict ctx used_univs in - let _, univs = Evd.check_univ_decl (Evd.from_ctx ctx_body) universe_decl in + let univs = Evd.check_univ_decl (Evd.from_ctx ctx_body) universe_decl in (initunivs, typ), ((body, Univ.ContextSet.of_context univs), eff) else (* Since the proof is computed now, we can simply have 1 set of @@ -370,7 +370,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now TODO: check if restrict is really necessary now. *) let used_univs = Univ.LSet.union used_univs_body used_univs_typ in let ctx = UState.restrict universes used_univs in - let _, univs = Evd.check_univ_decl (Evd.from_ctx ctx) universe_decl in + let univs = Evd.check_univ_decl (Evd.from_ctx ctx) universe_decl in (univs, typ), ((body, Univ.ContextSet.empty), eff) in fun t p -> Future.split2 (Future.chain p (make_body t)) @@ -383,7 +383,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now the declaration as well. If the declaration is non-extensible, this will prevent the body from adding universes and constraints. *) let bodyunivs = constrain_variables univctx (Future.force univs) in - let _, univs = Evd.check_univ_decl (Evd.from_ctx bodyunivs) universe_decl in + let univs = Evd.check_univ_decl (Evd.from_ctx bodyunivs) universe_decl in (pt,Univ.ContextSet.of_context univs),eff) in let entry_fn p (_, t) = diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 62f3866de..4d1b271d6 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -235,9 +235,9 @@ let inversion_scheme env sigma t sort dep_option inv_op = p, Evd.universe_context ~names:[] ~extensible:true sigma let add_inversion_lemma name env sigma t sort dep inv_op = - let invProof, ctx = inversion_scheme env sigma t sort dep inv_op in + let invProof, univs = inversion_scheme env sigma t sort dep inv_op in let entry = definition_entry ~poly:(Flags.use_polymorphic_flag ()) - ~univs:(snd ctx) invProof in + ~univs invProof in let _ = declare_constant name (DefinitionEntry entry, IsProof Lemma) in () diff --git a/vernac/class.ml b/vernac/class.ml index 44f20a088..e59482b71 100644 --- a/vernac/class.ml +++ b/vernac/class.ml @@ -212,7 +212,7 @@ let build_id_coercion idf_opt source poly = Id.of_string ("Id_"^(ident_key_of_class source)^"_"^ (ident_key_of_class cl)) in - let univs = (snd (Evd.universe_context ~names:[] ~extensible:true sigma)) in + let univs = Evd.universe_context ~names:[] ~extensible:true sigma in let constr_entry = (* Cast is necessary to express [val_f] is identity *) DefinitionEntry (definition_entry ~types:typ_f ~poly ~univs diff --git a/vernac/classes.ml b/vernac/classes.ml index ed37bab6d..6f5f96ee2 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -119,14 +119,14 @@ let declare_instance_constant k info global imps ?hook id decl poly evm term ter (Univops.universes_of_constr term) in Evd.restrict_universe_context evm levels in - let pl, uctx = Evd.check_univ_decl evm decl in + let uctx = Evd.check_univ_decl evm decl in let entry = Declare.definition_entry ~types:termtype ~poly ~univs:uctx term in let cdecl = (DefinitionEntry entry, kind) in let kn = Declare.declare_constant id cdecl in Declare.definition_message id; - Universes.register_universe_binders (ConstRef kn) pl; + Universes.register_universe_binders (ConstRef kn) (Evd.universe_binders evm); instance_hook k info global imps ?hook (ConstRef kn); id @@ -203,12 +203,12 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) nf t in Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr termtype); - let pl, ctx = Evd.check_univ_decl !evars decl in + let ctx = Evd.check_univ_decl !evars decl in let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id (ParameterEntry (None,poly,(termtype,ctx),None), Decl_kinds.IsAssumption Decl_kinds.Logical) in - Universes.register_universe_binders (ConstRef cst) pl; + Universes.register_universe_binders (ConstRef cst) (Evd.universe_binders !evars); instance_hook k pri global imps ?hook (ConstRef cst); id end else ( diff --git a/vernac/command.ml b/vernac/command.ml index 5d83de070..5a063f173 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -107,8 +107,9 @@ let interp_definition pl bl p red_option c ctypopt = let body = nf (it_mkLambda_or_LetIn c ctx) in let vars = Univops.universes_of_constr body in let evd = Evd.restrict_universe_context !evdref vars in - let pl, uctx = Evd.check_univ_decl evd decl in - imps1@(Impargs.lift_implicits nb_args imps2), pl, + let uctx = Evd.check_univ_decl evd decl in + let binders = Evd.universe_binders evd in + imps1@(Impargs.lift_implicits nb_args imps2), binders, definition_entry ~univs:uctx ~poly:p body | Some ctyp -> let ty, impsty = interp_type_evars_impls ~impls env_bl evdref ctyp in @@ -133,8 +134,9 @@ let interp_definition pl bl p red_option c ctypopt = let vars = Univ.LSet.union (Univops.universes_of_constr body) (Univops.universes_of_constr typ) in let ctx = Evd.restrict_universe_context !evdref vars in - let pl, uctx = Evd.check_univ_decl ctx decl in - imps1@(Impargs.lift_implicits nb_args impsty), pl, + let uctx = Evd.check_univ_decl ctx decl in + let binders = Evd.universe_binders evd in + imps1@(Impargs.lift_implicits nb_args impsty), binders, definition_entry ~types:typ ~poly:p ~univs:uctx body in @@ -277,9 +279,10 @@ let do_assumptions_bound_univs coe kind nl id pl c = let ty = nf ty in let vars = Univops.universes_of_constr ty in let evd = Evd.restrict_universe_context !evdref vars in - let pl, uctx = Evd.check_univ_decl evd decl in + let uctx = Evd.check_univ_decl evd decl in + let binders = Evd.universe_binders evd in let uctx = Univ.ContextSet.of_context uctx in - let (_, _, st) = declare_assumption coe kind (ty, uctx) pl impls false nl id in + let (_, _, st) = declare_assumption coe kind (ty, uctx) binders impls false nl id in st let do_assumptions kind nl l = match l with @@ -576,7 +579,7 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite = let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in let ctx_params = Context.Rel.map nf ctx_params in let evd = !evdref in - let pl, uctx = Evd.check_univ_decl evd decl in + let uctx = Evd.check_univ_decl evd decl in List.iter (fun c -> check_evars env_params Evd.empty evd (EConstr.of_constr c)) arities; Context.Rel.iter (fun c -> check_evars env0 Evd.empty evd (EConstr.of_constr c)) ctx_params; List.iter (fun (_,ctyps,_) -> @@ -617,7 +620,7 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite = in (if poly && cum then Inductiveops.infer_inductive_subtyping env_ar evd mind_ent - else mind_ent), pl, impls + else mind_ent), Evd.universe_binders evd, impls (* Very syntactical equality *) let eq_local_binders bl1 bl2 = @@ -1025,17 +1028,18 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = if List.length binders_rel > 1 then let name = add_suffix recname "_func" in let hook l gr _ = - let body = it_mkLambda_or_LetIn (mkApp (Evarutil.e_new_global evdref gr, [|make|])) binders_rel in - let ty = it_mkProd_or_LetIn top_arity binders_rel in - let ty = EConstr.Unsafe.to_constr ty in - let pl, univs = Evd.universe_context ~names:pl ~extensible:plext !evdref in - (*FIXME poly? *) - let ce = definition_entry ~poly ~types:ty ~univs (EConstr.to_constr !evdref body) in - (** FIXME: include locality *) - let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in - let gr = ConstRef c in - if Impargs.is_implicit_args () || not (List.is_empty impls) then - Impargs.declare_manual_implicits false gr [impls] + let body = it_mkLambda_or_LetIn (mkApp (Evarutil.e_new_global evdref gr, [|make|])) binders_rel in + let ty = it_mkProd_or_LetIn top_arity binders_rel in + let ty = EConstr.Unsafe.to_constr ty in + let univs = Evd.universe_context ~names:pl ~extensible:plext !evdref in + (*FIXME poly? *) + let ce = definition_entry ~poly ~types:ty ~univs (EConstr.to_constr !evdref body) in + (** FIXME: include locality *) + let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in + let gr = ConstRef c in + let () = Universes.register_universe_binders gr (Evd.universe_binders !evdref) in + if Impargs.is_implicit_args () || not (List.is_empty impls) then + Impargs.declare_manual_implicits false gr [impls] in let typ = it_mkProd_or_LetIn top_arity binders in hook, name, typ @@ -1168,7 +1172,8 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in let evd = Evd.from_ctx ctx in let evd = Evd.restrict_universe_context evd vars in - let pl, ctx = Evd.check_univ_decl evd pl in + let ctx = Evd.check_univ_decl evd pl in + let pl = Evd.universe_binders evd in let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in ignore (List.map4 (DeclareDef.declare_fix (local, poly, Fixpoint) pl ctx) fixnames fixdecls fixtypes fiximps); @@ -1200,7 +1205,8 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in let evd = Evd.from_ctx ctx in let evd = Evd.restrict_universe_context evd vars in - let pl, ctx = Evd.check_univ_decl evd pl in + let ctx = Evd.check_univ_decl evd pl in + let pl = Evd.universe_binders evd in ignore (List.map4 (DeclareDef.declare_fix (local, poly, CoFixpoint) pl ctx) fixnames fixdecls fixtypes fiximps); (* Declare the recursive definitions *) diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml index c0ddc7e2c..58f39e303 100644 --- a/vernac/indschemes.ml +++ b/vernac/indschemes.ml @@ -109,7 +109,7 @@ let _ = let define id internal ctx c t = let f = declare_constant ~internal in - let _, univs = Evd.universe_context ~names:[] ~extensible:true ctx in + let univs = Evd.universe_context ~names:[] ~extensible:true ctx in let univs = if Flags.is_universe_polymorphism () then Polymorphic_const_entry univs else Monomorphic_const_entry univs diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index a025bfff8..cd30c7a68 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -49,7 +49,7 @@ let retrieve_first_recthm uctx = function (NamedDecl.get_value (Global.lookup_named id),variable_opacity id) | ConstRef cst -> let cb = Global.lookup_constant cst in - let (_, uctx) = UState.universe_context ~names:[] ~extensible:true uctx in + let uctx = UState.universe_context ~names:[] ~extensible:true uctx in let inst = Univ.UContext.instance uctx in let map (c, ctx) = Vars.subst_instance_constr inst c in (Option.map map (Global.body_of_constant_body cb), is_opaque cb) @@ -223,7 +223,7 @@ let compute_proof_name locality = function let avoid = Id.Set.of_list (Proof_global.get_all_proof_names ()) in next_global_ident_away default_thm_id avoid -let save_remaining_recthms (locality,p,kind) norm ctx binders body opaq i (id,(t_i,(_,imps))) = +let save_remaining_recthms (locality,p,kind) norm ctx body opaq i (id,(t_i,(_,imps))) = let t_i = norm t_i in match body with | None -> @@ -420,9 +420,9 @@ let start_proof_with_initialization kind ctx decl 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 binders, ctx = Evd.check_univ_decl (Evd.from_ctx ctx) decl in + let 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 binders body opaq) 1 other_thms in + List.map_i (save_remaining_recthms kind norm ctx 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; @@ -513,9 +513,9 @@ let save_proof ?proof = function | _ -> None in let decl = Proof_global.get_universe_decl () in let evd = Evd.from_ctx universes in - let binders, ctx = Evd.check_univ_decl evd decl in + let ctx = Evd.check_univ_decl evd decl in let poly = pi2 k in - let binders = if poly then Some binders else None in + let binders = if poly then Some (UState.universe_binders universes) else None in Admitted(id,k,(sec_vars, poly, (typ, ctx), None), (universes, binders)) in diff --git a/vernac/obligations.ml b/vernac/obligations.ml index 590332d08..216801811 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -475,20 +475,17 @@ let declare_definition prg = (Evd.evar_universe_context_subst prg.prg_ctx) in let opaque = prg.prg_opaque in let fix_exn = Hook.get get_fix_exn () in - let pl, ctx = Evd.check_univ_decl (Evd.from_ctx prg.prg_ctx) prg.prg_univdecl in + let () = ignore(UState.check_univ_decl prg.prg_ctx prg.prg_univdecl) in let ce = definition_entry ~fix_exn ~opaque ~types:(nf typ) ~poly:(pi2 prg.prg_kind) ~univs:(Evd.evar_context_universe_context prg.prg_ctx) (nf body) in let () = progmap_remove prg in - let cst = - DeclareDef.declare_definition prg.prg_name - prg.prg_kind ce Universes.empty_binders prg.prg_implicits - (Lemmas.mk_hook (fun l r -> Lemmas.call_hook fix_exn prg.prg_hook l r prg.prg_ctx; r)) - in - Universes.register_universe_binders cst pl; - cst + let ubinders = UState.universe_binders prg.prg_ctx in + DeclareDef.declare_definition prg.prg_name + prg.prg_kind ce ubinders prg.prg_implicits + (Lemmas.mk_hook (fun l r -> Lemmas.call_hook fix_exn prg.prg_hook l r prg.prg_ctx; r)) let rec lam_index n t acc = match Constr.kind t with @@ -893,7 +890,7 @@ in let ctx' = Evd.merge_universe_subst evd (Evd.universe_subst (Evd.from_ctx ctx')) in Univ.Instance.empty, Evd.evar_universe_context ctx' else - let (_, uctx) = UState.universe_context ~names:[] ~extensible:true ctx' in + let uctx = UState.universe_context ~names:[] ~extensible:true ctx' in Univ.UContext.instance uctx, ctx' in let obl = { obl with obl_body = Some (DefinedObl (cst, inst)) } in diff --git a/vernac/record.ml b/vernac/record.ml index 6ffbd1584..9fef8f682 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -168,9 +168,10 @@ let typecheck_params_and_fields finite def id pl t ps nots fs = let typ = EConstr.to_constr evars typ in let ce t = Pretyping.check_evars env0 Evd.empty evars (EConstr.of_constr t) in let univs = Evd.check_univ_decl evars decl in + let ubinders = Evd.universe_binders evars in List.iter (iter_constr ce) (List.rev newps); List.iter (iter_constr ce) (List.rev newfs); - univs, typ, template, imps, newps, impls, newfs + ubinders, univs, typ, template, imps, newps, impls, newfs let degenerate_decl decl = let id = match RelDecl.get_name decl with @@ -605,7 +606,7 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf if isnot_class && List.exists (fun opt -> not (Option.is_empty opt)) priorities then user_err Pp.(str "Priorities only allowed for type class substructures"); (* Now, younger decl in params and fields is on top *) - let (pl, ctx), arity, template, implpars, params, implfs, fields = + let pl, ctx, arity, template, implpars, params, implfs, fields = States.with_state_protection (fun () -> typecheck_params_and_fields finite (kind = Class true) idstruc pl s ps notations fs) () in let sign = structure_signature (fields@params) in diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 62c7edb19..e2f28a371 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -1555,8 +1555,8 @@ let vernac_check_may_eval ?loc redexp glopt rc = let sigma' = Evarconv.solve_unif_constraints_with_heuristics env sigma' in Evarconv.check_problems_are_solved env sigma'; let sigma',nf = Evarutil.nf_evars_and_universes sigma' in - let pl, uctx = Evd.universe_context ~names:[] ~extensible:true sigma' in - let env = Environ.push_context uctx (Evarutil.nf_env_evar sigma' env) in + let uctx = Evd.universe_context_set sigma' in + let env = Environ.push_context_set uctx (Evarutil.nf_env_evar sigma' env) in let c = nf c in let j = if Evarutil.has_undefined_evars sigma' (EConstr.of_constr c) then @@ -1572,7 +1572,7 @@ let vernac_check_may_eval ?loc redexp glopt rc = let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' j.Environ.uj_type } in Feedback.msg_notice (print_judgment env sigma' j ++ pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++ - Printer.pr_universe_ctx sigma uctx) + Printer.pr_universe_ctx_set sigma uctx) | Some r -> let (sigma',r_interp) = Hook.get f_interp_redexp env sigma' r in let redfun env evm c = -- cgit v1.2.3 From 60770e86f4ec925fce52ad3565a92beb98d253c1 Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Fri, 15 Sep 2017 22:21:46 +0200 Subject: Stop exposing UState.universe_context and its Evd wrapper. We can enforce properties through check_univ_decl, or get an arbitrary ordered context with UState.context / Evd.to_universe_context (the later being a new wrapper of the former). --- API/API.mli | 4 ++-- engine/evd.ml | 3 +-- engine/evd.mli | 9 +++++++-- engine/uState.mli | 23 +++++++++-------------- interp/modintern.ml | 2 +- plugins/funind/functional_principles_types.ml | 2 +- plugins/funind/recdef.ml | 2 +- plugins/ltac/rewrite.ml | 2 +- plugins/setoid_ring/newring.ml | 2 +- proofs/proof_global.ml | 2 +- tactics/ind_tables.ml | 2 +- tactics/leminv.ml | 2 +- test-suite/success/polymorphism.v | 2 ++ vernac/class.ml | 2 +- vernac/command.ml | 4 +--- vernac/indschemes.ml | 2 +- vernac/lemmas.ml | 15 ++++++++------- vernac/obligations.ml | 13 +++++++------ 18 files changed, 47 insertions(+), 46 deletions(-) (limited to 'vernac') diff --git a/API/API.mli b/API/API.mli index 889050d90..5c7860671 100644 --- a/API/API.mli +++ b/API/API.mli @@ -2881,12 +2881,12 @@ sig val universe_context_set : evar_map -> Univ.ContextSet.t val evar_ident : evar -> evar_map -> Names.Id.t option val extract_all_conv_pbs : evar_map -> evar_map * evar_constraint list - val universe_context : names:(Names.Id.t Loc.located) list -> extensible:bool -> evar_map -> - Univ.UContext.t val universe_binders : evar_map -> Universes.universe_binders val nf_constraints : evar_map -> evar_map val from_ctx : UState.t -> evar_map + val to_universe_context : evar_map -> Univ.UContext.t + val meta_list : evar_map -> (Constr.metavariable * clbinding) list val meta_defined : evar_map -> Constr.metavariable -> bool diff --git a/engine/evd.ml b/engine/evd.ml index 67ae73ceb..eb0f2e3e7 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -756,8 +756,7 @@ let evar_universe_context d = d.universes let universe_context_set d = UState.context_set d.universes -let universe_context ~names ~extensible evd = - UState.universe_context ~names ~extensible evd.universes +let to_universe_context evd = UState.context evd.universes let check_univ_decl evd decl = UState.check_univ_decl evd.universes decl diff --git a/engine/evd.mli b/engine/evd.mli index b6157202d..49b6a7583 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -492,6 +492,8 @@ type 'a in_evar_universe_context = 'a * UState.t val evar_universe_context_set : UState.t -> Univ.ContextSet.t val evar_universe_context_constraints : UState.t -> Univ.constraints val evar_context_universe_context : UState.t -> Univ.UContext.t +[@@ocaml.deprecated "alias of UState.context"] + val evar_universe_context_of : Univ.ContextSet.t -> UState.t val empty_evar_universe_context : UState.t val union_evar_universe_context : UState.t -> UState.t -> @@ -552,11 +554,14 @@ val check_leq : evar_map -> Univ.Universe.t -> Univ.Universe.t -> bool val evar_universe_context : evar_map -> UState.t val universe_context_set : evar_map -> Univ.ContextSet.t -val universe_context : names:(Id.t located) list -> extensible:bool -> evar_map -> - Univ.UContext.t val universe_subst : evar_map -> Universes.universe_opt_subst val universes : evar_map -> UGraph.t +(** [to_universe_context evm] extracts the local universes and + constraints of [evm] and orders the universes the same as + [Univ.ContextSet.to_context]. *) +val to_universe_context : evar_map -> Univ.UContext.t + val check_univ_decl : evar_map -> UState.universe_decl -> Univ.UContext.t diff --git a/engine/uState.mli b/engine/uState.mli index 38af08f91..5279c6388 100644 --- a/engine/uState.mli +++ b/engine/uState.mli @@ -125,23 +125,18 @@ val refresh_undefined_univ_variables : t -> t * Univ.universe_level_subst val normalize : t -> t -(** [universe_context names extensible ctx] - - Return a universe context containing the local universes of [ctx] - and their constraints. The universes corresponding to [names] come - first in the order defined by that list. - - If [extensible] is false, check that the universes of [names] are - the only local universes. - - Also return the association list of universe names and universes - (including those not in [names]). *) -val universe_context : names:(Id.t Loc.located) list -> extensible:bool -> t -> - Univ.UContext.t - type universe_decl = (Names.Id.t Loc.located list, Univ.Constraint.t) Misctypes.gen_universe_decl +(** [check_univ_decl ctx decl] + + If non extensible in [decl], check that the local universes (resp. + universe constraints) in [ctx] are implied by [decl]. + + Return a universe context containing the local universes of [ctx] + and their constraints. The universes corresponding to + [decl.univdecl_instance] come first in the order defined by that + list. *) val check_univ_decl : t -> universe_decl -> Univ.UContext.t (** {5 TODO: Document me} *) diff --git a/interp/modintern.ml b/interp/modintern.ml index 08657936e..3eb91d8cd 100644 --- a/interp/modintern.ml +++ b/interp/modintern.ml @@ -62,7 +62,7 @@ let transl_with_decl env = function WithMod (fqid,lookup_module qid) | CWith_Definition ((_,fqid),c) -> let c, ectx = interp_constr env (Evd.from_env env) c in - let ctx = Evd.evar_context_universe_context ectx in + let ctx = UState.context ectx in WithDef (fqid,(c,ctx)) let loc_of_module l = l.CAst.loc diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index d5bff57af..076fba861 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -348,7 +348,7 @@ let generate_functional_principle (evd: Evd.evar_map ref) let evd',value = change_property_sort evd' s new_principle_type new_princ_name in let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr value)) in (* Pp.msgnl (str "new principle := " ++ pr_lconstr value); *) - let univs = Evd.universe_context ~names:[] ~extensible:true evd' in + let univs = Evd.to_universe_context evd' in let ce = Declare.definition_entry ~poly:(Flags.is_universe_polymorphism ()) ~univs value in ignore( Declare.declare_constant diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 63896bb56..320f7c7f3 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -1556,7 +1556,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num let functional_id = add_suffix function_name "_F" in let term_id = add_suffix function_name "_terminate" in let functional_ref = - let ctx = Evd.universe_context ~names:[] ~extensible:true evm in + let ctx = Evd.to_universe_context evm in declare_fun functional_id (IsDefinition Decl_kinds.Definition) ~ctx res in (* Refresh the global universes, now including those of _F *) diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 2eb660219..f6523d28c 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -1884,7 +1884,7 @@ let declare_projection n instance_id r = in it_mkProd_or_LetIn ccl ctx in let typ = it_mkProd_or_LetIn typ ctx in - let ctx = Evd.universe_context ~names:[] ~extensible:true sigma in + let ctx = Evd.to_universe_context sigma in let typ = EConstr.to_constr sigma typ in let term = EConstr.to_constr sigma term in let cst = diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index a37b35b3d..6d0b6fd09 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -220,7 +220,7 @@ let exec_tactic env evd n f args = let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic_ist ist (ltac_call f (args@[getter]))) gl in let evd, nf = Evarutil.nf_evars_and_universes (Refiner.project gls) in let nf c = nf (constr_of c) in - Array.map nf !tactic_res, Evd.universe_context ~names:[] ~extensible:true evd + Array.map nf !tactic_res, Evd.to_universe_context evd let stdlib_modules = [["Coq";"Setoids";"Setoid"]; diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index 9faff3211..ea0408169 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -353,7 +353,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now let used_univs_typ = Univops.universes_of_constr typ in if keep_body_ucst_separate || not (Safe_typing.empty_private_constants = eff) then - let initunivs = Evd.evar_context_universe_context initial_euctx in + let initunivs = UState.context initial_euctx in let ctx = constrain_variables initunivs universes in (* For vi2vo compilation proofs are computed now but we need to complement the univ constraints of the typ with the ones of diff --git a/tactics/ind_tables.ml b/tactics/ind_tables.ml index e7fa555c2..92c326b1e 100644 --- a/tactics/ind_tables.ml +++ b/tactics/ind_tables.ml @@ -123,7 +123,7 @@ let define internal id c p univs = let ctx = Evd.normalize_evar_universe_context univs in let c = Vars.subst_univs_fn_constr (Universes.make_opt_subst (Evd.evar_universe_context_subst ctx)) c in - let univs = Evd.evar_context_universe_context ctx in + let univs = UState.context ctx in let univs = if p then Polymorphic_const_entry univs else Monomorphic_const_entry univs diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 4d1b271d6..1ee873e0f 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -232,7 +232,7 @@ let inversion_scheme env sigma t sort dep_option inv_op = let invProof = it_mkNamedLambda_or_LetIn c !ownSign in let invProof = EConstr.Unsafe.to_constr invProof in let p = Evarutil.nf_evars_universes sigma invProof in - p, Evd.universe_context ~names:[] ~extensible:true sigma + p, Evd.to_universe_context sigma let add_inversion_lemma name env sigma t sort dep inv_op = let invProof, univs = inversion_scheme env sigma t sort dep inv_op in diff --git a/test-suite/success/polymorphism.v b/test-suite/success/polymorphism.v index 7eaafc354..00cab744e 100644 --- a/test-suite/success/polymorphism.v +++ b/test-suite/success/polymorphism.v @@ -190,6 +190,8 @@ Module binders. Fail Defined. Abort. + Fail Lemma bar@{u v | } : let x := (fun x => x) : Type@{u} -> Type@{v} in nat. + Lemma bar@{i j| i < j} : Type@{j}. Proof. exact Type@{i}. diff --git a/vernac/class.ml b/vernac/class.ml index e59482b71..68fd22cb6 100644 --- a/vernac/class.ml +++ b/vernac/class.ml @@ -212,7 +212,7 @@ let build_id_coercion idf_opt source poly = Id.of_string ("Id_"^(ident_key_of_class source)^"_"^ (ident_key_of_class cl)) in - let univs = Evd.universe_context ~names:[] ~extensible:true sigma in + let univs = Evd.to_universe_context sigma in let constr_entry = (* Cast is necessary to express [val_f] is identity *) DefinitionEntry (definition_entry ~types:typ_f ~poly ~univs diff --git a/vernac/command.ml b/vernac/command.ml index 5a063f173..80599c534 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -1022,8 +1022,6 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let binders_rel = nf_evar_context !evdref binders_rel in let binders = nf_evar_context !evdref binders in let top_arity = Evarutil.nf_evar !evdref top_arity in - let pl, plext = Option.cata - (fun d -> d.univdecl_instance, d.univdecl_extensible_instance) ([],true) pl in let hook, recname, typ = if List.length binders_rel > 1 then let name = add_suffix recname "_func" in @@ -1031,7 +1029,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let body = it_mkLambda_or_LetIn (mkApp (Evarutil.e_new_global evdref gr, [|make|])) binders_rel in let ty = it_mkProd_or_LetIn top_arity binders_rel in let ty = EConstr.Unsafe.to_constr ty in - let univs = Evd.universe_context ~names:pl ~extensible:plext !evdref in + let univs = Evd.check_univ_decl !evdref decl in (*FIXME poly? *) let ce = definition_entry ~poly ~types:ty ~univs (EConstr.to_constr !evdref body) in (** FIXME: include locality *) diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml index 58f39e303..c728c2671 100644 --- a/vernac/indschemes.ml +++ b/vernac/indschemes.ml @@ -109,7 +109,7 @@ let _ = let define id internal ctx c t = let f = declare_constant ~internal in - let univs = Evd.universe_context ~names:[] ~extensible:true ctx in + let univs = Evd.to_universe_context ctx in let univs = if Flags.is_universe_polymorphism () then Polymorphic_const_entry univs else Monomorphic_const_entry univs diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index cd30c7a68..71d80c4db 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -49,7 +49,8 @@ let retrieve_first_recthm uctx = function (NamedDecl.get_value (Global.lookup_named id),variable_opacity id) | ConstRef cst -> let cb = Global.lookup_constant cst in - let uctx = UState.universe_context ~names:[] ~extensible:true uctx in + (* we get the right order somehow but surely it could be enforced in a better way *) + let uctx = UState.context uctx in let inst = Univ.UContext.instance uctx in let map (c, ctx) = Vars.subst_instance_constr inst c in (Option.map map (Global.body_of_constant_body cb), is_opaque cb) @@ -420,8 +421,8 @@ let start_proof_with_initialization kind ctx decl 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 = Evd.check_univ_decl (Evd.from_ctx ctx) decl in - let body = Option.map norm body in + let ctx = UState.check_univ_decl 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 let thms_data = (strength,ref,imps)::other_thms_data in List.iter (fun (strength,ref,imps) -> @@ -453,9 +454,9 @@ 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 () = - if not decl.Misctypes.univdecl_extensible_instance then - ignore (Evd.universe_context evd ~names:decl.Misctypes.univdecl_instance ~extensible:false) - else () + let open Misctypes in + if not (decl.univdecl_extensible_instance && decl.univdecl_extensible_constraints) then + ignore (Evd.check_univ_decl evd decl) in let evd = if pi2 kind then evd @@ -490,7 +491,7 @@ let save_proof ?proof = function if const_entry_type = None then user_err Pp.(str "Admitted requires an explicit statement"); let typ = Option.get const_entry_type in - let ctx = Evd.evar_context_universe_context (fst universes) in + let ctx = UState.context (fst universes) in let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in Admitted(id, k, (sec_vars, pi2 k, (typ, ctx), None), universes) | None -> diff --git a/vernac/obligations.ml b/vernac/obligations.ml index 216801811..43af8968f 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -479,7 +479,7 @@ let declare_definition prg = let ce = definition_entry ~fix_exn ~opaque ~types:(nf typ) ~poly:(pi2 prg.prg_kind) - ~univs:(Evd.evar_context_universe_context prg.prg_ctx) (nf body) + ~univs:(UState.context prg.prg_ctx) (nf body) in let () = progmap_remove prg in let ubinders = UState.universe_binders prg.prg_ctx in @@ -549,7 +549,7 @@ let declare_mutual_definition l = mk_proof (mkCoFix (i,fixdecls))) 0 l in (* Declare the recursive definitions *) - let ctx = Evd.evar_context_universe_context first.prg_ctx in + let ctx = UState.context first.prg_ctx in let fix_exn = Hook.get get_fix_exn () in let kns = List.map4 (DeclareDef.declare_fix ~opaque (local, poly, kind) Universes.empty_binders ctx) fixnames fixdecls fixtypes fiximps in @@ -855,7 +855,7 @@ let obligation_terminator name num guard hook auto pf = | (_, status), Vernacexpr.Transparent -> status in let obl = { obl with obl_status = false, status } in - let uctx = Evd.evar_context_universe_context ctx in + let uctx = UState.context ctx in let (_, obl) = declare_obligation prg obl body ty uctx in let obls = Array.copy obls in let _ = obls.(num) <- obl in @@ -890,7 +890,8 @@ in let ctx' = Evd.merge_universe_subst evd (Evd.universe_subst (Evd.from_ctx ctx')) in Univ.Instance.empty, Evd.evar_universe_context ctx' else - let uctx = UState.universe_context ~names:[] ~extensible:true ctx' in + (* We get the right order somehow, but surely it could be enforced in a clearer way. *) + let uctx = UState.context ctx' in Univ.UContext.instance uctx, ctx' in let obl = { obl with obl_body = Some (DefinedObl (cst, inst)) } in @@ -969,7 +970,7 @@ and solve_obligation_by_tac prg obls i tac = solve_by_tac obl.obl_name (evar_of_obligation obl) tac (pi2 prg.prg_kind) (Evd.evar_universe_context evd) in - let uctx = Evd.evar_context_universe_context ctx in + let uctx = UState.context ctx in let prg = {prg with prg_ctx = ctx} in let def, obl' = declare_obligation prg obl t ty uctx in obls.(i) <- obl'; @@ -1120,7 +1121,7 @@ let admit_prog prg = match x.obl_body with | None -> let x = subst_deps_obl obls x in - let ctx = Evd.evar_context_universe_context prg.prg_ctx in + let ctx = UState.context prg.prg_ctx in let kn = Declare.declare_constant x.obl_name ~local:true (ParameterEntry (None,false,(x.obl_type,ctx),None), IsAssumption Conjectural) in -- cgit v1.2.3 From e414c07e193db7d4256c09167f6efd545831fa2b Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Fri, 15 Sep 2017 22:33:23 +0200 Subject: Register universe binders for record projections. --- vernac/record.ml | 35 +++++++++++++++++++---------------- vernac/record.mli | 24 ------------------------ 2 files changed, 19 insertions(+), 40 deletions(-) (limited to 'vernac') diff --git a/vernac/record.ml b/vernac/record.ml index 9fef8f682..e465e0616 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -266,7 +266,7 @@ let warn_non_primitive_record = strbrk" could not be defined as a primitive record"))) (* We build projections *) -let declare_projections indsp ?(kind=StructureComponent) binder_name coers fieldimpls fields = +let declare_projections indsp ?(kind=StructureComponent) binder_name coers ubinders fieldimpls fields = let env = Global.env() in let (mib,mip) = Global.lookup_inductive indsp in let poly = Declareops.inductive_is_polymorphic mib in @@ -305,9 +305,11 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field let kn, term = if is_local_assum decl && primitive then (** Already defined in the kernel silently *) - let kn = destConstRef (Nametab.locate (Libnames.qualid_of_ident fid)) in - Declare.definition_message fid; - kn, mkProj (Projection.make kn false,mkRel 1) + let gr = Nametab.locate (Libnames.qualid_of_ident fid) in + let kn = destConstRef gr in + Declare.definition_message fid; + Universes.register_universe_binders gr ubinders; + kn, mkProj (Projection.make kn false,mkRel 1) else let ccl = subst_projection fid subst ti in let body = match decl with @@ -344,8 +346,9 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field let constr_fip = let proj_args = (*Rel 1 refers to "x"*) paramargs@[mkRel 1] in applist (mkConstU (kn,u),proj_args) - in - Declare.definition_message fid; + in + Declare.definition_message fid; + Universes.register_universe_binders (ConstRef kn) ubinders; kn, constr_fip with Type_errors.TypeError (ctx,te) -> raise (NotDefinable (BadTypedProj (fid,ctx,te))) @@ -383,7 +386,7 @@ let structure_signature ctx = open Typeclasses -let declare_structure finite univs id idbuild paramimpls params arity template +let declare_structure finite ubinders univs id idbuild paramimpls params arity template fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers sign = let nparams = List.length params and nfields = List.length fields in let args = Context.Rel.to_extended_list mkRel nfields params in @@ -436,10 +439,11 @@ let declare_structure finite univs id idbuild paramimpls params arity template Context.Rel.map (fun c -> Vars.subst_univs_level_constr subst c) fields else fields in - let kinds,sp_projs = declare_projections rsp ~kind binder_name coers fieldimpls fields in + let kinds,sp_projs = declare_projections rsp ~kind binder_name coers ubinders fieldimpls fields in let build = ConstructRef cstr in let () = if is_coe then Class.try_add_new_coercion build ~local:false poly in Recordops.declare_structure(rsp,cstr,List.rev kinds,List.rev sp_projs); + Universes.register_universe_binders (IndRef rsp) ubinders; rsp let implicits_of_context ctx = @@ -451,7 +455,7 @@ let implicits_of_context ctx = in ExplByPos (i, explname), (true, true, true)) 1 (List.rev (Anonymous :: (List.map RelDecl.get_name ctx))) -let declare_class finite def cum poly ctx id idbuild paramimpls params arity +let declare_class finite def cum poly ubinders ctx id idbuild paramimpls params arity template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities sign = let fieldimpls = (* Make the class implicit in the projections, and the params if applicable. *) @@ -486,7 +490,9 @@ let declare_class finite def cum poly ctx id idbuild paramimpls params arity in let cref = ConstRef cst in Impargs.declare_manual_implicits false cref [paramimpls]; + Universes.register_universe_binders cref ubinders; Impargs.declare_manual_implicits false (ConstRef proj_cst) [List.hd fieldimpls]; + Universes.register_universe_binders (ConstRef proj_cst) ubinders; Classes.set_typeclass_transparency (EvalConstRef cst) false false; let sub = match List.hd coers with | Some b -> Some ((if b then Backward else Forward), List.hd priorities) @@ -503,7 +509,7 @@ let declare_class finite def cum poly ctx id idbuild paramimpls params arity else Monomorphic_ind_entry ctx in - let ind = declare_structure BiFinite univs (snd id) idbuild paramimpls + let ind = declare_structure BiFinite ubinders univs (snd id) idbuild paramimpls params arity template fieldimpls fields ~kind:Method ~name:binder_name false (List.map (fun _ -> false) fields) sign in @@ -610,10 +616,10 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf States.with_state_protection (fun () -> typecheck_params_and_fields finite (kind = Class true) idstruc pl s ps notations fs) () in let sign = structure_signature (fields@params) in - let gr = match kind with + match kind with | Class def -> let priorities = List.map (fun id -> {hint_priority = id; hint_pattern = None}) priorities in - let gr = declare_class finite def cum poly ctx (loc,idstruc) idbuild + let gr = declare_class finite def cum poly pl ctx (loc,idstruc) idbuild implpars params arity template implfs fields is_coe coers priorities sign in gr | _ -> @@ -630,10 +636,7 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf else Monomorphic_ind_entry ctx in - let ind = declare_structure finite univs idstruc + let ind = declare_structure finite pl univs idstruc idbuild implpars params arity template implfs fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) sign in IndRef ind - in - Universes.register_universe_binders gr pl; - gr diff --git a/vernac/record.mli b/vernac/record.mli index 33c2fba89..9fdd5e1c4 100644 --- a/vernac/record.mli +++ b/vernac/record.mli @@ -7,36 +7,12 @@ (************************************************************************) open Names -open Constr open Vernacexpr open Constrexpr -open Impargs open Globnames val primitive_flag : bool ref -(** [declare_projections ref name coers params fields] declare projections of - record [ref] (if allowed) using the given [name] as argument, and put them - as coercions accordingly to [coers]; it returns the absolute names of projections *) - -val declare_projections : - inductive -> ?kind:Decl_kinds.definition_object_kind -> Id.t -> - coercion_flag list -> manual_explicitation list list -> Context.Rel.t -> - (Name.t * bool) list * Constant.t option list - -val declare_structure : - Decl_kinds.recursivity_kind -> - Entries.inductive_universes -> - Id.t -> Id.t -> - manual_explicitation list -> Context.Rel.t -> (** params *) constr -> (** arity *) - bool (** template arity ? *) -> - Impargs.manual_explicitation list list -> Context.Rel.t -> (** fields *) - ?kind:Decl_kinds.definition_object_kind -> ?name:Id.t -> - bool -> (** coercion? *) - bool list -> (** field coercions *) - Evd.evar_map -> - inductive - val definition_structure : inductive_kind * Decl_kinds.cumulative_inductive_flag * Decl_kinds.polymorphic * Decl_kinds.recursivity_kind * ident_decl with_coercion * local_binder_expr list * -- cgit v1.2.3 From 4eb6d29d1ca7e0cc28d59d19a50adb83f7b30a2a Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Fri, 15 Sep 2017 23:39:11 +0200 Subject: Fix defining non primitive projections with abstracted universes. I think this only affects printing (in the new test you would get [Var (0)] when printing runwrap) but is still ugly. --- test-suite/output/UnivBinders.out | 51 ++++++++++++++++++++++++++++++++++++--- test-suite/output/UnivBinders.v | 26 ++++++++++++++++++-- vernac/record.ml | 16 +++--------- 3 files changed, 75 insertions(+), 18 deletions(-) (limited to 'vernac') diff --git a/test-suite/output/UnivBinders.out b/test-suite/output/UnivBinders.out index 904ff68aa..1bacb7513 100644 --- a/test-suite/output/UnivBinders.out +++ b/test-suite/output/UnivBinders.out @@ -1,12 +1,55 @@ +NonCumulative Inductive Empty@{u} : Type@{u} := +NonCumulative Record PWrap (A : Type@{u}) : Type@{u} := pwrap { punwrap : A } + +PWrap has primitive projections with eta conversion. +For PWrap: Argument scope is [type_scope] +For pwrap: Argument scopes are [type_scope _] +punwrap@{u} = +fun (A : Type@{u}) (p : PWrap@{u} A) => punwrap _ p + : forall A : Type@{u}, PWrap@{u} A -> A +(* u |= *) + +punwrap is universe polymorphic +Argument scopes are [type_scope _] +NonCumulative Record RWrap (A : Type@{u}) : Type@{u} := rwrap { runwrap : A } + +For RWrap: Argument scope is [type_scope] +For rwrap: Argument scopes are [type_scope _] +runwrap@{u} = +fun (A : Type@{u}) (r : RWrap@{u} A) => let (runwrap) := r in runwrap + : forall A : Type@{u}, RWrap@{u} A -> A +(* u |= *) + +runwrap is universe polymorphic +Argument scopes are [type_scope _] +Wrap@{u} = fun A : Type@{u} => A + : Type@{u} -> Type@{u} +(* u |= *) + +Wrap is universe polymorphic +Argument scope is [type_scope] +wrap@{u} = +fun (A : Type@{u}) (Wrap : Wrap@{u} A) => Wrap + : forall A : Type@{u}, Wrap@{u} A -> A +(* u |= *) + +wrap is universe polymorphic +Arguments A, Wrap are implicit and maximally inserted +Argument scopes are [type_scope _] bar@{u} = nat : Wrap@{u} Set (* u |= Set < u *) bar is universe polymorphic -foo@{u Top.8 v} = -Type@{Top.8} -> Type@{v} -> Type@{u} - : Type@{max(u+1, Top.8+1, v+1)} -(* u Top.8 v |= *) +foo@{u Top.17 v} = +Type@{Top.17} -> Type@{v} -> Type@{u} + : Type@{max(u+1, Top.17+1, v+1)} +(* u Top.17 v |= *) foo is universe polymorphic +Monomorphic mono = Type@{u} + : Type@{u+1} +(* u |= *) + +mono is not universe polymorphic diff --git a/test-suite/output/UnivBinders.v b/test-suite/output/UnivBinders.v index 8656ff1a3..2c378e795 100644 --- a/test-suite/output/UnivBinders.v +++ b/test-suite/output/UnivBinders.v @@ -2,12 +2,34 @@ Set Universe Polymorphism. Set Printing Universes. Unset Strict Universe Declaration. -Class Wrap A := wrap : A. +(* universe binders on inductive types and record projections *) +Inductive Empty@{u} : Type@{u} := . +Print Empty. -Instance bar@{u} : Wrap@{u} Set. Proof nat. +Set Primitive Projections. +Record PWrap@{u} (A:Type@{u}) := pwrap { punwrap : A }. +Print PWrap. +Print punwrap. + +Unset Primitive Projections. +Record RWrap@{u} (A:Type@{u}) := rwrap { runwrap : A }. +Print RWrap. +Print runwrap. + +(* universe binders also go on the constants for operational typeclasses. *) +Class Wrap@{u} (A:Type@{u}) := wrap : A. +Print Wrap. +Print wrap. + +(* Instance in lemma mode used to ignore the binders. *) +Instance bar@{u} : Wrap@{u} Set. Proof. exact nat. Qed. Print bar. (* The universes in the binder come first, then the extra universes in order of appearance. *) Definition foo@{u +} := Type -> Type@{v} -> Type@{u}. Print foo. + +(* Binders even work with monomorphic definitions! *) +Monomorphic Definition mono@{u} := Type@{u}. +Print mono. diff --git a/vernac/record.ml b/vernac/record.ml index e465e0616..0819dadb4 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -266,11 +266,10 @@ let warn_non_primitive_record = strbrk" could not be defined as a primitive record"))) (* We build projections *) -let declare_projections indsp ?(kind=StructureComponent) binder_name coers ubinders fieldimpls fields = +let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers ubinders fieldimpls fields = let env = Global.env() in let (mib,mip) = Global.lookup_inductive indsp in let poly = Declareops.inductive_is_polymorphic mib in - let ctx = Univ.AUContext.repr (Declareops.inductive_polymorphic_context mib) in let u = Univ.UContext.instance ctx in let paramdecls = Inductive.inductive_paramdecls (mib, u) in let indu = indsp, u in @@ -394,7 +393,7 @@ let declare_structure finite ubinders univs id idbuild paramimpls params arity t let type_constructor = it_mkProd_or_LetIn ind fields in let poly, ctx = match univs with - | Monomorphic_ind_entry ctx -> false, ctx + | Monomorphic_ind_entry ctx -> false, Univ.UContext.empty | Polymorphic_ind_entry ctx -> true, ctx | Cumulative_ind_entry cumi -> true, (Univ.CumulativityInfo.univ_context cumi) in @@ -430,20 +429,13 @@ let declare_structure finite ubinders univs id idbuild paramimpls params arity t else mie in - let kn = Command.declare_mutual_inductive_with_eliminations mie Universes.empty_binders [(paramimpls,[])] in + let kn = Command.declare_mutual_inductive_with_eliminations mie ubinders [(paramimpls,[])] in let rsp = (kn,0) in (* This is ind path of idstruc *) let cstr = (rsp,1) in - let fields = - if poly then - let subst, _ = Univ.abstract_universes ctx in - Context.Rel.map (fun c -> Vars.subst_univs_level_constr subst c) fields - else fields - in - let kinds,sp_projs = declare_projections rsp ~kind binder_name coers ubinders fieldimpls fields in + let kinds,sp_projs = declare_projections rsp ctx ~kind binder_name coers ubinders fieldimpls fields in let build = ConstructRef cstr in let () = if is_coe then Class.try_add_new_coercion build ~local:false poly in Recordops.declare_structure(rsp,cstr,List.rev kinds,List.rev sp_projs); - Universes.register_universe_binders (IndRef rsp) ubinders; rsp let implicits_of_context ctx = -- cgit v1.2.3 From 58c0784745f8b2ba7523f246c4611d780c9f3f70 Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Mon, 18 Sep 2017 14:50:07 +0200 Subject: When declaring constants/inductives use ContextSet if monomorphic. Also use constant_universes_entry instead of a bool flag to indicate polymorphism in ParameterEntry. There are a few places where we convert back to ContextSet because check_univ_decl returns a UContext, this could be improved. --- API/API.mli | 11 +++++----- checker/cic.mli | 4 ++-- checker/mod_checking.ml | 2 +- checker/values.ml | 6 ++--- interp/declare.ml | 6 +++-- kernel/declarations.ml | 6 ++--- kernel/declareops.ml | 4 ++-- kernel/entries.ml | 10 +++++---- kernel/indtypes.ml | 9 ++++---- kernel/mod_typing.ml | 26 ++++++++++++---------- kernel/safe_typing.ml | 12 +++++----- kernel/term_typing.ml | 46 ++++++++++++++++++++++----------------- kernel/term_typing.mli | 2 +- kernel/univ.ml | 1 + kernel/univ.mli | 5 ++++- kernel/vconv.ml | 2 +- plugins/ltac/rewrite.ml | 6 ++++- printing/prettyp.ml | 14 ++++++------ printing/printer.ml | 4 ++++ printing/printer.mli | 1 + proofs/proof_global.ml | 12 +++++----- tactics/ind_tables.ml | 5 ++--- test-suite/output/UnivBinders.out | 2 +- vernac/classes.ml | 34 +++++++++++++++++++---------- vernac/command.ml | 15 ++++++++----- vernac/indschemes.ml | 6 ++--- vernac/lemmas.ml | 19 ++++++++++++---- vernac/obligations.ml | 45 ++++++++++++++++++++++---------------- vernac/record.ml | 36 +++++++++++++++--------------- 29 files changed, 205 insertions(+), 146 deletions(-) (limited to 'vernac') diff --git a/API/API.mli b/API/API.mli index 5c7860671..ad5d5490b 100644 --- a/API/API.mli +++ b/API/API.mli @@ -1417,7 +1417,7 @@ sig | TemplateArity of 'b type constant_universes = - | Monomorphic_const of Univ.UContext.t + | Monomorphic_const of Univ.ContextSet.t | Polymorphic_const of Univ.AUContext.t type projection_body = { @@ -1484,7 +1484,7 @@ sig | MEwith of module_alg_expr * with_declaration type abstract_inductive_universes = - | Monomorphic_ind of Univ.UContext.t + | Monomorphic_ind of Univ.ContextSet.t | Polymorphic_ind of Univ.AUContext.t | Cumulative_ind of Univ.ACumulativityInfo.t @@ -1551,7 +1551,7 @@ sig | LocalAssumEntry of constr type inductive_universes = - | Monomorphic_ind_entry of Univ.UContext.t + | Monomorphic_ind_entry of Univ.ContextSet.t | Polymorphic_ind_entry of Univ.UContext.t | Cumulative_ind_entry of Univ.CumulativityInfo.t @@ -1580,8 +1580,9 @@ sig type 'a proof_output = Constr.t Univ.in_universe_context_set * 'a type 'a const_entry_body = 'a proof_output Future.computation type constant_universes_entry = - | Monomorphic_const_entry of Univ.UContext.t + | Monomorphic_const_entry of Univ.ContextSet.t | Polymorphic_const_entry of Univ.UContext.t + type 'a in_constant_universes_entry = 'a * constant_universes_entry type 'a definition_entry = { const_entry_body : 'a const_entry_body; (* List of section variables *) @@ -1592,7 +1593,7 @@ sig const_entry_universes : constant_universes_entry; const_entry_opaque : bool; const_entry_inline_code : bool } - type parameter_entry = Context.Named.t option * bool * Constr.types Univ.in_universe_context * inline + type parameter_entry = Context.Named.t option * Constr.types in_constant_universes_entry * inline type projection_entry = { proj_entry_ind : MutInd.t; diff --git a/checker/cic.mli b/checker/cic.mli index 354650964..4a0e706aa 100644 --- a/checker/cic.mli +++ b/checker/cic.mli @@ -208,7 +208,7 @@ type constant_def = | OpaqueDef of lazy_constr type constant_universes = - | Monomorphic_const of Univ.universe_context + | Monomorphic_const of Univ.ContextSet.t | Polymorphic_const of Univ.abstract_universe_context (** The [typing_flags] are instructions to the type-checker which @@ -303,7 +303,7 @@ type one_inductive_body = { } type abstract_inductive_universes = - | Monomorphic_ind of Univ.universe_context + | Monomorphic_ind of Univ.ContextSet.t | Polymorphic_ind of Univ.abstract_universe_context | Cumulative_ind of Univ.abstract_cumulativity_info diff --git a/checker/mod_checking.ml b/checker/mod_checking.ml index 63e28448f..4357a690e 100644 --- a/checker/mod_checking.ml +++ b/checker/mod_checking.ml @@ -29,7 +29,7 @@ let check_constant_declaration env kn cb = (** [env'] contains De Bruijn universe variables *) let env' = match cb.const_universes with - | Monomorphic_const ctx -> push_context ~strict:true ctx env + | Monomorphic_const ctx -> push_context_set ~strict:true ctx env | Polymorphic_const auctx -> let ctx = Univ.AUContext.repr auctx in push_context ~strict:false ctx env diff --git a/checker/values.ml b/checker/values.ml index 9e16c8435..5a371164c 100644 --- a/checker/values.ml +++ b/checker/values.ml @@ -13,7 +13,7 @@ To ensure this file is up-to-date, 'make' now compares the md5 of cic.mli with a copy we maintain here: -MD5 f4b00c567a972ae950b9ed10c533fda5 checker/cic.mli +MD5 56ac4cade33eff3d26ed5cdadb580c7e checker/cic.mli *) @@ -215,7 +215,7 @@ let v_projbody = let v_typing_flags = v_tuple "typing_flags" [|v_bool; v_bool|] -let v_const_univs = v_sum "constant_universes" 0 [|[|v_context|]; [|v_abs_context|]|] +let v_const_univs = v_sum "constant_universes" 0 [|[|v_context_set|]; [|v_abs_context|]|] let v_cb = v_tuple "constant_body" [|v_section_ctxt; @@ -265,7 +265,7 @@ let v_mind_record = Annot ("mind_record", let v_ind_pack_univs = v_sum "abstract_inductive_universes" 0 - [|[|v_context|]; [|v_abs_context|]; [|v_abs_cum_info|]|] + [|[|v_context_set|]; [|v_abs_context|]; [|v_abs_cum_info|]|] let v_ind_pack = v_tuple "mutual_inductive_body" [|Array v_one_ind; diff --git a/interp/declare.ml b/interp/declare.ml index dc77981f2..15999f1d1 100644 --- a/interp/declare.ml +++ b/interp/declare.ml @@ -207,7 +207,9 @@ let definition_entry ?fix_exn ?(opaque=false) ?(inline=false) ?types ?(poly=false) ?(univs=Univ.UContext.empty) ?(eff=Safe_typing.empty_private_constants) body = let univs = if poly then Polymorphic_const_entry univs - else Monomorphic_const_entry univs + else + (* FIXME be smarter about this *) + Monomorphic_const_entry (Univ.ContextSet.of_context univs) in { const_entry_body = Future.from_val ?fix_exn ((body,Univ.ContextSet.empty), eff); const_entry_secctx = None; @@ -340,7 +342,7 @@ let dummy_inductive_entry (_,m) = ([],{ mind_entry_record = None; mind_entry_finite = Decl_kinds.BiFinite; mind_entry_inds = List.map dummy_one_inductive_entry m.mind_entry_inds; - mind_entry_universes = Monomorphic_ind_entry Univ.UContext.empty; + mind_entry_universes = Monomorphic_ind_entry Univ.ContextSet.empty; mind_entry_private = None; }) diff --git a/kernel/declarations.ml b/kernel/declarations.ml index b95796fd8..d5312c500 100644 --- a/kernel/declarations.ml +++ b/kernel/declarations.ml @@ -63,7 +63,7 @@ type constant_def = | OpaqueDef of Opaqueproof.opaque (** or an opaque global definition *) type constant_universes = - | Monomorphic_const of Univ.UContext.t + | Monomorphic_const of Univ.ContextSet.t | Polymorphic_const of Univ.AUContext.t (** The [typing_flags] are instructions to the type-checker which @@ -168,9 +168,9 @@ type one_inductive_body = { } type abstract_inductive_universes = - | Monomorphic_ind of Univ.UContext.t + | Monomorphic_ind of Univ.ContextSet.t | Polymorphic_ind of Univ.AUContext.t - | Cumulative_ind of Univ.ACumulativityInfo.t + | Cumulative_ind of Univ.ACumulativityInfo.t type mutual_inductive_body = { diff --git a/kernel/declareops.ml b/kernel/declareops.ml index f5c26b33d..d8768a0fc 100644 --- a/kernel/declareops.ml +++ b/kernel/declareops.ml @@ -126,7 +126,7 @@ let hcons_const_def = function let hcons_const_universes cbu = match cbu with | Monomorphic_const ctx -> - Monomorphic_const (Univ.hcons_universe_context ctx) + Monomorphic_const (Univ.hcons_universe_context_set ctx) | Polymorphic_const ctx -> Polymorphic_const (Univ.hcons_abstract_universe_context ctx) @@ -274,7 +274,7 @@ let hcons_mind_packet oib = let hcons_mind_universes miu = match miu with - | Monomorphic_ind ctx -> Monomorphic_ind (Univ.hcons_universe_context ctx) + | Monomorphic_ind ctx -> Monomorphic_ind (Univ.hcons_universe_context_set ctx) | Polymorphic_ind ctx -> Polymorphic_ind (Univ.hcons_abstract_universe_context ctx) | Cumulative_ind cui -> Cumulative_ind (Univ.hcons_abstract_cumulativity_info cui) diff --git a/kernel/entries.ml b/kernel/entries.ml index 185dba409..c44a17df2 100644 --- a/kernel/entries.ml +++ b/kernel/entries.ml @@ -35,9 +35,9 @@ then, in i{^ th} block, [mind_entry_params] is [xn:Xn;...;x1:X1]; *) type inductive_universes = - | Monomorphic_ind_entry of Univ.UContext.t + | Monomorphic_ind_entry of Univ.ContextSet.t | Polymorphic_ind_entry of Univ.UContext.t - | Cumulative_ind_entry of Univ.CumulativityInfo.t + | Cumulative_ind_entry of Univ.CumulativityInfo.t type one_inductive_entry = { mind_entry_typename : Id.t; @@ -65,9 +65,11 @@ type 'a proof_output = constr Univ.in_universe_context_set * 'a type 'a const_entry_body = 'a proof_output Future.computation type constant_universes_entry = - | Monomorphic_const_entry of Univ.UContext.t + | Monomorphic_const_entry of Univ.ContextSet.t | Polymorphic_const_entry of Univ.UContext.t +type 'a in_constant_universes_entry = 'a * constant_universes_entry + type 'a definition_entry = { const_entry_body : 'a const_entry_body; (* List of section variables *) @@ -82,7 +84,7 @@ type 'a definition_entry = { type inline = int option (* inlining level, None for no inlining *) type parameter_entry = - Context.Named.t option * bool * types Univ.in_universe_context * inline + Context.Named.t option * types in_constant_universes_entry * inline type projection_entry = { proj_entry_ind : MutInd.t; diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml index 083b0ae40..8e9b606a5 100644 --- a/kernel/indtypes.ml +++ b/kernel/indtypes.ml @@ -265,13 +265,12 @@ let typecheck_inductive env mie = (* Check unicity of names *) mind_check_names mie; (* Params are typed-checked here *) - let univctx = + let env' = match mie.mind_entry_universes with - | Monomorphic_ind_entry ctx -> ctx - | Polymorphic_ind_entry ctx -> ctx - | Cumulative_ind_entry cumi -> Univ.CumulativityInfo.univ_context cumi + | Monomorphic_ind_entry ctx -> push_context_set ctx env + | Polymorphic_ind_entry ctx -> push_context ctx env + | Cumulative_ind_entry cumi -> push_context (Univ.CumulativityInfo.univ_context cumi) env in - let env' = push_context univctx env in let (env_params,paramsctxt) = infer_local_decls env' mie.mind_entry_params in (* We first type arity of each inductive definition *) (* This allows building the environment of arities and to share *) diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml index 8568bf14b..f7e755f00 100644 --- a/kernel/mod_typing.ml +++ b/kernel/mod_typing.ml @@ -79,18 +79,20 @@ let rec check_with_def env struc (idl,(c,ctx)) mp equiv = environment, because they do not appear in the type of the definition. Any inconsistency will be raised at a later stage when joining the environment. *) - let env' = Environ.push_context ~strict:true ctx env' in - let c',cst = match cb.const_body with - | Undef _ | OpaqueDef _ -> - let j = Typeops.infer env' c in - let typ = cb.const_type in - let cst' = Reduction.infer_conv_leq env' (Environ.universes env') - j.uj_type typ in - j.uj_val, cst' - | Def cs -> - let c' = Mod_subst.force_constr cs in - c, Reduction.infer_conv env' (Environ.universes env') c c' - in c', Monomorphic_const ctx, Univ.ContextSet.add_constraints cst (Univ.ContextSet.of_context ctx) + let env' = Environ.push_context ~strict:true ctx env' in + let c',cst = match cb.const_body with + | Undef _ | OpaqueDef _ -> + let j = Typeops.infer env' c in + let typ = cb.const_type in + let cst' = Reduction.infer_conv_leq env' (Environ.universes env') + j.uj_type typ in + j.uj_val, cst' + | Def cs -> + let c' = Mod_subst.force_constr cs in + c, Reduction.infer_conv env' (Environ.universes env') c c' + in + let ctx = Univ.ContextSet.of_context ctx in + c', Monomorphic_const ctx, Univ.ContextSet.add_constraints cst ctx | Polymorphic_const uctx -> let subst, ctx = Univ.abstract_universes ctx in let c = Vars.subst_univs_level_constr subst c in diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml index 0e416b3e5..0e41bfc3c 100644 --- a/kernel/safe_typing.ml +++ b/kernel/safe_typing.ml @@ -249,14 +249,14 @@ let universes_of_private eff = in match cb.const_universes with | Monomorphic_const ctx -> - (Univ.ContextSet.of_context ctx) :: acc + ctx :: acc | Polymorphic_const _ -> acc ) acc l | Entries.SEsubproof (c, cb, e) -> match cb.const_universes with | Monomorphic_const ctx -> - (Univ.ContextSet.of_context ctx) :: acc + ctx :: acc | Polymorphic_const _ -> acc ) [] (Term_typing.uniq_seff eff) @@ -389,7 +389,6 @@ let push_named_def (id,de) senv = | Monomorphic_const_entry _ -> false | Polymorphic_const_entry _ -> true in - let univs = Univ.ContextSet.of_context univs in let c, univs = match c with | Def c -> Mod_subst.force_constr c, univs | OpaqueDef o -> @@ -425,9 +424,8 @@ let labels_of_mib mib = let globalize_constant_universes env cb = match cb.const_universes with - | Monomorphic_const ctx -> - let cstrs = Univ.ContextSet.of_context ctx in - Now (false, cstrs) :: + | Monomorphic_const cstrs -> + Now (false, cstrs) :: (match cb.const_body with | (Undef _ | Def _) -> [] | OpaqueDef lc -> @@ -443,7 +441,7 @@ let globalize_constant_universes env cb = let globalize_mind_universes mb = match mb.mind_universes with | Monomorphic_ind ctx -> - [Now (false, Univ.ContextSet.of_context ctx)] + [Now (false, ctx)] | Polymorphic_ind _ -> [Now (true, Univ.ContextSet.empty)] | Cumulative_ind _ -> [Now (true, Univ.ContextSet.empty)] diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml index 4617f2d5f..70dd6438d 100644 --- a/kernel/term_typing.ml +++ b/kernel/term_typing.ml @@ -125,11 +125,10 @@ let inline_side_effects env body ctx side_eff = | _ -> assert false in match cb.const_universes with - | Monomorphic_const cnstctx -> + | Monomorphic_const univs -> (** Abstract over the term at the top of the proof *) let ty = cb.const_type in let subst = Cmap_env.add c (Inr var) subst in - let univs = Univ.ContextSet.of_context cnstctx in let ctx = Univ.ContextSet.union ctx univs in (subst, var + 1, ctx, (cname c, b, ty, opaque) :: args) | Polymorphic_const auctx -> @@ -228,19 +227,25 @@ let feedback_completion_typecheck = Option.iter (fun state_id -> feedback ~id:state_id Feedback.Complete) -let abstract_constant_universes abstract uctx = - if not abstract then +let abstract_constant_universes abstract = function + | Monomorphic_const_entry uctx -> Univ.empty_level_subst, Monomorphic_const uctx - else - let sbst, auctx = Univ.abstract_universes uctx in - sbst, Polymorphic_const auctx + | Polymorphic_const_entry uctx -> + if not abstract then + Univ.empty_level_subst, Monomorphic_const (Univ.ContextSet.of_context uctx) + else + let sbst, auctx = Univ.abstract_universes uctx in + sbst, Polymorphic_const auctx let infer_declaration (type a) ~(trust : a trust) env kn (dcl : a constant_entry) = match dcl with - | ParameterEntry (ctx,poly,(t,uctx),nl) -> - let env = push_context ~strict:(not poly) uctx env in + | ParameterEntry (ctx,(t,uctx),nl) -> + let env = match uctx with + | Monomorphic_const_entry uctx -> push_context_set ~strict:true uctx env + | Polymorphic_const_entry uctx -> push_context ~strict:false uctx env + in let j = infer env t in - let abstract = poly && not (Option.is_empty kn) in + let abstract = not (Option.is_empty kn) in let usubst, univs = abstract_constant_universes abstract uctx in @@ -262,7 +267,7 @@ let infer_declaration (type a) ~(trust : a trust) env kn (dcl : a constant_entry | DefinitionEntry ({ const_entry_type = Some typ; const_entry_opaque = true; const_entry_universes = Monomorphic_const_entry univs } as c) -> - let env = push_context ~strict:true univs env in + let env = push_context_set ~strict:true univs env in let { const_entry_body = body; const_entry_feedback = feedback_id } = c in let tyj = infer_type env typ in let proofterm = @@ -301,21 +306,22 @@ let infer_declaration (type a) ~(trust : a trust) env kn (dcl : a constant_entry let { const_entry_type = typ; const_entry_opaque = opaque } = c in let { const_entry_body = body; const_entry_feedback = feedback_id } = c in let (body, ctx), side_eff = Future.join body in - let poly, univs = match c.const_entry_universes with + let poly, univsctx = match c.const_entry_universes with | Monomorphic_const_entry univs -> false, univs - | Polymorphic_const_entry univs -> true, univs + | Polymorphic_const_entry univs -> true, Univ.ContextSet.of_context univs in - let univsctx = Univ.ContextSet.of_context univs in let ctx = Univ.ContextSet.union univsctx ctx in let body, ctx, _ = match trust with | Pure -> body, ctx, [] | SideEffects _ -> inline_side_effects env body ctx side_eff in let env = push_context_set ~strict:(not poly) ctx env in - let abstract = poly && not (Option.is_empty kn) in - let usubst, univs = - abstract_constant_universes abstract (Univ.ContextSet.to_context ctx) - in + let abstract = not (Option.is_empty kn) in + let ctx = if poly + then Polymorphic_const_entry (Univ.ContextSet.to_context ctx) + else Monomorphic_const_entry ctx + in + let usubst, univs = abstract_constant_universes abstract ctx in let j = infer env body in let typ = match typ with | None -> @@ -556,7 +562,7 @@ let export_side_effects mb env ce = let env = Environ.add_constant kn cb env in match cb.const_universes with | Monomorphic_const ctx -> - Environ.push_context ~strict:true ctx env + Environ.push_context_set ~strict:true ctx env | Polymorphic_const _ -> env end | kn, cb, `Opaque(_, ctx), _ -> @@ -564,7 +570,7 @@ let export_side_effects mb env ce = let env = Environ.add_constant kn cb env in match cb.const_universes with | Monomorphic_const cstctx -> - let env = Environ.push_context ~strict:true cstctx env in + let env = Environ.push_context_set ~strict:true cstctx env in Environ.push_context_set ~strict:true ctx env | Polymorphic_const _ -> env end diff --git a/kernel/term_typing.mli b/kernel/term_typing.mli index 9b35bfc6e..55da4197e 100644 --- a/kernel/term_typing.mli +++ b/kernel/term_typing.mli @@ -19,7 +19,7 @@ type _ trust = | SideEffects : structure_body -> side_effects trust val translate_local_def : 'a trust -> env -> Id.t -> 'a definition_entry -> - constant_def * types * Univ.UContext.t + constant_def * types * Univ.ContextSet.t val translate_local_assum : env -> types -> types diff --git a/kernel/univ.ml b/kernel/univ.ml index 7fe4f8274..64afb95d5 100644 --- a/kernel/univ.ml +++ b/kernel/univ.ml @@ -1053,6 +1053,7 @@ struct let constraints (univs, cst) = cst let levels (univs, cst) = univs + let size (univs,_) = LSet.cardinal univs end type universe_context_set = ContextSet.t diff --git a/kernel/univ.mli b/kernel/univ.mli index 8d46a8bee..c06ce2446 100644 --- a/kernel/univ.mli +++ b/kernel/univ.mli @@ -310,7 +310,7 @@ sig (** Keeps the order of the instances *) val union : t -> t -> t - (* the number of universes in the context *) + (** the number of universes in the context *) val size : t -> int end @@ -423,6 +423,9 @@ sig val constraints : t -> constraints val levels : t -> LSet.t + + (** the number of universes in the context *) + val size : t -> int end (** A set of universes with universe constraints. diff --git a/kernel/vconv.ml b/kernel/vconv.ml index 0e452621c..578a89371 100644 --- a/kernel/vconv.ml +++ b/kernel/vconv.ml @@ -93,7 +93,7 @@ and conv_atom env pb k a1 stk1 a2 stk2 cu = let mib = Environ.lookup_mind mi env in let ulen = match mib.Declarations.mind_universes with - | Declarations.Monomorphic_ind ctx -> Univ.UContext.size ctx + | Declarations.Monomorphic_ind ctx -> Univ.ContextSet.size ctx | Declarations.Polymorphic_ind auctx -> Univ.AUContext.size auctx | Declarations.Cumulative_ind cumi -> Univ.AUContext.size (Univ.ACumulativityInfo.univ_context cumi) diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index f6523d28c..a2f5a4577 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -1972,9 +1972,13 @@ let add_morphism_infer glob m n = let evd = Evd.from_env env in let uctx, instance = build_morphism_signature env evd m in if Lib.is_modtype () then + let uctx = if poly + then Entries.Polymorphic_const_entry (UState.context uctx) + else Entries.Monomorphic_const_entry (UState.context_set uctx) + in let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest instance_id (Entries.ParameterEntry - (None,poly,(instance,UState.context uctx),None), + (None,(instance,uctx),None), Decl_kinds.IsAssumption Decl_kinds.Logical) in add_instance (Typeclasses.new_instance diff --git a/printing/prettyp.ml b/printing/prettyp.ml index 8fc00ed96..b4ac94103 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -556,26 +556,26 @@ let print_constant with_values sep sp = Vars.subst_instance_constr inst cb.const_type in let univs = + let open Entries in let otab = Global.opaque_tables () in match cb.const_body with | Undef _ | Def _ -> begin match cb.const_universes with - | Monomorphic_const ctx -> ctx + | Monomorphic_const ctx -> Monomorphic_const_entry ctx | Polymorphic_const ctx -> let inst = Univ.AUContext.instance ctx in - Univ.UContext.make (inst, Univ.AUContext.instantiate inst ctx) + Polymorphic_const_entry (Univ.UContext.make (inst, Univ.AUContext.instantiate inst ctx)) end | OpaqueDef o -> let body_uctxs = Opaqueproof.force_constraints otab o in match cb.const_universes with | Monomorphic_const ctx -> - let uctxs = Univ.ContextSet.of_context ctx in - Univ.ContextSet.to_context (Univ.ContextSet.union body_uctxs uctxs) + Monomorphic_const_entry (Univ.ContextSet.union body_uctxs ctx) | Polymorphic_const ctx -> assert(Univ.ContextSet.is_empty body_uctxs); let inst = Univ.AUContext.instance ctx in - Univ.UContext.make (inst, Univ.AUContext.instantiate inst ctx) + Polymorphic_const_entry (Univ.UContext.make (inst, Univ.AUContext.instantiate inst ctx)) in let ctx = Evd.evar_universe_context_of_binders @@ -589,12 +589,12 @@ let print_constant with_values sep sp = str"*** [ " ++ print_basename sp ++ print_instance sigma cb ++ str " : " ++ cut () ++ pr_ltype typ ++ str" ]" ++ - Printer.pr_universe_ctx sigma univs + Printer.pr_constant_universes sigma univs | Some (c, ctx) -> let c = Vars.subst_instance_constr (Univ.AUContext.instance ctx) c in print_basename sp ++ print_instance sigma cb ++ str sep ++ cut () ++ (if with_values then print_typed_body env sigma (Some c,typ) else pr_ltype typ)++ - Printer.pr_universe_ctx sigma univs) + Printer.pr_constant_universes sigma univs) let gallina_print_constant_with_infos sp = print_constant true " = " sp ++ diff --git a/printing/printer.ml b/printing/printer.ml index 4d5cfcba3..ae6292024 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -270,6 +270,10 @@ let pr_universe_ctx sigma c = else mt() +let pr_constant_universes sigma = function + | Entries.Monomorphic_const_entry ctx -> pr_universe_ctx_set sigma ctx + | Entries.Polymorphic_const_entry ctx -> pr_universe_ctx sigma ctx + let pr_cumulativity_info sigma cumi = if !Detyping.print_universes && not (Univ.UContext.is_empty (Univ.CumulativityInfo.univ_context cumi)) then diff --git a/printing/printer.mli b/printing/printer.mli index a3d0eaac2..67128da40 100644 --- a/printing/printer.mli +++ b/printing/printer.mli @@ -122,6 +122,7 @@ val pr_cumulative : bool -> bool -> Pp.t val pr_universe_instance : evar_map -> Univ.UContext.t -> Pp.t val pr_universe_ctx : evar_map -> Univ.UContext.t -> Pp.t val pr_universe_ctx_set : evar_map -> Univ.ContextSet.t -> Pp.t +val pr_constant_universes : evar_map -> Entries.constant_universes_entry -> Pp.t val pr_cumulativity_info : evar_map -> Univ.CumulativityInfo.t -> Pp.t (** Printing global references using names as short as possible *) diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index ea0408169..905173efc 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -331,7 +331,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now let initial_euctx = Proof.initial_euctx proof in let fpl, univs = Future.split2 fpl in let universes = if poly || now then Future.force univs else initial_euctx in - let univctx = Evd.check_univ_decl (Evd.from_ctx universes) universe_decl in + let univctx = UState.check_univ_decl universes universe_decl in let binders = if poly then Some (UState.universe_binders universes) else None in (* Because of dependent subgoals at the beginning of proofs, we could have existential variables in the initial types of goals, we need to @@ -360,7 +360,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now the body. So we keep the two sets distinct. *) let used_univs = Univ.LSet.union used_univs_body used_univs_typ in let ctx_body = UState.restrict ctx used_univs in - let univs = Evd.check_univ_decl (Evd.from_ctx ctx_body) universe_decl in + let univs = UState.check_univ_decl ctx_body universe_decl in (initunivs, typ), ((body, Univ.ContextSet.of_context univs), eff) else (* Since the proof is computed now, we can simply have 1 set of @@ -370,7 +370,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now TODO: check if restrict is really necessary now. *) let used_univs = Univ.LSet.union used_univs_body used_univs_typ in let ctx = UState.restrict universes used_univs in - let univs = Evd.check_univ_decl (Evd.from_ctx ctx) universe_decl in + let univs = UState.check_univ_decl ctx universe_decl in (univs, typ), ((body, Univ.ContextSet.empty), eff) in fun t p -> Future.split2 (Future.chain p (make_body t)) @@ -383,7 +383,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now the declaration as well. If the declaration is non-extensible, this will prevent the body from adding universes and constraints. *) let bodyunivs = constrain_variables univctx (Future.force univs) in - let univs = Evd.check_univ_decl (Evd.from_ctx bodyunivs) universe_decl in + let univs = UState.check_univ_decl bodyunivs universe_decl in (pt,Univ.ContextSet.of_context univs),eff) in let entry_fn p (_, t) = @@ -392,7 +392,9 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now let univs, typ = Future.force univstyp in let univs = if poly then Entries.Polymorphic_const_entry univs - else Entries.Monomorphic_const_entry univs + else + (* FIXME be smarter about the unnecessary context linearisation in make_body *) + Entries.Monomorphic_const_entry (Univ.ContextSet.of_context univs) in {Entries. const_entry_body = body; diff --git a/tactics/ind_tables.ml b/tactics/ind_tables.ml index 92c326b1e..e1bf32f3c 100644 --- a/tactics/ind_tables.ml +++ b/tactics/ind_tables.ml @@ -123,10 +123,9 @@ let define internal id c p univs = let ctx = Evd.normalize_evar_universe_context univs in let c = Vars.subst_univs_fn_constr (Universes.make_opt_subst (Evd.evar_universe_context_subst ctx)) c in - let univs = UState.context ctx in let univs = - if p then Polymorphic_const_entry univs - else Monomorphic_const_entry univs + if p then Polymorphic_const_entry (UState.context ctx) + else Monomorphic_const_entry (UState.context_set ctx) in let entry = { const_entry_body = diff --git a/test-suite/output/UnivBinders.out b/test-suite/output/UnivBinders.out index 1bacb7513..42fb52a3b 100644 --- a/test-suite/output/UnivBinders.out +++ b/test-suite/output/UnivBinders.out @@ -50,6 +50,6 @@ Type@{Top.17} -> Type@{v} -> Type@{u} foo is universe polymorphic Monomorphic mono = Type@{u} : Type@{u+1} -(* u |= *) +(* {u} |= *) mono is not universe polymorphic diff --git a/vernac/classes.ml b/vernac/classes.ml index 6f5f96ee2..c99eba2cc 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -200,16 +200,22 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) let nf, subst = Evarutil.e_nf_evars_and_universes evars in let termtype = let t = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in - nf t - in - Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr termtype); + nf t + in + Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr termtype); let ctx = Evd.check_univ_decl !evars decl in - let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id - (ParameterEntry - (None,poly,(termtype,ctx),None), Decl_kinds.IsAssumption Decl_kinds.Logical) - in + let ctx = if poly + then Polymorphic_const_entry ctx + else + (* FIXME be smarter about this *) + Monomorphic_const_entry (Univ.ContextSet.of_context ctx) + in + let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id + (ParameterEntry + (None,(termtype,ctx),None), Decl_kinds.IsAssumption Decl_kinds.Logical) + in Universes.register_universe_binders (ConstRef cst) (Evd.universe_binders !evars); - instance_hook k pri global imps ?hook (ConstRef cst); id + instance_hook k pri global imps ?hook (ConstRef cst); id end else ( let props = @@ -384,14 +390,20 @@ let context poly l = let uctx = ref (Evd.universe_context_set !evars) in let fn status (id, b, t) = if Lib.is_modtype () && not (Lib.sections_are_opened ()) then - let ctx = Univ.ContextSet.to_context !uctx in (* Declare the universe context once *) + let univs = !uctx in let () = uctx := Univ.ContextSet.empty in let decl = match b with | None -> - (ParameterEntry (None,poly,(t,ctx),None), IsAssumption Logical) + let univs = if poly + then Polymorphic_const_entry (Univ.ContextSet.to_context univs) + else Monomorphic_const_entry univs + in + (ParameterEntry (None,(t,univs),None), IsAssumption Logical) | Some b -> - let entry = Declare.definition_entry ~poly ~univs:ctx ~types:t b in + (* FIXME be smarter about this *) + let univs = Univ.ContextSet.to_context univs in + let entry = Declare.definition_entry ~poly ~univs ~types:t b in (DefinitionEntry entry, IsAssumption Logical) in let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in diff --git a/vernac/command.ml b/vernac/command.ml index 80599c534..06d0c8470 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -197,8 +197,11 @@ match local with | DefaultInline -> Some (Flags.get_inline_level()) | InlineAt i -> Some i in - let ctx = Univ.ContextSet.to_context ctx in - let decl = (ParameterEntry (None,p,(c,ctx),inl), IsAssumption kind) in + let ctx = if p + then Polymorphic_const_entry (Univ.ContextSet.to_context ctx) + else Monomorphic_const_entry ctx + in + let decl = (ParameterEntry (None,(c,ctx),inl), IsAssumption kind) in let kn = declare_constant ident ~local decl in let gr = ConstRef kn in let () = maybe_declare_manual_implicits false gr imps in @@ -206,9 +209,9 @@ match local with let () = assumption_message ident in let () = Typeclasses.declare_instance None false gr in let () = if is_coe then Class.try_add_new_coercion gr ~local p in - let inst = - if p (* polymorphic *) then Univ.UContext.instance ctx - else Univ.Instance.empty + let inst = match ctx with + | Polymorphic_const_entry ctx -> Univ.UContext.instance ctx + | Monomorphic_const_entry _ -> Univ.Instance.empty in (gr,inst,Lib.is_modtype_strict ()) @@ -606,7 +609,7 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite = Cumulative_ind_entry (Universes.univ_inf_ind_from_universe_context uctx) else Polymorphic_ind_entry uctx else - Monomorphic_ind_entry uctx + Monomorphic_ind_entry (Univ.ContextSet.of_context uctx) in (* Build the mutual inductive entry *) let mind_ent = diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml index c728c2671..e4ca98749 100644 --- a/vernac/indschemes.ml +++ b/vernac/indschemes.ml @@ -109,10 +109,10 @@ let _ = let define id internal ctx c t = let f = declare_constant ~internal in - let univs = Evd.to_universe_context ctx in let univs = - if Flags.is_universe_polymorphism () then Polymorphic_const_entry univs - else Monomorphic_const_entry univs + if Flags.is_universe_polymorphism () + then Polymorphic_const_entry (Evd.to_universe_context ctx) + else Monomorphic_const_entry (Evd.universe_context_set ctx) in let kn = f id (DefinitionEntry diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index 71d80c4db..f59b5fcae 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -242,7 +242,11 @@ let save_remaining_recthms (locality,p,kind) norm ctx body opaq i (id,(t_i,(_,im | Global -> false | Discharge -> assert false in - let decl = (ParameterEntry (None,p,(t_i,ctx),None), k) in + let ctx = if p + then Polymorphic_const_entry ctx + else Monomorphic_const_entry (Univ.ContextSet.of_context ctx) + in + let decl = (ParameterEntry (None,(t_i,ctx),None), k) in let kn = declare_constant id ~local decl in (locality,ConstRef kn,imps)) | Some body -> @@ -491,9 +495,12 @@ let save_proof ?proof = function if const_entry_type = None then user_err Pp.(str "Admitted requires an explicit statement"); let typ = Option.get const_entry_type in - let ctx = UState.context (fst universes) in + let ctx = if pi2 k (* polymorphic *) + then Polymorphic_const_entry (UState.context (fst universes)) + else Monomorphic_const_entry (UState.context_set (fst universes)) + in let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in - Admitted(id, k, (sec_vars, pi2 k, (typ, ctx), None), universes) + Admitted(id, k, (sec_vars, (typ, ctx), None), universes) | None -> let pftree = Proof_global.give_me_the_proof () in let id, k, typ = Pfedit.current_proof_statement () in @@ -516,8 +523,12 @@ let save_proof ?proof = function let evd = Evd.from_ctx universes in let ctx = Evd.check_univ_decl evd decl in let poly = pi2 k in + let ctx = if poly + then Polymorphic_const_entry ctx + else Monomorphic_const_entry (Univ.ContextSet.of_context ctx) + in let binders = if poly then Some (UState.universe_binders universes) else None in - Admitted(id,k,(sec_vars, poly, (typ, ctx), None), + Admitted(id,k,(sec_vars, (typ, ctx), None), (universes, binders)) in Proof_global.apply_terminator (Proof_global.get_terminator ()) pe diff --git a/vernac/obligations.ml b/vernac/obligations.ml index 43af8968f..a9110c76c 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -633,12 +633,11 @@ let declare_obligation prg obl body ty uctx = shrink_body body ty else [], body, ty, [||] in let body = ((body,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in - let univs = if poly then Polymorphic_const_entry uctx else Monomorphic_const_entry uctx in let ce = { const_entry_body = Future.from_val ~fix_exn:(fun x -> x) body; const_entry_secctx = None; const_entry_type = ty; - const_entry_universes = univs; + const_entry_universes = uctx; const_entry_opaque = opaque; const_entry_inline_code = false; const_entry_feedback = None; @@ -647,13 +646,15 @@ let declare_obligation prg obl body ty uctx = let constant = Declare.declare_constant obl.obl_name ~local:true (DefinitionEntry ce,IsProof Property) in - if not opaque then add_hint (Locality.make_section_locality None) prg constant; - definition_message obl.obl_name; - true, { obl with obl_body = - if poly then - Some (DefinedObl (constant, Univ.UContext.instance uctx)) - else - Some (TermObl (it_mkLambda_or_LetIn_or_clean (mkApp (mkConst constant, args)) ctx)) } + if not opaque then add_hint (Locality.make_section_locality None) prg constant; + definition_message obl.obl_name; + let body = match uctx with + | Polymorphic_const_entry uctx -> + Some (DefinedObl (constant, Univ.UContext.instance uctx)) + | Monomorphic_const_entry _ -> + Some (TermObl (it_mkLambda_or_LetIn_or_clean (mkApp (mkConst constant, args)) ctx)) + in + true, { obl with obl_body = body } let init_prog_info ?(opaque = false) sign n udecl b t ctx deps fixkind notations obls impls kind reduce hook = @@ -855,7 +856,10 @@ let obligation_terminator name num guard hook auto pf = | (_, status), Vernacexpr.Transparent -> status in let obl = { obl with obl_status = false, status } in - let uctx = UState.context ctx in + let uctx = if pi2 prg.prg_kind + then Polymorphic_const_entry (UState.context ctx) + else Monomorphic_const_entry (UState.context_set ctx) + in let (_, obl) = declare_obligation prg obl body ty uctx in let obls = Array.copy obls in let _ = obls.(num) <- obl in @@ -967,13 +971,16 @@ and solve_obligation_by_tac prg obls i tac = let evd = Evd.from_ctx prg.prg_ctx in let evd = Evd.update_sigma_env evd (Global.env ()) in let t, ty, ctx = - solve_by_tac obl.obl_name (evar_of_obligation obl) tac - (pi2 prg.prg_kind) (Evd.evar_universe_context evd) - in - let uctx = UState.context ctx in - let prg = {prg with prg_ctx = ctx} in - let def, obl' = declare_obligation prg obl t ty uctx in - obls.(i) <- obl'; + solve_by_tac obl.obl_name (evar_of_obligation obl) tac + (pi2 prg.prg_kind) (Evd.evar_universe_context evd) + in + let uctx = if pi2 prg.prg_kind + then Polymorphic_const_entry (UState.context ctx) + else Monomorphic_const_entry (UState.context_set ctx) + in + let prg = {prg with prg_ctx = ctx} in + let def, obl' = declare_obligation prg obl t ty uctx in + obls.(i) <- obl'; if def && not (pi2 prg.prg_kind) then ( (* Declare the term constraints with the first obligation only *) let evd = Evd.from_env (Global.env ()) in @@ -1121,9 +1128,9 @@ let admit_prog prg = match x.obl_body with | None -> let x = subst_deps_obl obls x in - let ctx = UState.context prg.prg_ctx in + let ctx = Monomorphic_const_entry (UState.context_set prg.prg_ctx) in let kn = Declare.declare_constant x.obl_name ~local:true - (ParameterEntry (None,false,(x.obl_type,ctx),None), IsAssumption Conjectural) + (ParameterEntry (None,(x.obl_type,ctx),None), IsAssumption Conjectural) in assumption_message x.obl_name; obls.(i) <- { x with obl_body = Some (DefinedObl (kn, Univ.Instance.empty)) } diff --git a/vernac/record.ml b/vernac/record.ml index 0819dadb4..faf277d86 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -270,7 +270,10 @@ let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers u let env = Global.env() in let (mib,mip) = Global.lookup_inductive indsp in let poly = Declareops.inductive_is_polymorphic mib in - let u = Univ.UContext.instance ctx in + let u = match ctx with + | Polymorphic_const_entry ctx -> Univ.UContext.instance ctx + | Monomorphic_const_entry ctx -> Univ.Instance.empty + in let paramdecls = Inductive.inductive_paramdecls (mib, u) in let indu = indsp, u in let r = mkIndU (indsp,u) in @@ -327,16 +330,12 @@ let declare_projections indsp ctx ?(kind=StructureComponent) binder_name coers u let projtyp = it_mkProd_or_LetIn (mkProd (x,rp,ccl)) paramdecls in try - let univs = - if poly then Polymorphic_const_entry ctx - else Monomorphic_const_entry ctx - in let entry = { const_entry_body = Future.from_val (Safe_typing.mk_pure_proof proj); const_entry_secctx = None; const_entry_type = Some projtyp; - const_entry_universes = univs; + const_entry_universes = ctx; const_entry_opaque = false; const_entry_inline_code = false; const_entry_feedback = None } in @@ -391,11 +390,14 @@ let declare_structure finite ubinders univs id idbuild paramimpls params arity t let args = Context.Rel.to_extended_list mkRel nfields params in let ind = applist (mkRel (1+nparams+nfields), args) in let type_constructor = it_mkProd_or_LetIn ind fields in - let poly, ctx = + let template, ctx = match univs with - | Monomorphic_ind_entry ctx -> false, Univ.UContext.empty - | Polymorphic_ind_entry ctx -> true, ctx - | Cumulative_ind_entry cumi -> true, (Univ.CumulativityInfo.univ_context cumi) + | Monomorphic_ind_entry ctx -> + template, Monomorphic_const_entry Univ.ContextSet.empty + | Polymorphic_ind_entry ctx -> + false, Polymorphic_const_entry ctx + | Cumulative_ind_entry cumi -> + false, Polymorphic_const_entry (Univ.CumulativityInfo.univ_context cumi) in let binder_name = match name with @@ -405,7 +407,7 @@ let declare_structure finite ubinders univs id idbuild paramimpls params arity t let mie_ind = { mind_entry_typename = id; mind_entry_arity = arity; - mind_entry_template = not poly && template; + mind_entry_template = template; mind_entry_consnames = [idbuild]; mind_entry_lc = [type_constructor] } in @@ -419,14 +421,13 @@ let declare_structure finite ubinders univs id idbuild paramimpls params arity t } in let mie = - if poly then - begin + match ctx with + | Polymorphic_const_entry ctx -> let env = Global.env () in let env' = Environ.push_context ctx env in let evd = Evd.from_env env' in Inductiveops.infer_inductive_subtyping env' evd mie - end - else + | Monomorphic_const_entry _ -> mie in let kn = Command.declare_mutual_inductive_with_eliminations mie ubinders [(paramimpls,[])] in @@ -434,6 +435,7 @@ let declare_structure finite ubinders univs id idbuild paramimpls params arity t let cstr = (rsp,1) in let kinds,sp_projs = declare_projections rsp ctx ~kind binder_name coers ubinders fieldimpls fields in let build = ConstructRef cstr in + let poly = match ctx with | Polymorphic_const_entry _ -> true | Monomorphic_const_entry _ -> false in let () = if is_coe then Class.try_add_new_coercion build ~local:false poly in Recordops.declare_structure(rsp,cstr,List.rev kinds,List.rev sp_projs); rsp @@ -499,7 +501,7 @@ let declare_class finite def cum poly ubinders ctx id idbuild paramimpls params else Polymorphic_ind_entry ctx else - Monomorphic_ind_entry ctx + Monomorphic_ind_entry (Univ.ContextSet.of_context ctx) in let ind = declare_structure BiFinite ubinders univs (snd id) idbuild paramimpls params arity template fieldimpls fields @@ -626,7 +628,7 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf else Polymorphic_ind_entry ctx else - Monomorphic_ind_entry ctx + Monomorphic_ind_entry (Univ.ContextSet.of_context ctx) in let ind = declare_structure finite pl univs idstruc idbuild implpars params arity template implfs -- cgit v1.2.3 From 34d85e1e899f8a045659ccc53bfd6a1f5104130b Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Mon, 18 Sep 2017 17:22:24 +0200 Subject: Use Entries.constant_universes_entry more. This reduces conversions between ContextSet/UContext and encodes whether we are polymorphic by which constructor we use rather than using some boolean. --- API/API.mli | 11 ++- engine/evd.ml | 5 +- engine/evd.mli | 8 +- engine/uState.ml | 107 ++++++++++++++++++-------- engine/uState.mli | 17 +++- interp/declare.ml | 13 +--- interp/declare.mli | 6 +- plugins/funind/functional_principles_types.ml | 7 +- plugins/funind/recdef.ml | 8 +- plugins/ltac/extratactics.ml4 | 6 +- plugins/ltac/rewrite.ml | 9 +-- plugins/setoid_ring/newring.ml | 10 +-- proofs/proof_global.ml | 31 +++----- tactics/leminv.ml | 11 ++- vernac/class.ml | 4 +- vernac/classes.ml | 34 ++++---- vernac/command.ml | 44 ++++++----- vernac/declareDef.ml | 4 +- vernac/declareDef.mli | 7 +- vernac/lemmas.ml | 36 ++++----- vernac/obligations.ml | 17 ++-- vernac/record.ml | 54 +++++++------ 22 files changed, 251 insertions(+), 198 deletions(-) (limited to 'vernac') diff --git a/API/API.mli b/API/API.mli index ad5d5490b..4405bf8da 100644 --- a/API/API.mli +++ b/API/API.mli @@ -2764,6 +2764,9 @@ sig val context_set : t -> Univ.ContextSet.t val of_context_set : Univ.ContextSet.t -> t + val const_univ_entry : poly:bool -> t -> Entries.constant_universes_entry + val ind_univ_entry : poly:bool -> t -> Entries.inductive_universes + type rigid = | UnivRigid | UnivFlexible of bool @@ -2887,6 +2890,8 @@ sig val from_ctx : UState.t -> evar_map val to_universe_context : evar_map -> Univ.UContext.t + val const_univ_entry : poly:bool -> evar_map -> Entries.constant_universes_entry + val ind_univ_entry : poly:bool -> evar_map -> Entries.inductive_universes val meta_list : evar_map -> (Constr.metavariable * clbinding) list @@ -4669,11 +4674,11 @@ sig val declare_definition : ?internal:internal_flag -> ?opaque:bool -> ?kind:Decl_kinds.definition_object_kind -> - ?local:bool -> ?poly:Decl_kinds.polymorphic -> Names.Id.t -> ?types:Constr.t -> - Constr.t Univ.in_universe_context_set -> Names.Constant.t + ?local:bool -> Names.Id.t -> ?types:Constr.t -> + Constr.t Entries.in_constant_universes_entry -> Names.Constant.t val definition_entry : ?fix_exn:Future.fix_exn -> ?opaque:bool -> ?inline:bool -> ?types:Constr.types -> - ?poly:Decl_kinds.polymorphic -> ?univs:Univ.UContext.t -> + ?univs:Entries.constant_universes_entry -> ?eff:Safe_typing.private_constants -> Constr.t -> Safe_typing.private_constants Entries.definition_entry val definition_message : Names.Id.t -> unit val declare_variable : Names.Id.t -> variable_declaration -> Libnames.object_name diff --git a/engine/evd.ml b/engine/evd.ml index eb0f2e3e7..9ea2cc00f 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -758,7 +758,10 @@ let universe_context_set d = UState.context_set d.universes let to_universe_context evd = UState.context evd.universes -let check_univ_decl evd decl = UState.check_univ_decl evd.universes decl +let const_univ_entry ~poly evd = UState.const_univ_entry ~poly evd.universes +let ind_univ_entry ~poly evd = UState.ind_univ_entry ~poly evd.universes + +let check_univ_decl ~poly evd decl = UState.check_univ_decl ~poly evd.universes decl let restrict_universe_context evd vars = { evd with universes = UState.restrict evd.universes vars } diff --git a/engine/evd.mli b/engine/evd.mli index 49b6a7583..f06fb8d3b 100644 --- a/engine/evd.mli +++ b/engine/evd.mli @@ -562,8 +562,12 @@ val universes : evar_map -> UGraph.t [Univ.ContextSet.to_context]. *) val to_universe_context : evar_map -> Univ.UContext.t -val check_univ_decl : evar_map -> UState.universe_decl -> - Univ.UContext.t +val const_univ_entry : poly:bool -> evar_map -> Entries.constant_universes_entry + +(** NB: [ind_univ_entry] cannot create cumulative entries. *) +val ind_univ_entry : poly:bool -> evar_map -> Entries.inductive_universes + +val check_univ_decl : poly:bool -> evar_map -> UState.universe_decl -> Entries.constant_universes_entry val merge_universe_context : evar_map -> UState.t -> evar_map val set_universe_context : evar_map -> UState.t -> evar_map diff --git a/engine/uState.ml b/engine/uState.ml index ff91493ee..dadc004c5 100644 --- a/engine/uState.ml +++ b/engine/uState.ml @@ -87,6 +87,17 @@ let constraints ctx = snd ctx.uctx_local let context ctx = Univ.ContextSet.to_context ctx.uctx_local +let const_univ_entry ~poly uctx = + let open Entries in + if poly then Polymorphic_const_entry (context uctx) + else Monomorphic_const_entry (context_set uctx) + +(* does not support cumulativity since you need more info *) +let ind_univ_entry ~poly uctx = + let open Entries in + if poly then Polymorphic_ind_entry (context uctx) + else Monomorphic_ind_entry (context_set uctx) + let of_context_set ctx = { empty with uctx_local = ctx } let subst ctx = ctx.uctx_univ_variables @@ -260,58 +271,92 @@ let pr_uctx_level uctx = type universe_decl = (Names.Id.t Loc.located list, Univ.Constraint.t) Misctypes.gen_universe_decl +let error_unbound_universes left uctx = + let open Univ in + let n = LSet.cardinal left in + let loc = + try + let info = + LMap.find (LSet.choose left) (snd uctx.uctx_names) in + info.uloc + with Not_found -> None + in + user_err ?loc ~hdr:"universe_context" + ((str(CString.plural n "Universe") ++ spc () ++ + LSet.pr (pr_uctx_level uctx) left ++ + spc () ++ str (CString.conjugate_verb_to_be n) ++ + str" unbound.")) + let universe_context ~names ~extensible uctx = - let levels = Univ.ContextSet.levels uctx.uctx_local in + let open Univ in + let levels = ContextSet.levels uctx.uctx_local in let newinst, left = List.fold_right (fun (loc,id) (newinst, acc) -> let l = try UNameMap.find id (fst uctx.uctx_names) - with Not_found -> - user_err ?loc ~hdr:"universe_context" - (str"Universe " ++ Id.print id ++ str" is not bound anymore.") - in (l :: newinst, Univ.LSet.remove l acc)) + with Not_found -> assert false + in (l :: newinst, LSet.remove l acc)) names ([], levels) in - if not extensible && not (Univ.LSet.is_empty left) then - let n = Univ.LSet.cardinal left in - let loc = - try - let info = - Univ.LMap.find (Univ.LSet.choose left) (snd uctx.uctx_names) in - info.uloc - with Not_found -> None - in - user_err ?loc ~hdr:"universe_context" - ((str(CString.plural n "Universe") ++ spc () ++ - Univ.LSet.pr (pr_uctx_level uctx) left ++ - spc () ++ str (CString.conjugate_verb_to_be n) ++ - str" unbound.")) + if not extensible && not (LSet.is_empty left) + then error_unbound_universes left uctx else - let left = Univ.ContextSet.sort_levels (Array.of_list (Univ.LSet.elements left)) in + let left = ContextSet.sort_levels (Array.of_list (LSet.elements left)) in let inst = Array.append (Array.of_list newinst) left in - let inst = Univ.Instance.of_array inst in - let ctx = Univ.UContext.make (inst, - Univ.ContextSet.constraints uctx.uctx_local) in + let inst = Instance.of_array inst in + let ctx = UContext.make (inst, ContextSet.constraints uctx.uctx_local) in ctx -let check_implication uctx cstrs ctx = +let check_universe_context_set ~names ~extensible uctx = + if extensible then () + else + let open Univ in + let left = List.fold_left (fun left (loc,id) -> + let l = + try UNameMap.find id (fst uctx.uctx_names) + with Not_found -> assert false + in LSet.remove l left) + (ContextSet.levels uctx.uctx_local) names + in + if not (LSet.is_empty left) + then error_unbound_universes left uctx + +let check_implication uctx cstrs cstrs' = let gr = initial_graph uctx in let grext = UGraph.merge_constraints cstrs gr in - let cstrs' = Univ.UContext.constraints ctx in if UGraph.check_constraints cstrs' grext then () else CErrors.user_err ~hdr:"check_univ_decl" (str "Universe constraints are not implied by the ones declared.") -let check_univ_decl uctx decl = +let check_mono_univ_decl uctx decl = + let open Misctypes in + let () = + let names = decl.univdecl_instance in + let extensible = decl.univdecl_extensible_instance in + check_universe_context_set ~names ~extensible uctx + in + if not decl.univdecl_extensible_constraints then + check_implication uctx + decl.univdecl_constraints + (Univ.ContextSet.constraints uctx.uctx_local); + uctx.uctx_local + +let check_univ_decl ~poly uctx decl = let open Misctypes in - let ctx = universe_context - ~names:decl.univdecl_instance - ~extensible:decl.univdecl_extensible_instance - uctx + let ctx = + let names = decl.univdecl_instance in + let extensible = decl.univdecl_extensible_instance in + if poly + then Entries.Polymorphic_const_entry (universe_context ~names ~extensible uctx) + else + let () = check_universe_context_set ~names ~extensible uctx in + Entries.Monomorphic_const_entry uctx.uctx_local in if not decl.univdecl_extensible_constraints then - check_implication uctx decl.univdecl_constraints ctx; + check_implication uctx + decl.univdecl_constraints + (Univ.ContextSet.constraints uctx.uctx_local); ctx let restrict ctx vars = diff --git a/engine/uState.mli b/engine/uState.mli index 5279c6388..4265f2b20 100644 --- a/engine/uState.mli +++ b/engine/uState.mli @@ -59,6 +59,13 @@ val constraints : t -> Univ.constraints val context : t -> Univ.UContext.t (** Shorthand for {!context_set} with {!Context_set.to_context}. *) +val const_univ_entry : poly:bool -> t -> Entries.constant_universes_entry +(** Pick from {!context} or {!context_set} based on [poly]. *) + +val ind_univ_entry : poly:bool -> t -> Entries.inductive_universes +(** Pick from {!context} or {!context_set} based on [poly]. + Cannot create cumulative entries. *) + (** {5 Constraints handling} *) val add_constraints : t -> Univ.constraints -> t @@ -133,11 +140,15 @@ type universe_decl = If non extensible in [decl], check that the local universes (resp. universe constraints) in [ctx] are implied by [decl]. - Return a universe context containing the local universes of [ctx] - and their constraints. The universes corresponding to + Return a [Entries.constant_universes_entry] containing the local + universes of [ctx] and their constraints. + + When polymorphic, the universes corresponding to [decl.univdecl_instance] come first in the order defined by that list. *) -val check_univ_decl : t -> universe_decl -> Univ.UContext.t +val check_univ_decl : poly:bool -> t -> universe_decl -> Entries.constant_universes_entry + +val check_mono_univ_decl : t -> universe_decl -> Univ.ContextSet.t (** {5 TODO: Document me} *) diff --git a/interp/declare.ml b/interp/declare.ml index 15999f1d1..1b4645aff 100644 --- a/interp/declare.ml +++ b/interp/declare.ml @@ -203,14 +203,9 @@ let declare_constant_common id cst = update_tables c; c +let default_univ_entry = Monomorphic_const_entry Univ.ContextSet.empty let definition_entry ?fix_exn ?(opaque=false) ?(inline=false) ?types - ?(poly=false) ?(univs=Univ.UContext.empty) ?(eff=Safe_typing.empty_private_constants) body = - let univs = - if poly then Polymorphic_const_entry univs - else - (* FIXME be smarter about this *) - Monomorphic_const_entry (Univ.ContextSet.of_context univs) - in + ?(univs=default_univ_entry) ?(eff=Safe_typing.empty_private_constants) body = { const_entry_body = Future.from_val ?fix_exn ((body,Univ.ContextSet.empty), eff); const_entry_secctx = None; const_entry_type = types; @@ -263,9 +258,9 @@ let declare_constant ?(internal = UserIndividualRequest) ?(local = false) id ?(e let declare_definition ?(internal=UserIndividualRequest) ?(opaque=false) ?(kind=Decl_kinds.Definition) ?(local = false) - ?(poly=false) id ?types (body,ctx) = + id ?types (body,univs) = let cb = - definition_entry ?types ~poly ~univs:(Univ.ContextSet.to_context ctx) ~opaque body + definition_entry ?types ~univs ~opaque body in declare_constant ~internal ~local id (Entries.DefinitionEntry cb, Decl_kinds.IsDefinition kind) diff --git a/interp/declare.mli b/interp/declare.mli index 9b3194dec..d50d37368 100644 --- a/interp/declare.mli +++ b/interp/declare.mli @@ -42,7 +42,7 @@ type internal_flag = (* Defaut definition entries, transparent with no secctx or proj information *) val definition_entry : ?fix_exn:Future.fix_exn -> ?opaque:bool -> ?inline:bool -> ?types:types -> - ?poly:polymorphic -> ?univs:Univ.UContext.t -> + ?univs:Entries.constant_universes_entry -> ?eff:Safe_typing.private_constants -> constr -> Safe_typing.private_constants definition_entry (** [declare_constant id cd] declares a global declaration @@ -56,8 +56,8 @@ val declare_constant : val declare_definition : ?internal:internal_flag -> ?opaque:bool -> ?kind:definition_object_kind -> - ?local:bool -> ?poly:polymorphic -> Id.t -> ?types:constr -> - constr Univ.in_universe_context_set -> Constant.t + ?local:bool -> Id.t -> ?types:constr -> + constr Entries.in_constant_universes_entry -> Constant.t (** Since transparent constants' side effects are globally declared, we * need that *) diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index 076fba861..7a9bbd92c 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -348,8 +348,11 @@ let generate_functional_principle (evd: Evd.evar_map ref) let evd',value = change_property_sort evd' s new_principle_type new_princ_name in let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr value)) in (* Pp.msgnl (str "new principle := " ++ pr_lconstr value); *) - let univs = Evd.to_universe_context evd' in - let ce = Declare.definition_entry ~poly:(Flags.is_universe_polymorphism ()) ~univs value in + let univs = + let poly = Flags.is_universe_polymorphism () in + Evd.const_univ_entry ~poly evd' + in + let ce = Declare.definition_entry ~univs value in ignore( Declare.declare_constant name diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 320f7c7f3..b0a76137b 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -66,8 +66,8 @@ let find_reference sl s = let dp = Names.DirPath.make (List.rev_map Id.of_string sl) in locate (make_qualid dp (Id.of_string s)) -let declare_fun f_id kind ?(ctx=Univ.UContext.empty) value = - let ce = definition_entry ~univs:ctx value (*FIXME *) in +let declare_fun f_id kind ?univs value = + let ce = definition_entry ?univs value (*FIXME *) in ConstRef(declare_constant f_id (DefinitionEntry ce, kind));; let defined () = Lemmas.save_proof (Vernacexpr.(Proved (Transparent,None))) @@ -1556,8 +1556,8 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num let functional_id = add_suffix function_name "_F" in let term_id = add_suffix function_name "_terminate" in let functional_ref = - let ctx = Evd.to_universe_context evm in - declare_fun functional_id (IsDefinition Decl_kinds.Definition) ~ctx res + let univs = Entries.Monomorphic_const_entry (Evd.universe_context_set evm) in + declare_fun functional_id (IsDefinition Decl_kinds.Definition) ~univs res in (* Refresh the global universes, now including those of _F *) let evm = Evd.from_env (Global.env ()) in diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index d6cfa3cf9..750517e91 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -313,10 +313,10 @@ let project_hint pri l2r r = let id = Nameops.add_suffix (Nametab.basename_of_global gr) ("_proj_" ^ (if l2r then "l2r" else "r2l")) in - let ctx = Evd.universe_context_set sigma in - let c = EConstr.to_constr sigma c in let poly = Flags.use_polymorphic_flag () in - let c = Declare.declare_definition ~poly ~internal:Declare.InternalTacticRequest id (c,ctx) in + let ctx = Evd.const_univ_entry ~poly sigma in + let c = EConstr.to_constr sigma c in + let c = Declare.declare_definition ~internal:Declare.InternalTacticRequest id (c,ctx) in let info = {Vernacexpr.hint_priority = pri; hint_pattern = None} in (info,false,true,Hints.PathAny, Hints.IsGlobRef (Globnames.ConstRef c)) diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index a2f5a4577..c0060c5a7 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -1884,11 +1884,11 @@ let declare_projection n instance_id r = in it_mkProd_or_LetIn ccl ctx in let typ = it_mkProd_or_LetIn typ ctx in - let ctx = Evd.to_universe_context sigma in + let univs = Evd.const_univ_entry ~poly sigma in let typ = EConstr.to_constr sigma typ in let term = EConstr.to_constr sigma term in let cst = - Declare.definition_entry ~types:typ ~poly ~univs:ctx term + Declare.definition_entry ~types:typ ~univs term in ignore(Declare.declare_constant n (Entries.DefinitionEntry cst, Decl_kinds.IsDefinition Decl_kinds.Definition)) @@ -1972,10 +1972,7 @@ let add_morphism_infer glob m n = let evd = Evd.from_env env in let uctx, instance = build_morphism_signature env evd m in if Lib.is_modtype () then - let uctx = if poly - then Entries.Polymorphic_const_entry (UState.context uctx) - else Entries.Monomorphic_const_entry (UState.context_set uctx) - in + let uctx = UState.const_univ_entry ~poly uctx in let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest instance_id (Entries.ParameterEntry (None,(instance,uctx),None), diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 6d0b6fd09..f22f00839 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -150,13 +150,13 @@ let ic_unsafe c = (*FIXME remove *) let sigma = Evd.from_env env in EConstr.of_constr (fst (Constrintern.interp_constr env sigma c)) -let decl_constant na ctx c = +let decl_constant na univs c = let open Constr in let vars = Univops.universes_of_constr c in - let ctx = Univops.restrict_universe_context (Univ.ContextSet.of_context ctx) vars in + let univs = Univops.restrict_universe_context univs vars in + let univs = Monomorphic_const_entry univs in mkConst(declare_constant (Id.of_string na) - (DefinitionEntry (definition_entry ~opaque:true - ~univs:(Univ.ContextSet.to_context ctx) c), + (DefinitionEntry (definition_entry ~opaque:true ~univs c), IsProof Lemma)) (* Calling a global tactic *) @@ -220,7 +220,7 @@ let exec_tactic env evd n f args = let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic_ist ist (ltac_call f (args@[getter]))) gl in let evd, nf = Evarutil.nf_evars_and_universes (Refiner.project gls) in let nf c = nf (constr_of c) in - Array.map nf !tactic_res, Evd.to_universe_context evd + Array.map nf !tactic_res, Evd.universe_context_set evd let stdlib_modules = [["Coq";"Setoids";"Setoid"]; diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index 905173efc..bfd64e21b 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -316,10 +316,6 @@ let get_open_goals () = (List.map (fun (l1,l2) -> List.length l1 + List.length l2) gll) + List.length shelf -let constrain_variables init uctx = - let levels = Univ.Instance.levels (Univ.UContext.instance init) in - UState.constrain_variables levels uctx - type closed_proof_output = (Constr.t * Safe_typing.private_constants) list * UState.t let close_proof ~keep_body_ucst_separate ?feedback_id ~now @@ -329,9 +325,12 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now let poly = pi2 strength (* Polymorphic *) in let initial_goals = Proof.initial_goals proof in let initial_euctx = Proof.initial_euctx proof in + let constrain_variables ctx = + UState.constrain_variables (fst (UState.context_set initial_euctx)) ctx + in let fpl, univs = Future.split2 fpl in let universes = if poly || now then Future.force univs else initial_euctx in - let univctx = UState.check_univ_decl universes universe_decl in + let univctx = UState.check_univ_decl ~poly universes universe_decl in let binders = if poly then Some (UState.universe_binders universes) else None in (* Because of dependent subgoals at the beginning of proofs, we could have existential variables in the initial types of goals, we need to @@ -353,15 +352,15 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now let used_univs_typ = Univops.universes_of_constr typ in if keep_body_ucst_separate || not (Safe_typing.empty_private_constants = eff) then - let initunivs = UState.context initial_euctx in - let ctx = constrain_variables initunivs universes in + let initunivs = UState.const_univ_entry ~poly initial_euctx in + let ctx = constrain_variables universes in (* For vi2vo compilation proofs are computed now but we need to complement the univ constraints of the typ with the ones of the body. So we keep the two sets distinct. *) let used_univs = Univ.LSet.union used_univs_body used_univs_typ in let ctx_body = UState.restrict ctx used_univs in - let univs = UState.check_univ_decl ctx_body universe_decl in - (initunivs, typ), ((body, Univ.ContextSet.of_context univs), eff) + let univs = UState.check_mono_univ_decl ctx_body universe_decl in + (initunivs, typ), ((body, univs), eff) else (* Since the proof is computed now, we can simply have 1 set of constraints in which we merge the ones for the body and the ones @@ -370,7 +369,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now TODO: check if restrict is really necessary now. *) let used_univs = Univ.LSet.union used_univs_body used_univs_typ in let ctx = UState.restrict universes used_univs in - let univs = UState.check_univ_decl ctx universe_decl in + let univs = UState.check_univ_decl ~poly ctx universe_decl in (univs, typ), ((body, Univ.ContextSet.empty), eff) in fun t p -> Future.split2 (Future.chain p (make_body t)) @@ -382,20 +381,14 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now the initial universes, ensure that the final universes respect the declaration as well. If the declaration is non-extensible, this will prevent the body from adding universes and constraints. *) - let bodyunivs = constrain_variables univctx (Future.force univs) in - let univs = UState.check_univ_decl bodyunivs universe_decl in - (pt,Univ.ContextSet.of_context univs),eff) + let bodyunivs = constrain_variables (Future.force univs) in + let univs = UState.check_mono_univ_decl bodyunivs universe_decl in + (pt,univs),eff) in let entry_fn p (_, t) = let t = EConstr.Unsafe.to_constr t in let univstyp, body = make_body t p in let univs, typ = Future.force univstyp in - let univs = - if poly then Entries.Polymorphic_const_entry univs - else - (* FIXME be smarter about the unnecessary context linearisation in make_body *) - Entries.Monomorphic_const_entry (Univ.ContextSet.of_context univs) - in {Entries. const_entry_body = body; const_entry_secctx = section_vars; diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 1ee873e0f..1ae3577ed 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -232,12 +232,15 @@ let inversion_scheme env sigma t sort dep_option inv_op = let invProof = it_mkNamedLambda_or_LetIn c !ownSign in let invProof = EConstr.Unsafe.to_constr invProof in let p = Evarutil.nf_evars_universes sigma invProof in - p, Evd.to_universe_context sigma + p, sigma let add_inversion_lemma name env sigma t sort dep inv_op = - let invProof, univs = inversion_scheme env sigma t sort dep inv_op in - let entry = definition_entry ~poly:(Flags.use_polymorphic_flag ()) - ~univs invProof in + let invProof, sigma = inversion_scheme env sigma t sort dep inv_op in + let univs = + let poly = Flags.use_polymorphic_flag () in + Evd.const_univ_entry ~poly sigma + in + let entry = definition_entry ~univs invProof in let _ = declare_constant name (DefinitionEntry entry, IsProof Lemma) in () diff --git a/vernac/class.ml b/vernac/class.ml index 68fd22cb6..943da8fa8 100644 --- a/vernac/class.ml +++ b/vernac/class.ml @@ -212,10 +212,10 @@ let build_id_coercion idf_opt source poly = Id.of_string ("Id_"^(ident_key_of_class source)^"_"^ (ident_key_of_class cl)) in - let univs = Evd.to_universe_context sigma in + let univs = Evd.const_univ_entry ~poly sigma in let constr_entry = (* Cast is necessary to express [val_f] is identity *) DefinitionEntry - (definition_entry ~types:typ_f ~poly ~univs + (definition_entry ~types:typ_f ~univs ~inline:true (mkCast (val_f, DEFAULTcast, typ_f))) in let decl = (constr_entry, IsDefinition IdentityCoercion) in diff --git a/vernac/classes.ml b/vernac/classes.ml index c99eba2cc..1833b7a1d 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -119,9 +119,9 @@ let declare_instance_constant k info global imps ?hook id decl poly evm term ter (Univops.universes_of_constr term) in Evd.restrict_universe_context evm levels in - let uctx = Evd.check_univ_decl evm decl in + let uctx = Evd.check_univ_decl ~poly evm decl in let entry = - Declare.definition_entry ~types:termtype ~poly ~univs:uctx term + Declare.definition_entry ~types:termtype ~univs:uctx term in let cdecl = (DefinitionEntry entry, kind) in let kn = Declare.declare_constant id cdecl in @@ -203,16 +203,10 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) nf t in Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr termtype); - let ctx = Evd.check_univ_decl !evars decl in - let ctx = if poly - then Polymorphic_const_entry ctx - else - (* FIXME be smarter about this *) - Monomorphic_const_entry (Univ.ContextSet.of_context ctx) - in + let univs = Evd.check_univ_decl ~poly !evars decl in let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id (ParameterEntry - (None,(termtype,ctx),None), Decl_kinds.IsAssumption Decl_kinds.Logical) + (None,(termtype,univs),None), Decl_kinds.IsAssumption Decl_kinds.Logical) in Universes.register_universe_binders (ConstRef cst) (Evd.universe_binders !evars); instance_hook k pri global imps ?hook (ConstRef cst); id @@ -391,19 +385,16 @@ let context poly l = let fn status (id, b, t) = if Lib.is_modtype () && not (Lib.sections_are_opened ()) then (* Declare the universe context once *) - let univs = !uctx in + let univs = if poly + then Polymorphic_const_entry (Univ.ContextSet.to_context !uctx) + else Monomorphic_const_entry !uctx + in let () = uctx := Univ.ContextSet.empty in let decl = match b with | None -> - let univs = if poly - then Polymorphic_const_entry (Univ.ContextSet.to_context univs) - else Monomorphic_const_entry univs - in (ParameterEntry (None,(t,univs),None), IsAssumption Logical) | Some b -> - (* FIXME be smarter about this *) - let univs = Univ.ContextSet.to_context univs in - let entry = Declare.definition_entry ~poly ~univs ~types:t b in + let entry = Declare.definition_entry ~univs ~types:t b in (DefinitionEntry entry, IsAssumption Logical) in let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in @@ -426,9 +417,12 @@ let context poly l = pi3 (Command.declare_assumption false decl (t, !uctx) Universes.empty_binders [] impl Vernacexpr.NoInline (Loc.tag id)) | Some b -> - let ctx = Univ.ContextSet.to_context !uctx in + let univs = if poly + then Polymorphic_const_entry (Univ.ContextSet.to_context !uctx) + else Monomorphic_const_entry !uctx + in let decl = (Discharge, poly, Definition) in - let entry = Declare.definition_entry ~poly ~univs:ctx ~types:t b in + let entry = Declare.definition_entry ~univs ~types:t b in let hook = Lemmas.mk_hook (fun _ gr -> gr) in let _ = DeclareDef.declare_definition id decl entry Universes.empty_binders [] hook in Lib.sections_are_opened () || Lib.is_modtype_strict () diff --git a/vernac/command.ml b/vernac/command.ml index 06d0c8470..1a8b16a89 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -88,7 +88,7 @@ let warn_implicits_in_term = strbrk "Implicit arguments declaration relies on type." ++ spc () ++ strbrk "The term declares more implicits than the type here.") -let interp_definition pl bl p red_option c ctypopt = +let interp_definition pl bl poly red_option c ctypopt = let env = Global.env() in let evd, decl = Univdecls.interp_univ_decl_opt env pl in let evdref = ref evd in @@ -105,15 +105,15 @@ let interp_definition pl bl p red_option c ctypopt = let c = EConstr.Unsafe.to_constr c in let nf,subst = Evarutil.e_nf_evars_and_universes evdref in let body = nf (it_mkLambda_or_LetIn c ctx) in - let vars = Univops.universes_of_constr body in - let evd = Evd.restrict_universe_context !evdref vars in - let uctx = Evd.check_univ_decl evd decl in + let vars = Univops.universes_of_constr body in + let evd = Evd.restrict_universe_context !evdref vars in + let uctx = Evd.check_univ_decl ~poly evd decl in let binders = Evd.universe_binders evd in imps1@(Impargs.lift_implicits nb_args imps2), binders, - definition_entry ~univs:uctx ~poly:p body + definition_entry ~univs:uctx body | Some ctyp -> - let ty, impsty = interp_type_evars_impls ~impls env_bl evdref ctyp in - let subst = evd_comb0 Evd.nf_univ_variables evdref in + let ty, impsty = interp_type_evars_impls ~impls env_bl evdref ctyp in + let subst = evd_comb0 Evd.nf_univ_variables evdref in let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in let env_bl = push_rel_context ctx env in let c, imps2 = interp_casted_constr_evars_impls ~impls env_bl evdref c ty in @@ -134,10 +134,10 @@ let interp_definition pl bl p red_option c ctypopt = let vars = Univ.LSet.union (Univops.universes_of_constr body) (Univops.universes_of_constr typ) in let ctx = Evd.restrict_universe_context !evdref vars in - let uctx = Evd.check_univ_decl ctx decl in + let uctx = Evd.check_univ_decl ~poly ctx decl in let binders = Evd.universe_binders evd in imps1@(Impargs.lift_implicits nb_args impsty), binders, - definition_entry ~types:typ ~poly:p + definition_entry ~types:typ ~univs:uctx body in red_constant_entry (Context.Rel.length ctx) ce !evdref red_option, !evdref, decl, pl, imps @@ -282,9 +282,14 @@ let do_assumptions_bound_univs coe kind nl id pl c = let ty = nf ty in let vars = Univops.universes_of_constr ty in let evd = Evd.restrict_universe_context !evdref vars in - let uctx = Evd.check_univ_decl evd decl in + let uctx = Evd.check_univ_decl ~poly:(pi2 kind) evd decl in let binders = Evd.universe_binders evd in - let uctx = Univ.ContextSet.of_context uctx in + let uctx = match uctx with + | Polymorphic_const_entry uctx -> + (* ?? *) + Univ.ContextSet.of_context uctx + | Monomorphic_const_entry uctx -> uctx + in let (_, _, st) = declare_assumption coe kind (ty, uctx) binders impls false nl id in st @@ -582,7 +587,7 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite = let constructors = List.map (fun (idl,cl,impsl) -> (idl,List.map nf' cl,impsl)) constructors in let ctx_params = Context.Rel.map nf ctx_params in let evd = !evdref in - let uctx = Evd.check_univ_decl evd decl in + let uctx = Evd.check_univ_decl ~poly evd decl in List.iter (fun c -> check_evars env_params Evd.empty evd (EConstr.of_constr c)) arities; Context.Rel.iter (fun c -> check_evars env0 Evd.empty evd (EConstr.of_constr c)) ctx_params; List.iter (fun (_,ctyps,_) -> @@ -604,12 +609,13 @@ let interp_mutual_inductive (paramsl,indl) notations cum poly prv finite = userimpls @ (lift_implicits len impls)) cimpls) indimpls constructors in let univs = - if poly then + match uctx with + | Polymorphic_const_entry uctx -> if cum then Cumulative_ind_entry (Universes.univ_inf_ind_from_universe_context uctx) else Polymorphic_ind_entry uctx - else - Monomorphic_ind_entry (Univ.ContextSet.of_context uctx) + | Monomorphic_const_entry uctx -> + Monomorphic_ind_entry uctx in (* Build the mutual inductive entry *) let mind_ent = @@ -1032,9 +1038,9 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation = let body = it_mkLambda_or_LetIn (mkApp (Evarutil.e_new_global evdref gr, [|make|])) binders_rel in let ty = it_mkProd_or_LetIn top_arity binders_rel in let ty = EConstr.Unsafe.to_constr ty in - let univs = Evd.check_univ_decl !evdref decl in + let univs = Evd.check_univ_decl ~poly !evdref decl in (*FIXME poly? *) - let ce = definition_entry ~poly ~types:ty ~univs (EConstr.to_constr !evdref body) in + let ce = definition_entry ~types:ty ~univs (EConstr.to_constr !evdref body) in (** FIXME: include locality *) let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in let gr = ConstRef c in @@ -1173,7 +1179,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in let evd = Evd.from_ctx ctx in let evd = Evd.restrict_universe_context evd vars in - let ctx = Evd.check_univ_decl evd pl in + let ctx = Evd.check_univ_decl ~poly evd pl in let pl = Evd.universe_binders evd in let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in ignore (List.map4 (DeclareDef.declare_fix (local, poly, Fixpoint) pl ctx) @@ -1206,7 +1212,7 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in let evd = Evd.from_ctx ctx in let evd = Evd.restrict_universe_context evd vars in - let ctx = Evd.check_univ_decl evd pl in + let ctx = Evd.check_univ_decl ~poly evd pl in let pl = Evd.universe_binders evd in ignore (List.map4 (DeclareDef.declare_fix (local, poly, CoFixpoint) pl ctx) fixnames fixdecls fixtypes fiximps); diff --git a/vernac/declareDef.ml b/vernac/declareDef.ml index c18744052..980db4109 100644 --- a/vernac/declareDef.ml +++ b/vernac/declareDef.ml @@ -57,7 +57,7 @@ let declare_definition ident (local, p, k) ce pl imps hook = declare_global_definition ident ce local k pl imps in Lemmas.call_hook fix_exn hook local r -let declare_fix ?(opaque = false) (_,poly,_ as kind) pl ctx f ((def,_),eff) t imps = - let ce = definition_entry ~opaque ~types:t ~poly ~univs:ctx ~eff def in +let declare_fix ?(opaque = false) (_,poly,_ as kind) pl univs f ((def,_),eff) t imps = + let ce = definition_entry ~opaque ~types:t ~univs ~eff def in declare_definition f kind ce pl imps (Lemmas.mk_hook (fun _ r -> r)) diff --git a/vernac/declareDef.mli b/vernac/declareDef.mli index 01a87818a..55f7c7861 100644 --- a/vernac/declareDef.mli +++ b/vernac/declareDef.mli @@ -15,5 +15,8 @@ val declare_definition : Id.t -> definition_kind -> Safe_typing.private_constants Entries.definition_entry -> Universes.universe_binders -> Impargs.manual_implicits -> Globnames.global_reference Lemmas.declaration_hook -> Globnames.global_reference -val declare_fix : ?opaque:bool -> definition_kind -> Universes.universe_binders -> Univ.UContext.t -> Id.t -> - Safe_typing.private_constants Entries.proof_output -> Constr.types -> Impargs.manual_implicits -> Globnames.global_reference +val declare_fix : ?opaque:bool -> definition_kind -> + Universes.universe_binders -> Entries.constant_universes_entry -> + Id.t -> Safe_typing.private_constants Entries.proof_output -> + Constr.types -> Impargs.manual_implicits -> + Globnames.global_reference diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index f59b5fcae..a787ec6fa 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -224,7 +224,7 @@ let compute_proof_name locality = function let avoid = Id.Set.of_list (Proof_global.get_all_proof_names ()) in next_global_ident_away default_thm_id avoid -let save_remaining_recthms (locality,p,kind) norm ctx body opaq i (id,(t_i,(_,imps))) = +let save_remaining_recthms (locality,p,kind) norm univs body opaq i (id,(t_i,(_,imps))) = let t_i = norm t_i in match body with | None -> @@ -232,7 +232,13 @@ let save_remaining_recthms (locality,p,kind) norm ctx body opaq i (id,(t_i,(_,im | Discharge -> let impl = false in (* copy values from Vernacentries *) let k = IsAssumption Conjectural in - let c = SectionLocalAssum ((t_i,Univ.ContextSet.of_context ctx),p,impl) in + let univs = match univs with + | Polymorphic_const_entry univs -> + (* What is going on here? *) + Univ.ContextSet.of_context univs + | Monomorphic_const_entry univs -> univs + in + let c = SectionLocalAssum ((t_i, univs),p,impl) in let _ = declare_variable id (Lib.cwd(),c,k) in (Discharge, VarRef id,imps) | Local | Global -> @@ -242,11 +248,7 @@ let save_remaining_recthms (locality,p,kind) norm ctx body opaq i (id,(t_i,(_,im | Global -> false | Discharge -> assert false in - let ctx = if p - then Polymorphic_const_entry ctx - else Monomorphic_const_entry (Univ.ContextSet.of_context ctx) - in - let decl = (ParameterEntry (None,(t_i,ctx),None), k) in + let decl = (ParameterEntry (None,(t_i,univs),None), k) in let kn = declare_constant id ~local decl in (locality,ConstRef kn,imps)) | Some body -> @@ -264,8 +266,7 @@ let save_remaining_recthms (locality,p,kind) norm ctx body opaq i (id,(t_i,(_,im let body_i = body_i body in match locality with | Discharge -> - let const = definition_entry ~types:t_i ~opaque:opaq ~poly:p - ~univs:ctx body_i in + let const = definition_entry ~types:t_i ~opaque:opaq ~univs body_i in let c = SectionLocalDef const in let _ = declare_variable id (Lib.cwd(), c, k) in (Discharge,VarRef id,imps) @@ -276,7 +277,7 @@ let save_remaining_recthms (locality,p,kind) norm ctx body opaq i (id,(t_i,(_,im | Discharge -> assert false in let const = - Declare.definition_entry ~types:t_i ~poly:p ~univs:ctx ~opaque:opaq body_i + Declare.definition_entry ~types:t_i ~univs ~opaque:opaq body_i in let kn = declare_constant id ~local (DefinitionEntry const, k) in (locality,ConstRef kn,imps) @@ -425,7 +426,7 @@ let start_proof_with_initialization kind ctx decl 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.check_univ_decl ctx decl in + let ctx = UState.check_univ_decl ~poly:(pi2 kind) 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 let thms_data = (strength,ref,imps)::other_thms_data in @@ -460,7 +461,7 @@ let start_proof_com ?inference_hook kind thms hook = let () = let open Misctypes in if not (decl.univdecl_extensible_instance && decl.univdecl_extensible_constraints) then - ignore (Evd.check_univ_decl evd decl) + ignore (Evd.check_univ_decl ~poly:(pi2 kind) evd decl) in let evd = if pi2 kind then evd @@ -495,10 +496,7 @@ let save_proof ?proof = function if const_entry_type = None then user_err Pp.(str "Admitted requires an explicit statement"); let typ = Option.get const_entry_type in - let ctx = if pi2 k (* polymorphic *) - then Polymorphic_const_entry (UState.context (fst universes)) - else Monomorphic_const_entry (UState.context_set (fst universes)) - in + let ctx = UState.const_univ_entry ~poly:(pi2 k) (fst universes) in let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in Admitted(id, k, (sec_vars, (typ, ctx), None), universes) | None -> @@ -521,12 +519,8 @@ let save_proof ?proof = function | _ -> None in let decl = Proof_global.get_universe_decl () in let evd = Evd.from_ctx universes in - let ctx = Evd.check_univ_decl evd decl in let poly = pi2 k in - let ctx = if poly - then Polymorphic_const_entry ctx - else Monomorphic_const_entry (Univ.ContextSet.of_context ctx) - in + let ctx = Evd.check_univ_decl ~poly evd decl in let binders = if poly then Some (UState.universe_binders universes) else None in Admitted(id,k,(sec_vars, (typ, ctx), None), (universes, binders)) diff --git a/vernac/obligations.ml b/vernac/obligations.ml index a9110c76c..97cdd7977 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -475,12 +475,8 @@ let declare_definition prg = (Evd.evar_universe_context_subst prg.prg_ctx) in let opaque = prg.prg_opaque in let fix_exn = Hook.get get_fix_exn () in - let () = ignore(UState.check_univ_decl prg.prg_ctx prg.prg_univdecl) in - let ce = - definition_entry ~fix_exn - ~opaque ~types:(nf typ) ~poly:(pi2 prg.prg_kind) - ~univs:(UState.context prg.prg_ctx) (nf body) - in + let univs = UState.check_univ_decl ~poly:(pi2 prg.prg_kind) prg.prg_ctx prg.prg_univdecl in + let ce = definition_entry ~fix_exn ~opaque ~types:(nf typ) ~univs (nf body) in let () = progmap_remove prg in let ubinders = UState.universe_binders prg.prg_ctx in DeclareDef.declare_definition prg.prg_name @@ -549,9 +545,9 @@ let declare_mutual_definition l = mk_proof (mkCoFix (i,fixdecls))) 0 l in (* Declare the recursive definitions *) - let ctx = UState.context first.prg_ctx in + let univs = UState.const_univ_entry ~poly first.prg_ctx in let fix_exn = Hook.get get_fix_exn () in - let kns = List.map4 (DeclareDef.declare_fix ~opaque (local, poly, kind) Universes.empty_binders ctx) + let kns = List.map4 (DeclareDef.declare_fix ~opaque (local, poly, kind) Universes.empty_binders univs) fixnames fixdecls fixtypes fiximps in (* Declare notations *) List.iter (Metasyntax.add_notation_interpretation (Global.env())) first.prg_notations; @@ -856,10 +852,7 @@ let obligation_terminator name num guard hook auto pf = | (_, status), Vernacexpr.Transparent -> status in let obl = { obl with obl_status = false, status } in - let uctx = if pi2 prg.prg_kind - then Polymorphic_const_entry (UState.context ctx) - else Monomorphic_const_entry (UState.context_set ctx) - in + let uctx = UState.const_univ_entry ~poly:(pi2 prg.prg_kind) ctx in let (_, obl) = declare_obligation prg obl body ty uctx in let obls = Array.copy obls in let _ = obls.(num) <- obl in diff --git a/vernac/record.ml b/vernac/record.ml index faf277d86..1d255b08e 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -95,7 +95,7 @@ let binder_of_decl = function let binders_of_decls = List.map binder_of_decl -let typecheck_params_and_fields finite def id pl t ps nots fs = +let typecheck_params_and_fields finite def id poly pl t ps nots fs = let env0 = Global.env () in let evd, decl = Univdecls.interp_univ_decl_opt env0 pl in let evars = ref evd in @@ -167,7 +167,7 @@ let typecheck_params_and_fields finite def id pl t ps nots fs = let newps = List.map (EConstr.to_rel_decl evars) newps in let typ = EConstr.to_constr evars typ in let ce t = Pretyping.check_evars env0 Evd.empty evars (EConstr.of_constr t) in - let univs = Evd.check_univ_decl evars decl in + let univs = Evd.check_univ_decl ~poly evars decl in let ubinders = Evd.universe_binders evars in List.iter (iter_constr ce) (List.rev newps); List.iter (iter_constr ce) (List.rev newfs); @@ -449,7 +449,7 @@ let implicits_of_context ctx = in ExplByPos (i, explname), (true, true, true)) 1 (List.rev (Anonymous :: (List.map RelDecl.get_name ctx))) -let declare_class finite def cum poly ubinders ctx id idbuild paramimpls params arity +let declare_class finite def cum ubinders univs id idbuild paramimpls params arity template fieldimpls fields ?(kind=StructureComponent) is_coe coers priorities sign = let fieldimpls = (* Make the class implicit in the projections, and the params if applicable. *) @@ -464,21 +464,21 @@ let declare_class finite def cum poly ubinders ctx id idbuild paramimpls params let class_body = it_mkLambda_or_LetIn field params in let class_type = it_mkProd_or_LetIn arity params in let class_entry = - Declare.definition_entry ~types:class_type ~poly ~univs:ctx class_body in + Declare.definition_entry ~types:class_type ~univs class_body in let cst = Declare.declare_constant (snd id) (DefinitionEntry class_entry, IsDefinition Definition) in - let cstu = (cst, if poly then Univ.UContext.instance ctx else Univ.Instance.empty) in + let cstu = (cst, match univs with + | Polymorphic_const_entry univs -> Univ.UContext.instance univs + | Monomorphic_const_entry _ -> Univ.Instance.empty) + in let inst_type = appvectc (mkConstU cstu) (Termops.rel_vect 0 (List.length params)) in let proj_type = it_mkProd_or_LetIn (mkProd(Name binder_name, inst_type, lift 1 field)) params in let proj_body = it_mkLambda_or_LetIn (mkLambda (Name binder_name, inst_type, mkRel 1)) params in - let proj_entry = - Declare.definition_entry ~types:proj_type ~poly - ~univs:(if poly then ctx else Univ.UContext.empty) proj_body - in + let proj_entry = Declare.definition_entry ~types:proj_type ~univs proj_body in let proj_cst = Declare.declare_constant proj_name (DefinitionEntry proj_entry, IsDefinition Definition) in @@ -495,13 +495,14 @@ let declare_class finite def cum poly ubinders ctx id idbuild paramimpls params cref, [Name proj_name, sub, Some proj_cst] | _ -> let univs = - if poly then + match univs with + | Polymorphic_const_entry univs -> if cum then - Cumulative_ind_entry (Universes.univ_inf_ind_from_universe_context ctx) + Cumulative_ind_entry (Universes.univ_inf_ind_from_universe_context univs) else - Polymorphic_ind_entry ctx - else - Monomorphic_ind_entry (Univ.ContextSet.of_context ctx) + Polymorphic_ind_entry univs + | Monomorphic_const_entry univs -> + Monomorphic_ind_entry univs in let ind = declare_structure BiFinite ubinders univs (snd id) idbuild paramimpls params arity template fieldimpls fields @@ -524,13 +525,15 @@ let declare_class finite def cum poly ubinders ctx id idbuild paramimpls params params, params in let univs, ctx_context, fields = - if poly then - let usubst, auctx = Univ.abstract_universes ctx in + match univs with + | Polymorphic_const_entry univs -> + let usubst, auctx = Univ.abstract_universes univs in let map c = Vars.subst_univs_level_constr usubst c in let fields = Context.Rel.map map fields in let ctx_context = on_snd (fun d -> Context.Rel.map map d) ctx_context in auctx, ctx_context, fields - else Univ.AUContext.empty, ctx_context, fields + | Monomorphic_const_entry _ -> + Univ.AUContext.empty, ctx_context, fields in let k = { cl_univs = univs; @@ -606,14 +609,14 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf if isnot_class && List.exists (fun opt -> not (Option.is_empty opt)) priorities then user_err Pp.(str "Priorities only allowed for type class substructures"); (* Now, younger decl in params and fields is on top *) - let pl, ctx, arity, template, implpars, params, implfs, fields = + let pl, univs, arity, template, implpars, params, implfs, fields = States.with_state_protection (fun () -> - typecheck_params_and_fields finite (kind = Class true) idstruc pl s ps notations fs) () in + typecheck_params_and_fields finite (kind = Class true) idstruc poly pl s ps notations fs) () in let sign = structure_signature (fields@params) in match kind with | Class def -> let priorities = List.map (fun id -> {hint_priority = id; hint_pattern = None}) priorities in - let gr = declare_class finite def cum poly pl ctx (loc,idstruc) idbuild + let gr = declare_class finite def cum pl univs (loc,idstruc) idbuild implpars params arity template implfs fields is_coe coers priorities sign in gr | _ -> @@ -622,13 +625,14 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf (succ (List.length params)) impls) implfs in let univs = - if poly then + match univs with + | Polymorphic_const_entry univs -> if cum then - Cumulative_ind_entry (Universes.univ_inf_ind_from_universe_context ctx) + Cumulative_ind_entry (Universes.univ_inf_ind_from_universe_context univs) else - Polymorphic_ind_entry ctx - else - Monomorphic_ind_entry (Univ.ContextSet.of_context ctx) + Polymorphic_ind_entry univs + | Monomorphic_const_entry univs -> + Monomorphic_ind_entry univs in let ind = declare_structure finite pl univs idstruc idbuild implpars params arity template implfs -- cgit v1.2.3 From e6c87412d70b71daaf417bd4b8e4ae6f1f28515b Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Fri, 24 Nov 2017 18:32:35 +0100 Subject: Fix obligations handling of universes anticipating stronger restrict --- vernac/obligations.ml | 94 +++++++++++++++++++++++++++++---------------------- 1 file changed, 54 insertions(+), 40 deletions(-) (limited to 'vernac') diff --git a/vernac/obligations.ml b/vernac/obligations.ml index 97cdd7977..17a596c42 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -824,49 +824,63 @@ let obligation_terminator name num guard hook auto pf = match pf with | Admitted _ -> apply_terminator term pf | Proved (opq, id, proof) -> - if not !shrink_obligations then apply_terminator term pf - else - let (_, (entry, uctx, _)) = Pfedit.cook_this_proof proof in - let env = Global.env () in - let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in - let ty = entry.Entries.const_entry_type in - let (body, cstr), () = Future.force entry.Entries.const_entry_body in - let sigma = Evd.from_ctx (fst uctx) in - let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in - Inductiveops.control_only_guard (Global.env ()) body; - (** Declare the obligation ourselves and drop the hook *) - let prg = get_info (ProgMap.find name !from_prg) in - (** Ensure universes are substituted properly in body and type *) - let body = EConstr.to_constr sigma (EConstr.of_constr body) in - let ty = Option.map (fun x -> EConstr.to_constr sigma (EConstr.of_constr x)) ty in - let ctx = Evd.evar_universe_context sigma in - let prg = { prg with prg_ctx = ctx } in - let obls, rem = prg.prg_obligations in - let obl = obls.(num) in - let status = - match obl.obl_status, opq with - | (_, Evar_kinds.Expand), Vernacexpr.Opaque -> err_not_transp () - | (true, _), Vernacexpr.Opaque -> err_not_transp () - | (false, _), Vernacexpr.Opaque -> Evar_kinds.Define true - | (_, Evar_kinds.Define true), Vernacexpr.Transparent -> Evar_kinds.Define false - | (_, status), Vernacexpr.Transparent -> status - in - let obl = { obl with obl_status = false, status } in - let uctx = UState.const_univ_entry ~poly:(pi2 prg.prg_kind) ctx in - let (_, obl) = declare_obligation prg obl body ty uctx in - let obls = Array.copy obls in - let _ = obls.(num) <- obl in - try + let (_, (entry, uctx, _)) = Pfedit.cook_this_proof proof in + let env = Global.env () in + let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in + let ty = entry.Entries.const_entry_type in + let (body, cstr), () = Future.force entry.Entries.const_entry_body in + let sigma = Evd.from_ctx (fst uctx) in + let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in + Inductiveops.control_only_guard (Global.env ()) body; + (** Declare the obligation ourselves and drop the hook *) + let prg = get_info (ProgMap.find name !from_prg) in + (** Ensure universes are substituted properly in body and type *) + let body = EConstr.to_constr sigma (EConstr.of_constr body) in + let ty = Option.map (fun x -> EConstr.to_constr sigma (EConstr.of_constr x)) ty in + let ctx = Evd.evar_universe_context sigma in + let obls, rem = prg.prg_obligations in + let obl = obls.(num) in + let status = + match obl.obl_status, opq with + | (_, Evar_kinds.Expand), Vernacexpr.Opaque -> err_not_transp () + | (true, _), Vernacexpr.Opaque -> err_not_transp () + | (false, _), Vernacexpr.Opaque -> Evar_kinds.Define true + | (_, Evar_kinds.Define true), Vernacexpr.Transparent -> + Evar_kinds.Define false + | (_, status), Vernacexpr.Transparent -> status + in + let obl = { obl with obl_status = false, status } in + let ctx = + if pi2 prg.prg_kind then ctx + else UState.union prg.prg_ctx ctx + in + let uctx = UState.const_univ_entry ~poly:(pi2 prg.prg_kind) ctx in + let (_, obl) = declare_obligation prg obl body ty uctx in + let obls = Array.copy obls in + let _ = obls.(num) <- obl in + let prg_ctx = + if pi2 (prg.prg_kind) then (* Polymorphic *) + (** We merge the new universes and constraints of the + polymorphic obligation with the existing ones *) + UState.union prg.prg_ctx ctx + else + (** The first obligation declares the univs of the constant, + each subsequent obligation declares its own additional + universes and constraints if any *) + UState.make (Global.universes ()) + in + let prg = { prg with prg_ctx } in + try ignore (update_obls prg obls (pred rem)); if pred rem > 0 then begin - let deps = dependencies obls num in - if not (Int.Set.is_empty deps) then - ignore (auto (Some name) None deps) - end - with e when CErrors.noncritical e -> - let e = CErrors.push e in - pperror (CErrors.iprint (ExplainErr.process_vernac_interp_error e)) + let deps = dependencies obls num in + if not (Int.Set.is_empty deps) then + ignore (auto (Some name) None deps) + end + with e when CErrors.noncritical e -> + let e = CErrors.push e in + pperror (CErrors.iprint (ExplainErr.process_vernac_interp_error e)) let obligation_hook prg obl num auto ctx' _ gr = let obls, rem = prg.prg_obligations in -- cgit v1.2.3 From 280c922fc55b57c430cad721c83650a796a375fd Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Wed, 20 Sep 2017 19:45:43 +0200 Subject: Allow local universe renaming in Print. --- doc/refman/RefMan-oth.tex | 7 +++ engine/universes.ml | 15 +++++++ engine/universes.mli | 11 +++++ intf/vernacexpr.ml | 6 ++- parsing/g_vernac.ml4 | 11 +++-- printing/ppvernac.ml | 13 ++++-- printing/prettyp.ml | 93 +++++++++++++++++++++++---------------- printing/prettyp.mli | 10 +++-- printing/printmod.ml | 24 ++++++---- printing/printmod.mli | 4 +- test-suite/output/UnivBinders.out | 31 +++++++++++++ test-suite/output/UnivBinders.v | 19 ++++++++ vernac/vernacentries.ml | 11 ++--- 13 files changed, 189 insertions(+), 66 deletions(-) (limited to 'vernac') diff --git a/doc/refman/RefMan-oth.tex b/doc/refman/RefMan-oth.tex index 60cd8b73a..3ebeba178 100644 --- a/doc/refman/RefMan-oth.tex +++ b/doc/refman/RefMan-oth.tex @@ -10,6 +10,8 @@ defined object referred by {\qualid}. \begin{ErrMsgs} \item {\qualid} \errindex{not a defined object} +\item \errindex{Universe instance should have length} $n$. +\item \errindex{This object does not support universe names.} \end{ErrMsgs} \begin{Variants} @@ -27,6 +29,11 @@ constructor, abbreviation, \ldots), long name, type, implicit arguments and argument scopes. It does not print the body of definitions or proofs. +\item {\tt Print {\qualid}@\{names\}.}\\ +This locally renames the polymorphic universes of {\qualid}. +An underscore means the raw universe is printed. +This form can be used with {\tt Print Term} and {\tt About}. + %\item {\tt Print Proof {\qualid}.}\comindex{Print Proof}\\ %In case \qualid\ denotes an opaque theorem defined in a section, %it is stored on a special unprintable form and displayed as diff --git a/engine/universes.ml b/engine/universes.ml index 459c53002..194de2cee 100644 --- a/engine/universes.ml +++ b/engine/universes.ml @@ -35,6 +35,21 @@ let universe_binders_of_global ref : universe_binders = let register_universe_binders ref l = universe_binders_table := Refmap.add ref l !universe_binders_table +type univ_name_list = Name.t Loc.located list + +let universe_binders_with_opt_names ref levels = function + | None -> universe_binders_of_global ref + | Some udecl -> + if Int.equal(List.length levels) (List.length udecl) + then + List.fold_left2 (fun acc (_,na) lvl -> match na with + | Anonymous -> acc + | Name na -> Names.Id.Map.add na lvl acc) + empty_binders udecl levels + else + CErrors.user_err ~hdr:"universe_binders_with_opt_names" + Pp.(str "Universe instance should have length " ++ int (List.length levels)) + (* To disallow minimization to Set *) let set_minimization = ref true diff --git a/engine/universes.mli b/engine/universes.mli index 621ca5e84..1401c4ee8 100644 --- a/engine/universes.mli +++ b/engine/universes.mli @@ -28,6 +28,17 @@ val empty_binders : universe_binders val register_universe_binders : Globnames.global_reference -> universe_binders -> unit val universe_binders_of_global : Globnames.global_reference -> universe_binders +type univ_name_list = Name.t Loc.located list + +(** [universe_binders_with_opt_names ref u l] + + If [l] is [Some univs] return the universe binders naming the levels of [u] by [univs] (skipping Anonymous). + May error if the lengths mismatch. + + Otherwise return [universe_binders_of_global ref]. *) +val universe_binders_with_opt_names : Globnames.global_reference -> + Univ.Level.t list -> univ_name_list option -> universe_binders + (** The global universe counter *) val set_remote_new_univ_level : Level.t RemoteCounter.installer diff --git a/intf/vernacexpr.ml b/intf/vernacexpr.ml index 9aef4b131..96bcba5e8 100644 --- a/intf/vernacexpr.ml +++ b/intf/vernacexpr.ml @@ -40,6 +40,8 @@ type goal_reference = | NthGoal of int | GoalId of Id.t +type univ_name_list = Name.t Loc.located list + type printable = | PrintTables | PrintFullContext @@ -54,7 +56,7 @@ type printable = | PrintMLLoadPath | PrintMLModules | PrintDebugGC - | PrintName of reference or_by_notation + | PrintName of reference or_by_notation * univ_name_list option | PrintGraph | PrintClasses | PrintTypeClasses @@ -70,7 +72,7 @@ type printable = | PrintScopes | PrintScope of string | PrintVisibility of string option - | PrintAbout of reference or_by_notation * goal_selector option + | PrintAbout of reference or_by_notation * univ_name_list option * goal_selector option | PrintImplicit of reference or_by_notation | PrintAssumptions of bool * bool * reference or_by_notation | PrintStrategy of reference or_by_notation option diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index a01ea26af..0e585cff7 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -875,7 +875,7 @@ GEXTEND Gram (* Printing (careful factorization of entries) *) | IDENT "Print"; p = printable -> VernacPrint p - | IDENT "Print"; qid = smart_global -> VernacPrint (PrintName qid) + | IDENT "Print"; qid = smart_global; l = OPT univ_name_list -> VernacPrint (PrintName (qid,l)) | IDENT "Print"; IDENT "Module"; "Type"; qid = global -> VernacPrint (PrintModuleType qid) | IDENT "Print"; IDENT "Module"; qid = global -> @@ -940,8 +940,8 @@ GEXTEND Gram | IDENT "Check"; c = lconstr; "." -> fun g -> VernacCheckMayEval (None, g, c) (* Searching the environment *) - | IDENT "About"; qid = smart_global; "." -> - fun g -> VernacPrint (PrintAbout (qid,g)) + | IDENT "About"; qid = smart_global; l = OPT univ_name_list; "." -> + fun g -> VernacPrint (PrintAbout (qid,l,g)) | IDENT "SearchHead"; c = constr_pattern; l = in_or_out_modules; "." -> fun g -> VernacSearch (SearchHead c,g, l) | IDENT "SearchPattern"; c = constr_pattern; l = in_or_out_modules; "." -> @@ -960,7 +960,7 @@ GEXTEND Gram ] ] ; printable: - [ [ IDENT "Term"; qid = smart_global -> PrintName qid + [ [ IDENT "Term"; qid = smart_global; l = OPT univ_name_list -> PrintName (qid,l) | IDENT "All" -> PrintFullContext | IDENT "Section"; s = global -> PrintSectionContext s | IDENT "Grammar"; ent = IDENT -> @@ -1060,6 +1060,9 @@ GEXTEND Gram | -> ([],SearchOutside []) ] ] ; + univ_name_list: + [ [ "@{" ; l = LIST0 name; "}" -> l ] ] + ; END; GEXTEND Gram diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index e897b1938..1a74538aa 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -41,6 +41,11 @@ open Decl_kinds pr_glob_level l ++ spc () ++ Univ.pr_constraint_type d ++ spc () ++ pr_glob_level r + let pr_univ_name_list = function + | None -> mt () + | Some l -> + str "@{" ++ prlist_with_sep spc pr_lname l ++ str"}" + let pr_univdecl_instance l extensible = prlist_with_sep spc pr_lident l ++ (if extensible then str"+" else mt ()) @@ -488,8 +493,8 @@ open Decl_kinds else "Print Universes" in keyword cmd ++ pr_opt str fopt - | PrintName qid -> - keyword "Print" ++ spc() ++ pr_smart_global qid + | PrintName (qid,udecl) -> + keyword "Print" ++ spc() ++ pr_smart_global qid ++ pr_univ_name_list udecl | PrintModuleType qid -> keyword "Print Module Type" ++ spc() ++ pr_reference qid | PrintModule qid -> @@ -502,9 +507,9 @@ open Decl_kinds keyword "Print Scope" ++ spc() ++ str s | PrintVisibility s -> keyword "Print Visibility" ++ pr_opt str s - | PrintAbout (qid,gopt) -> + | PrintAbout (qid,l,gopt) -> pr_opt (fun g -> Proof_bullet.pr_goal_selector g ++ str ":"++ spc()) gopt - ++ keyword "About" ++ spc() ++ pr_smart_global qid + ++ keyword "About" ++ spc() ++ pr_smart_global qid ++ pr_univ_name_list l | PrintImplicit qid -> keyword "Print Implicit" ++ spc() ++ pr_smart_global qid (* spiwack: command printing all the axioms and section variables used in a diff --git a/printing/prettyp.ml b/printing/prettyp.ml index b4ac94103..602b7a496 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -32,8 +32,8 @@ open Context.Rel.Declaration module NamedDecl = Context.Named.Declaration type object_pr = { - print_inductive : MutInd.t -> Pp.t; - print_constant_with_infos : Constant.t -> Pp.t; + print_inductive : MutInd.t -> Universes.univ_name_list option -> Pp.t; + print_constant_with_infos : Constant.t -> Universes.univ_name_list option -> Pp.t; print_section_variable : env -> Evd.evar_map -> variable -> Pp.t; print_syntactic_def : env -> KerName.t -> Pp.t; print_module : bool -> ModPath.t -> Pp.t; @@ -68,7 +68,7 @@ let int_or_no n = if Int.equal n 0 then str "no" else int n let print_basename sp = pr_global (ConstRef sp) -let print_ref reduce ref = +let print_ref reduce ref udecl = let typ, ctx = Global.type_of_global_in_context (Global.env ()) ref in let typ = Vars.subst_instance_constr (Univ.AUContext.instance ctx) typ in let typ = EConstr.of_constr typ in @@ -81,7 +81,8 @@ let print_ref reduce ref = let inst = Univ.AUContext.instance univs in let univs = Univ.UContext.make (inst, Univ.AUContext.instantiate inst univs) in let env = Global.env () in - let bl = Universes.universe_binders_of_global ref in + let bl = Universes.universe_binders_with_opt_names ref + (Array.to_list (Univ.Instance.to_array inst)) udecl in let sigma = Evd.from_ctx (Evd.evar_universe_context_of_binders bl) in let inst = if Global.is_polymorphic ref then Printer.pr_universe_instance sigma univs @@ -150,7 +151,7 @@ let print_impargs ref = let has_impl = not (List.is_empty impl) in (* Need to reduce since implicits are computed with products flattened *) pr_infos_list - ([ print_ref (need_expansion (select_impargs_size 0 impl) ref) ref; + ([ print_ref (need_expansion (select_impargs_size 0 impl) ref) ref None; blankline ] @ (if has_impl then print_impargs_list (mt()) impl else [str "No implicit arguments"])) @@ -256,7 +257,7 @@ let print_name_infos ref = if need_expansion (select_impargs_size 0 impls) ref then (* Need to reduce since implicits are computed with products flattened *) [str "Expanded type for implicit arguments"; - print_ref true ref; blankline] + print_ref true ref None; blankline] else [] in print_polymorphism ref @ @@ -512,11 +513,11 @@ let assumptions_for_print lna = (*********************) (* *) -let gallina_print_inductive sp = +let gallina_print_inductive sp udecl = let env = Global.env() in let mib = Environ.lookup_mind sp env in let mipv = mib.mind_packets in - pr_mutual_inductive_body env sp mib ++ + pr_mutual_inductive_body env sp mib udecl ++ with_line_skip (print_primitive_record mib.mind_finite mipv mib.mind_record @ print_inductive_renames sp mipv @ @@ -545,7 +546,7 @@ let print_instance sigma cb = pr_universe_instance sigma univs else mt() -let print_constant with_values sep sp = +let print_constant with_values sep sp udecl = let cb = Global.lookup_constant sp in let val_0 = Global.body_of_constant_body cb in let typ = @@ -555,31 +556,34 @@ let print_constant with_values sep sp = let inst = Univ.AUContext.instance univs in Vars.subst_instance_constr inst cb.const_type in - let univs = + let univs, ulist = let open Entries in + let open Univ in let otab = Global.opaque_tables () in match cb.const_body with | Undef _ | Def _ -> begin match cb.const_universes with - | Monomorphic_const ctx -> Monomorphic_const_entry ctx + | Monomorphic_const ctx -> Monomorphic_const_entry ctx, [] | Polymorphic_const ctx -> - let inst = Univ.AUContext.instance ctx in - Polymorphic_const_entry (Univ.UContext.make (inst, Univ.AUContext.instantiate inst ctx)) + let inst = AUContext.instance ctx in + Polymorphic_const_entry (UContext.make (inst, AUContext.instantiate inst ctx)), + Array.to_list (Instance.to_array inst) end | OpaqueDef o -> let body_uctxs = Opaqueproof.force_constraints otab o in match cb.const_universes with | Monomorphic_const ctx -> - Monomorphic_const_entry (Univ.ContextSet.union body_uctxs ctx) + Monomorphic_const_entry (ContextSet.union body_uctxs ctx), [] | Polymorphic_const ctx -> - assert(Univ.ContextSet.is_empty body_uctxs); - let inst = Univ.AUContext.instance ctx in - Polymorphic_const_entry (Univ.UContext.make (inst, Univ.AUContext.instantiate inst ctx)) + assert(ContextSet.is_empty body_uctxs); + let inst = AUContext.instance ctx in + Polymorphic_const_entry (UContext.make (inst, AUContext.instantiate inst ctx)), + Array.to_list (Instance.to_array inst) in let ctx = Evd.evar_universe_context_of_binders - (Universes.universe_binders_of_global (ConstRef sp)) + (Universes.universe_binders_with_opt_names (ConstRef sp) ulist udecl) in let env = Global.env () and sigma = Evd.from_ctx ctx in let pr_ltype = pr_ltype_env env sigma in @@ -596,8 +600,8 @@ let print_constant with_values sep sp = (if with_values then print_typed_body env sigma (Some c,typ) else pr_ltype typ)++ Printer.pr_constant_universes sigma univs) -let gallina_print_constant_with_infos sp = - print_constant true " = " sp ++ +let gallina_print_constant_with_infos sp udecl = + print_constant true " = " sp udecl ++ with_line_skip (print_name_infos (ConstRef sp)) let gallina_print_syntactic_def env kn = @@ -622,9 +626,9 @@ let gallina_print_leaf_entry env sigma with_values ((sp,kn as oname),lobj) = constraints *) (try Some(print_named_decl env sigma (basename sp)) with Not_found -> None) | (_,"CONSTANT") -> - Some (print_constant with_values sep (Constant.make1 kn)) + Some (print_constant with_values sep (Constant.make1 kn) None) | (_,"INDUCTIVE") -> - Some (gallina_print_inductive (MutInd.make1 kn)) + Some (gallina_print_inductive (MutInd.make1 kn) None) | (_,"MODULE") -> let (mp,_,l) = KerName.repr kn in Some (print_module with_values (MPdot (mp,l))) @@ -745,7 +749,7 @@ let print_full_pure_context env sigma = | "INDUCTIVE" -> let mind = Global.mind_of_delta_kn kn in let mib = Global.lookup_mind mind in - pr_mutual_inductive_body (Global.env()) mind mib ++ + pr_mutual_inductive_body (Global.env()) mind mib None ++ str "." ++ fnl () ++ fnl () | "MODULE" -> (* TODO: make it reparsable *) @@ -792,10 +796,19 @@ let print_sec_context env sigma sec = let print_sec_context_typ env sigma sec = print_context env sigma false None (read_sec_context sec) -let print_any_name env sigma = function - | Term (ConstRef sp) -> print_constant_with_infos sp - | Term (IndRef (sp,_)) -> print_inductive sp - | Term (ConstructRef ((sp,_),_)) -> print_inductive sp +let maybe_error_reject_univ_decl na udecl = + match na, udecl with + | _, None | Term (ConstRef _ | IndRef _ | ConstructRef _), Some _ -> () + | (Term (VarRef _) | Syntactic _ | Dir _ | ModuleType _ | Other _ | Undefined _), Some udecl -> + (* TODO Print na somehow *) + user_err ~hdr:"reject_univ_decl" (str "This object does not support universe names.") + +let print_any_name env sigma na udecl = + maybe_error_reject_univ_decl na udecl; + match na with + | Term (ConstRef sp) -> print_constant_with_infos sp udecl + | Term (IndRef (sp,_)) -> print_inductive sp udecl + | Term (ConstructRef ((sp,_),_)) -> print_inductive sp udecl | Term (VarRef sp) -> print_section_variable env sigma sp | Syntactic kn -> print_syntactic_def env kn | Dir (DirModule(dirpath,(mp,_))) -> print_module (printable_body dirpath) mp @@ -812,24 +825,26 @@ let print_any_name env sigma = function user_err ~hdr:"print_name" (pr_qualid qid ++ spc () ++ str "not a defined object.") -let print_name env sigma = function +let print_name env sigma na udecl = + match na with | ByNotation (loc,(ntn,sc)) -> print_any_name env sigma (Term (Notation.interp_notation_as_global_reference ?loc (fun _ -> true) ntn sc)) + udecl | AN ref -> - print_any_name env sigma (locate_any_name ref) + print_any_name env sigma (locate_any_name ref) udecl let print_opaque_name env sigma qid = match Nametab.global qid with | ConstRef cst -> let cb = Global.lookup_constant cst in if Declareops.constant_has_body cb then - print_constant_with_infos cst + print_constant_with_infos cst None else user_err Pp.(str "Not a defined constant.") | IndRef (sp,_) -> - print_inductive sp + print_inductive sp None | ConstructRef cstr as gr -> let ty, ctx = Global.type_of_global_in_context env gr in let inst = Univ.AUContext.instance ctx in @@ -840,13 +855,14 @@ let print_opaque_name env sigma qid = | VarRef id -> env |> lookup_named id |> print_named_decl env sigma -let print_about_any ?loc env sigma k = +let print_about_any ?loc env sigma k udecl = + maybe_error_reject_univ_decl k udecl; match k with | Term ref -> let rb = Reductionops.ReductionBehaviour.print ref in Dumpglob.add_glob ?loc ref; pr_infos_list - (print_ref false ref :: blankline :: + (print_ref false ref udecl :: blankline :: print_name_infos ref @ (if Pp.ismt rb then [] else [rb]) @ print_opacity ref @ @@ -862,13 +878,14 @@ let print_about_any ?loc env sigma k = hov 0 (pr_located_qualid k) | Other (obj, info) -> hov 0 (info.about obj) -let print_about env sigma = function +let print_about env sigma na udecl = + match na with | ByNotation (loc,(ntn,sc)) -> print_about_any ?loc env sigma (Term (Notation.interp_notation_as_global_reference ?loc (fun _ -> true) - ntn sc)) + ntn sc)) udecl | AN ref -> - print_about_any ?loc:(loc_of_reference ref) env sigma (locate_any_name ref) + print_about_any ?loc:(loc_of_reference ref) env sigma (locate_any_name ref) udecl (* for debug *) let inspect env sigma depth = @@ -940,7 +957,7 @@ let print_canonical_projections env sigma = open Typeclasses let pr_typeclass env t = - print_ref false t.cl_impl + print_ref false t.cl_impl None let print_typeclasses () = let env = Global.env () in @@ -949,7 +966,7 @@ let print_typeclasses () = let pr_instance env i = (* gallina_print_constant_with_infos i.is_impl *) (* lighter *) - print_ref false (instance_impl i) ++ + print_ref false (instance_impl i) None ++ begin match hint_priority i with | None -> mt () | Some i -> spc () ++ str "|" ++ spc () ++ int i diff --git a/printing/prettyp.mli b/printing/prettyp.mli index 89099a043..8f3a6ddc4 100644 --- a/printing/prettyp.mli +++ b/printing/prettyp.mli @@ -31,9 +31,11 @@ val print_eval : reduction_function -> env -> Evd.evar_map -> Constrexpr.constr_expr -> EConstr.unsafe_judgment -> Pp.t -val print_name : env -> Evd.evar_map -> reference or_by_notation -> Pp.t +val print_name : env -> Evd.evar_map -> reference or_by_notation -> + Vernacexpr.univ_name_list option -> Pp.t val print_opaque_name : env -> Evd.evar_map -> reference -> Pp.t -val print_about : env -> Evd.evar_map -> reference or_by_notation -> Pp.t +val print_about : env -> Evd.evar_map -> reference or_by_notation -> + Vernacexpr.univ_name_list option -> Pp.t val print_impargs : reference or_by_notation -> Pp.t (** Pretty-printing functions for classes and coercions *) @@ -80,8 +82,8 @@ val print_located_module : reference -> Pp.t val print_located_other : string -> reference -> Pp.t type object_pr = { - print_inductive : MutInd.t -> Pp.t; - print_constant_with_infos : Constant.t -> Pp.t; + print_inductive : MutInd.t -> Universes.univ_name_list option -> Pp.t; + print_constant_with_infos : Constant.t -> Universes.univ_name_list option -> Pp.t; print_section_variable : env -> Evd.evar_map -> variable -> Pp.t; print_syntactic_def : env -> KerName.t -> Pp.t; print_module : bool -> ModPath.t -> Pp.t; diff --git a/printing/printmod.ml b/printing/printmod.ml index 13a03e9b4..c34543bbd 100644 --- a/printing/printmod.ml +++ b/printing/printmod.ml @@ -121,7 +121,7 @@ let instantiate_cumulativity_info cumi = in CumulativityInfo.make (expose univs, expose subtyp) -let print_mutual_inductive env mind mib = +let print_mutual_inductive env mind mib udecl = let inds = List.init (Array.length mib.mind_packets) (fun x -> (mind, x)) in let keyword = @@ -131,7 +131,14 @@ let print_mutual_inductive env mind mib = | BiFinite -> "Variant" | CoFinite -> "CoInductive" in - let bl = Universes.universe_binders_of_global (IndRef (mind, 0)) in + let univs = + let open Univ in + if Declareops.inductive_is_polymorphic mib then + Array.to_list (Instance.to_array + (AUContext.instance (Declareops.inductive_polymorphic_context mib))) + else [] + in + let bl = Universes.universe_binders_with_opt_names (IndRef (mind, 0)) univs udecl in let sigma = Evd.from_ctx (Evd.evar_universe_context_of_binders bl) in hov 0 (Printer.pr_polymorphic (Declareops.inductive_is_polymorphic mib) ++ Printer.pr_cumulative @@ -159,7 +166,7 @@ let get_fields = in prodec_rec [] [] -let print_record env mind mib = +let print_record env mind mib udecl = let u = if Declareops.inductive_is_polymorphic mib then Univ.AUContext.instance (Declareops.inductive_polymorphic_context mib) @@ -173,7 +180,8 @@ let print_record env mind mib = let cstrtype = hnf_prod_applist env cstrtypes.(0) args in let fields = get_fields cstrtype in let envpar = push_rel_context params env in - let bl = Universes.universe_binders_of_global (IndRef (mind,0)) in + let bl = Universes.universe_binders_with_opt_names (IndRef (mind,0)) + (Array.to_list (Univ.Instance.to_array u)) udecl in let sigma = Evd.from_ctx (Evd.evar_universe_context_of_binders bl) in let keyword = let open Decl_kinds in @@ -205,11 +213,11 @@ let print_record env mind mib = sigma (instantiate_cumulativity_info cumi) ) -let pr_mutual_inductive_body env mind mib = +let pr_mutual_inductive_body env mind mib udecl = if mib.mind_record <> None && not !Flags.raw_print then - print_record env mind mib + print_record env mind mib udecl else - print_mutual_inductive env mind mib + print_mutual_inductive env mind mib udecl (** Modpaths *) @@ -335,7 +343,7 @@ let print_body is_impl env mp (l,body) = | SFBmind mib -> try let env = Option.get env in - pr_mutual_inductive_body env (MutInd.make2 mp l) mib + pr_mutual_inductive_body env (MutInd.make2 mp l) mib None with e when CErrors.noncritical e -> let keyword = let open Decl_kinds in diff --git a/printing/printmod.mli b/printing/printmod.mli index b0bbb21e0..97ed063fe 100644 --- a/printing/printmod.mli +++ b/printing/printmod.mli @@ -11,6 +11,8 @@ open Names (** false iff the module is an element of an open module type *) val printable_body : DirPath.t -> bool -val pr_mutual_inductive_body : Environ.env -> MutInd.t -> Declarations.mutual_inductive_body -> Pp.t +val pr_mutual_inductive_body : Environ.env -> + MutInd.t -> Declarations.mutual_inductive_body -> + Vernacexpr.univ_name_list option -> Pp.t val print_module : bool -> ModPath.t -> Pp.t val print_modtype : ModPath.t -> Pp.t diff --git a/test-suite/output/UnivBinders.out b/test-suite/output/UnivBinders.out index 42fb52a3b..6fb8cba18 100644 --- a/test-suite/output/UnivBinders.out +++ b/test-suite/output/UnivBinders.out @@ -53,3 +53,34 @@ Monomorphic mono = Type@{u} (* {u} |= *) mono is not universe polymorphic +foo@{E M N} = +Type@{M} -> Type@{N} -> Type@{E} + : Type@{max(E+1, M+1, N+1)} +(* E M N |= *) + +foo is universe polymorphic +foo@{Top.16 Top.17 Top.18} = +Type@{Top.17} -> Type@{Top.18} -> Type@{Top.16} + : Type@{max(Top.16+1, Top.17+1, Top.18+1)} +(* Top.16 Top.17 Top.18 |= *) + +foo is universe polymorphic +NonCumulative Inductive Empty@{E} : Type@{E} := +NonCumulative Record PWrap (A : Type@{E}) : Type@{E} := pwrap { punwrap : A } + +PWrap has primitive projections with eta conversion. +For PWrap: Argument scope is [type_scope] +For pwrap: Argument scopes are [type_scope _] +punwrap@{K} : forall A : Type@{K}, PWrap@{K} A -> A +(* K |= *) + +punwrap is universe polymorphic +Argument scopes are [type_scope _] +punwrap is transparent +Expands to: Constant Top.punwrap +The command has indeed failed with message: +Universe instance should have length 3 +The command has indeed failed with message: +Universe instance should have length 0 +The command has indeed failed with message: +This object does not support universe names. diff --git a/test-suite/output/UnivBinders.v b/test-suite/output/UnivBinders.v index 2c378e795..46904b384 100644 --- a/test-suite/output/UnivBinders.v +++ b/test-suite/output/UnivBinders.v @@ -33,3 +33,22 @@ Print foo. (* Binders even work with monomorphic definitions! *) Monomorphic Definition mono@{u} := Type@{u}. Print mono. + +(* Using local binders for printing. *) +Print foo@{E M N}. +(* Underscores discard the name if there's one. *) +Print foo@{_ _ _}. + +(* Also works for inductives and records. *) +Print Empty@{E}. +Print PWrap@{E}. + +(* Also works for About. *) +About punwrap@{K}. + +(* Instance length check. *) +Fail Print foo@{E}. +Fail Print mono@{E}. + +(* Not everything can be printed with custom universe names. *) +Fail Print Coq.Init.Logic@{E}. diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index e2f28a371..22e3e9c99 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -1609,9 +1609,10 @@ exception NoHyp (* Printing "About" information of a hypothesis of the current goal. We only print the type and a small statement to this comes from the goal. Precondition: there must be at least one current goal. *) -let print_about_hyp_globs ?loc ref_or_by_not glopt = +let print_about_hyp_globs ?loc ref_or_by_not udecl glopt = let open Context.Named.Declaration in try + (* FIXME error on non None udecl if we find the hyp. *) let glnumopt = query_command_selector ?loc glopt in let gl,id = match glnumopt,ref_or_by_not with @@ -1634,7 +1635,7 @@ let print_about_hyp_globs ?loc ref_or_by_not glopt = with (* fallback to globals *) | NoHyp | Not_found -> let sigma, env = Pfedit.get_current_context () in - print_about env sigma ref_or_by_not + print_about env sigma ref_or_by_not udecl let vernac_print ?loc env sigma = let open Feedback in function @@ -1651,7 +1652,7 @@ let vernac_print ?loc env sigma = let open Feedback in function | PrintMLLoadPath -> msg_notice (Mltop.print_ml_path ()) | PrintMLModules -> msg_notice (Mltop.print_ml_modules ()) | PrintDebugGC -> msg_notice (Mltop.print_gc ()) - | PrintName qid -> dump_global qid; msg_notice (print_name env sigma qid) + | PrintName (qid,udecl) -> dump_global qid; msg_notice (print_name env sigma qid udecl) | PrintGraph -> msg_notice (Prettyp.print_graph()) | PrintClasses -> msg_notice (Prettyp.print_classes()) | PrintTypeClasses -> msg_notice (Prettyp.print_typeclasses()) @@ -1681,8 +1682,8 @@ let vernac_print ?loc env sigma = let open Feedback in function msg_notice (Notation.pr_scope (Constrextern.without_symbols (pr_lglob_constr_env env)) s) | PrintVisibility s -> msg_notice (Notation.pr_visibility (Constrextern.without_symbols (pr_lglob_constr_env env)) s) - | PrintAbout (ref_or_by_not,glnumopt) -> - msg_notice (print_about_hyp_globs ?loc ref_or_by_not glnumopt) + | PrintAbout (ref_or_by_not,udecl,glnumopt) -> + msg_notice (print_about_hyp_globs ?loc ref_or_by_not udecl glnumopt) | PrintImplicit qid -> dump_global qid; msg_notice (print_impargs qid) | PrintAssumptions (o,t,r) -> -- cgit v1.2.3 From 765392492df2f5e065b2b5e706b6620846337cc0 Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Thu, 21 Sep 2017 00:53:10 +0200 Subject: Universe binders survive sections, modules and compilation. --- engine/universes.ml | 30 +++++++++++++++++++++++++- test-suite/Makefile | 3 ++- test-suite/output/UnivBinders.out | 41 ++++++++++++++++++++++++++++++++++++ test-suite/output/UnivBinders.v | 27 ++++++++++++++++++++++++ test-suite/prerequisite/bind_univs.v | 5 +++++ vernac/lemmas.ml | 4 ++-- 6 files changed, 106 insertions(+), 4 deletions(-) create mode 100644 test-suite/prerequisite/bind_univs.v (limited to 'vernac') diff --git a/engine/universes.ml b/engine/universes.ml index 194de2cee..f2942be6d 100644 --- a/engine/universes.ml +++ b/engine/universes.ml @@ -32,9 +32,37 @@ let universe_binders_of_global ref : universe_binders = let l = Refmap.find ref !universe_binders_table in l with Not_found -> Names.Id.Map.empty -let register_universe_binders ref l = +let cache_ubinder (_,(ref,l)) = universe_binders_table := Refmap.add ref l !universe_binders_table +let subst_ubinder (subst,(ref,l as orig)) = + let ref' = fst (Globnames.subst_global subst ref) in + if ref == ref' then orig else ref', l + +let discharge_ubinder (_,(ref,l)) = + Some (Lib.discharge_global ref, l) + +let ubinder_obj : Globnames.global_reference * universe_binders -> Libobject.obj = + let open Libobject in + declare_object { (default_object "universe binder") with + cache_function = cache_ubinder; + load_function = (fun _ x -> cache_ubinder x); + classify_function = (fun x -> Substitute x); + subst_function = subst_ubinder; + discharge_function = discharge_ubinder; + rebuild_function = (fun x -> x); } + +let register_universe_binders ref ubinders = + (* Add the polymorphic (section) universes *) + let open Names in + let ubinders = Idmap.fold (fun id (poly,lvl) ubinders -> + if poly then Id.Map.add id lvl ubinders + else ubinders) + (fst (Global.global_universe_names ())) ubinders + in + if not (Id.Map.is_empty ubinders) + then Lib.add_anonymous_leaf (ubinder_obj (ref,ubinders)) + type univ_name_list = Name.t Loc.located list let universe_binders_with_opt_names ref levels = function diff --git a/test-suite/Makefile b/test-suite/Makefile index f169f86e8..6865dcc76 100644 --- a/test-suite/Makefile +++ b/test-suite/Makefile @@ -95,7 +95,8 @@ VSUBSYSTEMS := prerequisite success failure $(BUGS) output \ SUBSYSTEMS := $(VSUBSYSTEMS) misc bugs ide vio coqchk coqwc coq-makefile PREREQUISITELOG = prerequisite/admit.v.log \ - prerequisite/make_local.v.log prerequisite/make_notation.v.log + prerequisite/make_local.v.log prerequisite/make_notation.v.log \ + prerequisite/bind_univs.v.log ####################################################################### # Phony targets diff --git a/test-suite/output/UnivBinders.out b/test-suite/output/UnivBinders.out index 6fb8cba18..04bd169bd 100644 --- a/test-suite/output/UnivBinders.out +++ b/test-suite/output/UnivBinders.out @@ -84,3 +84,44 @@ The command has indeed failed with message: Universe instance should have length 0 The command has indeed failed with message: This object does not support universe names. +Monomorphic bind_univs.mono = Type@{u} + : Type@{u+1} +(* {u} |= *) + +bind_univs.mono is not universe polymorphic +bind_univs.poly@{u} = Type@{u} + : Type@{u+1} +(* u |= *) + +bind_univs.poly is universe polymorphic +insec@{v} = Type@{u} -> Type@{v} + : Type@{max(u+1, v+1)} +(* v |= *) + +insec is universe polymorphic +insec@{u v} = Type@{u} -> Type@{v} + : Type@{max(u+1, v+1)} +(* u v |= *) + +insec is universe polymorphic +inmod@{u} = Type@{u} + : Type@{u+1} +(* u |= *) + +inmod is universe polymorphic +SomeMod.inmod@{u} = Type@{u} + : Type@{u+1} +(* u |= *) + +SomeMod.inmod is universe polymorphic +inmod@{u} = Type@{u} + : Type@{u+1} +(* u |= *) + +inmod is universe polymorphic +Applied.infunct@{u v} = +inmod@{u} -> Type@{v} + : Type@{max(u+1, v+1)} +(* u v |= *) + +Applied.infunct is universe polymorphic diff --git a/test-suite/output/UnivBinders.v b/test-suite/output/UnivBinders.v index 46904b384..f0a990986 100644 --- a/test-suite/output/UnivBinders.v +++ b/test-suite/output/UnivBinders.v @@ -52,3 +52,30 @@ Fail Print mono@{E}. (* Not everything can be printed with custom universe names. *) Fail Print Coq.Init.Logic@{E}. + +(* Universe binders survive through compilation, sections and modules. *) +Require bind_univs. +Print bind_univs.mono. +Print bind_univs.poly. + +Section SomeSec. + Universe u. + Definition insec@{v} := Type@{u} -> Type@{v}. + Print insec. +End SomeSec. +Print insec. + +Module SomeMod. + Definition inmod@{u} := Type@{u}. + Print inmod. +End SomeMod. +Print SomeMod.inmod. +Import SomeMod. +Print inmod. + +Module Type SomeTyp. Definition inmod := Type. End SomeTyp. +Module SomeFunct (In : SomeTyp). + Definition infunct@{u v} := In.inmod@{u} -> Type@{v}. +End SomeFunct. +Module Applied := SomeFunct(SomeMod). +Print Applied.infunct. diff --git a/test-suite/prerequisite/bind_univs.v b/test-suite/prerequisite/bind_univs.v new file mode 100644 index 000000000..f17c00e9d --- /dev/null +++ b/test-suite/prerequisite/bind_univs.v @@ -0,0 +1,5 @@ +(* Used in output/UnivBinders.v *) + +Monomorphic Definition mono@{u} := Type@{u}. + +Polymorphic Definition poly@{u} := Type@{u}. diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index a787ec6fa..42631a15b 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -204,7 +204,7 @@ let save ?export_seff id const cstrs pl do_guard (locality,poly,kind) hook = (locality, ConstRef kn) in definition_message id; - Option.iter (Universes.register_universe_binders r) pl; + Universes.register_universe_binders r (Option.default Universes.empty_binders pl); call_hook (fun exn -> exn) hook l r with e when CErrors.noncritical e -> let e = CErrors.push e in @@ -312,7 +312,7 @@ let admit (id,k,e) pl hook () = | Local, _, _ | Discharge, _, _ -> warn_let_as_axiom id in let () = assumption_message id in - Option.iter (Universes.register_universe_binders (ConstRef kn)) pl; + Universes.register_universe_binders (ConstRef kn) (Option.default Universes.empty_binders pl); call_hook (fun exn -> exn) hook Global (ConstRef kn) (* Starting a goal *) -- cgit v1.2.3 From 4940f99922b0784d726b7c50abe63395713f012f Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Fri, 17 Nov 2017 23:48:21 +0100 Subject: Fix #5347: unify declaration of axioms with and without bound univs. Note that this makes the following syntax valid: Axiom foo@{i} bar : Type@{i}. (ie putting a universe declaration on the first axiom in the list, the declaration then holds for the whole list). --- test-suite/bugs/closed/5347.v | 10 +++ test-suite/output/UnivBinders.out | 24 +++++++ test-suite/output/UnivBinders.v | 13 ++++ vernac/classes.ml | 10 +-- vernac/command.ml | 145 +++++++++++++++++++------------------- vernac/command.mli | 2 +- 6 files changed, 124 insertions(+), 80 deletions(-) create mode 100644 test-suite/bugs/closed/5347.v (limited to 'vernac') diff --git a/test-suite/bugs/closed/5347.v b/test-suite/bugs/closed/5347.v new file mode 100644 index 000000000..9267b3eb6 --- /dev/null +++ b/test-suite/bugs/closed/5347.v @@ -0,0 +1,10 @@ +Set Universe Polymorphism. + +Axiom X : Type. +(* Used to declare [x0@{u1 u2} : X@{u1}] and [x1@{} : X@{u2}] leaving + the type of x1 with undeclared universes. After PR #891 this should + error at declaration time. *) +Axiom x₀ x₁ : X. +Axiom Xᵢ : X -> Type. + +Check Xᵢ x₁. (* conversion test raised anomaly universe undefined *) diff --git a/test-suite/output/UnivBinders.out b/test-suite/output/UnivBinders.out index 9eb162fc0..f69379a57 100644 --- a/test-suite/output/UnivBinders.out +++ b/test-suite/output/UnivBinders.out @@ -129,3 +129,27 @@ inmod@{u} -> Type@{v} (* u v |= *) Applied.infunct is universe polymorphic +axfoo@{i Top.33 Top.34} : Type@{Top.33} -> Type@{i} +(* i Top.33 Top.34 |= *) + +axfoo is universe polymorphic +Argument scope is [type_scope] +Expands to: Constant Top.axfoo +axbar@{i Top.33 Top.34} : Type@{Top.34} -> Type@{i} +(* i Top.33 Top.34 |= *) + +axbar is universe polymorphic +Argument scope is [type_scope] +Expands to: Constant Top.axbar +axfoo' : Type@{Top.36} -> Type@{i} + +axfoo' is not universe polymorphic +Argument scope is [type_scope] +Expands to: Constant Top.axfoo' +axbar' : Type@{Top.36} -> Type@{i} + +axbar' is not universe polymorphic +Argument scope is [type_scope] +Expands to: Constant Top.axbar' +The command has indeed failed with message: +When declaring multiple axioms in one command, only the first is allowed a universe binder (which will be shared by the whole block). diff --git a/test-suite/output/UnivBinders.v b/test-suite/output/UnivBinders.v index a65a1fdf3..116d5e5e9 100644 --- a/test-suite/output/UnivBinders.v +++ b/test-suite/output/UnivBinders.v @@ -86,3 +86,16 @@ Module SomeFunct (In : SomeTyp). End SomeFunct. Module Applied := SomeFunct(SomeMod). Print Applied.infunct. + +(* Multi-axiom declaration + + In polymorphic mode the domain Type gets separate universes for the + different axioms, but all axioms have to declare all universes. In + polymorphic mode they get the same universes, ie the type is only + interpd once. *) +Axiom axfoo@{i+} axbar : Type -> Type@{i}. +Monomorphic Axiom axfoo'@{i+} axbar' : Type -> Type@{i}. + +About axfoo. About axbar. About axfoo'. About axbar'. + +Fail Axiom failfoo failbar@{i} : Type. diff --git a/vernac/classes.ml b/vernac/classes.ml index 1833b7a1d..b80741269 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -412,15 +412,15 @@ let context poly l = in let impl = List.exists test impls in let decl = (Discharge, poly, Definitional) in + let univs = if poly + then Polymorphic_const_entry (Univ.ContextSet.to_context !uctx) + else Monomorphic_const_entry !uctx + in let nstatus = match b with | None -> - pi3 (Command.declare_assumption false decl (t, !uctx) Universes.empty_binders [] impl + pi3 (Command.declare_assumption false decl (t, univs) Universes.empty_binders [] impl Vernacexpr.NoInline (Loc.tag id)) | Some b -> - let univs = if poly - then Polymorphic_const_entry (Univ.ContextSet.to_context !uctx) - else Monomorphic_const_entry !uctx - in let decl = (Discharge, poly, Definition) in let entry = Declare.definition_entry ~univs ~types:t b in let hook = Lemmas.mk_hook (fun _ gr -> gr) in diff --git a/vernac/command.ml b/vernac/command.ml index 1a8b16a89..373e5a1be 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -177,6 +177,10 @@ let do_definition ident k univdecl bl red_option c ctypopt hook = let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl (_,ident) = match local with | Discharge when Lib.sections_are_opened () -> + let ctx = match ctx with + | Monomorphic_const_entry ctx -> ctx + | Polymorphic_const_entry ctx -> Univ.ContextSet.of_context ctx + in let decl = (Lib.cwd(), SectionLocalAssum ((c,ctx),p,impl), IsAssumption kind) in let _ = declare_variable ident decl in let () = assumption_message ident in @@ -197,10 +201,6 @@ match local with | DefaultInline -> Some (Flags.get_inline_level()) | InlineAt i -> Some i in - let ctx = if p - then Polymorphic_const_entry (Univ.ContextSet.to_context ctx) - else Monomorphic_const_entry ctx - in let decl = (ParameterEntry (None,(c,ctx),inl), IsAssumption kind) in let kn = declare_constant ident ~local decl in let gr = ConstRef kn in @@ -221,26 +221,63 @@ let interp_assumption evdref env impls bl c = let ty = EConstr.Unsafe.to_constr ty in (ty, impls) -let declare_assumptions idl is_coe k (c,ctx) pl imps impl_is_on nl = +(* When monomorphic the universe constraints are declared with the first declaration only. *) +let next_uctx = + let empty_uctx = Monomorphic_const_entry Univ.ContextSet.empty in + function + | Polymorphic_const_entry _ as uctx -> uctx + | Monomorphic_const_entry _ -> empty_uctx + +let declare_assumptions idl is_coe k (c,uctx) pl imps nl = let refs, status, _ = - List.fold_left (fun (refs,status,ctx) id -> + List.fold_left (fun (refs,status,uctx) id -> let ref',u',status' = - declare_assumption is_coe k (c,ctx) pl imps impl_is_on nl id in - (ref',u')::refs, status' && status, Univ.ContextSet.empty) - ([],true,ctx) idl + declare_assumption is_coe k (c,uctx) pl imps false nl id in + (ref',u')::refs, status' && status, next_uctx uctx) + ([],true,uctx) idl in List.rev refs, status -let do_assumptions_unbound_univs (_, poly, _ as kind) nl l = + +let maybe_error_many_udecls = function + | ((loc,id), Some _) -> + user_err ?loc ~hdr:"many_universe_declarations" + Pp.(str "When declaring multiple axioms in one command, " ++ + str "only the first is allowed a universe binder " ++ + str "(which will be shared by the whole block).") + | (_, None) -> () + +let process_assumptions_udecls kind l = + let udecl, first_id = match l with + | (coe, ((id, udecl)::rest, c))::rest' -> + List.iter maybe_error_many_udecls rest; + List.iter (fun (coe, (idl, c)) -> List.iter maybe_error_many_udecls idl) rest'; + udecl, id + | (_, ([], _))::_ | [] -> assert false + in + let () = match kind, udecl with + | (Discharge, _, _), Some _ when Lib.sections_are_opened () -> + let loc = fst first_id in + let msg = Pp.str "Section variables cannot be polymorphic." in + user_err ?loc msg + | _ -> () + in + udecl, List.map (fun (coe, (idl, c)) -> coe, (List.map fst idl, c)) l + +let do_assumptions kind nl l = let open Context.Named.Declaration in let env = Global.env () in - let evdref = ref (Evd.from_env env) in - let l = - if poly then + let udecl, l = process_assumptions_udecls kind l in + let evdref, udecl = + let evd, udecl = Univdecls.interp_univ_decl_opt env udecl in + ref evd, udecl + in + let l = + if pi2 kind (* poly *) then (* Separate declarations so that A B : Type puts A and B in different levels. *) List.fold_right (fun (is_coe,(idl,c)) acc -> - List.fold_right (fun id acc -> - (is_coe, ([id], c)) :: acc) idl acc) + List.fold_right (fun id acc -> + (is_coe, ([id], c)) :: acc) idl acc) l [] else l in @@ -252,71 +289,31 @@ let do_assumptions_unbound_univs (_, poly, _ as kind) nl l = let ienv = List.fold_right (fun (_,id) ienv -> let impls = compute_internalization_data env Variable t imps in Id.Map.add id impls ienv) idl ienv in - ((env,ienv),((is_coe,idl),t,imps))) + ((env,ienv),((is_coe,idl),t,imps))) (env,empty_internalization_env) l in let evd = solve_remaining_evars all_and_fail_flags env !evdref Evd.empty in (* The universe constraints come from the whole telescope. *) let evd = Evd.nf_constraints evd in - let ctx = Evd.universe_context_set evd in - let nf_evar c = EConstr.Unsafe.to_constr (nf_evar evd (EConstr.of_constr c)) in - let l = List.map (on_pi2 nf_evar) l in - pi2 (List.fold_left (fun (subst,status,ctx) ((is_coe,idl),t,imps) -> - let t = replace_vars subst t in - let (refs,status') = declare_assumptions idl is_coe kind (t,ctx) Universes.empty_binders imps false nl in - let subst' = List.map2 - (fun (_,id) (c,u) -> (id,Universes.constr_of_global_univ (c,u))) - idl refs - in - (subst'@subst, status' && status, - (* The universe constraints are declared with the first declaration only. *) - Univ.ContextSet.empty)) ([],true,ctx) l) - -let do_assumptions_bound_univs coe kind nl id pl c = - let env = Global.env () in - let evd, decl = Univdecls.interp_univ_decl_opt env pl in - let evdref = ref evd in - let ty, impls = interp_type_evars_impls env evdref c in - let nf, subst = Evarutil.e_nf_evars_and_universes evdref in - let ty = EConstr.Unsafe.to_constr ty in - let ty = nf ty in - let vars = Univops.universes_of_constr ty in - let evd = Evd.restrict_universe_context !evdref vars in - let uctx = Evd.check_univ_decl ~poly:(pi2 kind) evd decl in - let binders = Evd.universe_binders evd in - let uctx = match uctx with - | Polymorphic_const_entry uctx -> - (* ?? *) - Univ.ContextSet.of_context uctx - | Monomorphic_const_entry uctx -> uctx + let nf_evar c = EConstr.to_constr evd (EConstr.of_constr c) in + let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) -> + let t = nf_evar t in + let uvars = Univ.LSet.union uvars (Univops.universes_of_constr t) in + uvars, (coe,t,imps)) + Univ.LSet.empty l in - let (_, _, st) = declare_assumption coe kind (ty, uctx) binders impls false nl id in - st - -let do_assumptions kind nl l = match l with -| [coe, ([id, Some pl], c)] -> - let () = match kind with - | (Discharge, _, _) when Lib.sections_are_opened () -> - let loc = fst id in - let msg = Pp.str "Section variables cannot be polymorphic." in - user_err ?loc msg - | _ -> () - in - do_assumptions_bound_univs coe kind nl id (Some pl) c -| _ -> - let map (coe, (idl, c)) = - let map (id, univs) = match univs with - | None -> id - | Some _ -> - let loc = fst id in - let msg = - Pp.str "Assumptions with bound universes can only be defined one at a time." in - user_err ?loc msg - in - (coe, (List.map map idl, c)) - in - let l = List.map map l in - do_assumptions_unbound_univs kind nl l + let evd = Evd.restrict_universe_context evd uvars in + let uctx = Evd.check_univ_decl ~poly:(pi2 kind) evd udecl in + let ubinders = Evd.universe_binders evd in + pi2 (List.fold_left (fun (subst,status,uctx) ((is_coe,idl),t,imps) -> + let t = replace_vars subst t in + let refs, status' = declare_assumptions idl is_coe kind (t,uctx) ubinders imps nl in + let subst' = List.map2 + (fun (_,id) (c,u) -> (id, Universes.constr_of_global_univ (c,u))) + idl refs + in + subst'@subst, status' && status, next_uctx uctx) + ([], true, uctx) l) (* 3a| Elimination schemes for mutual inductive definitions *) diff --git a/vernac/command.mli b/vernac/command.mli index fb99a717b..070f3e112 100644 --- a/vernac/command.mli +++ b/vernac/command.mli @@ -43,7 +43,7 @@ val do_definition : Id.t -> definition_kind -> Vernacexpr.universe_decl_expr opt (** returns [false] if the assumption is neither local to a section, nor in a module type and meant to be instantiated. *) val declare_assumption : coercion_flag -> assumption_kind -> - types Univ.in_universe_context_set -> + types in_constant_universes_entry -> Universes.universe_binders -> Impargs.manual_implicits -> bool (** implicit *) -> Vernacexpr.inline -> variable Loc.located -> global_reference * Univ.Instance.t * bool -- cgit v1.2.3 From edb2b2535d62b37c324c98b28802c1b1699d4014 Mon Sep 17 00:00:00 2001 From: Gaëtan Gilbert Date: Sat, 25 Nov 2017 22:08:01 +0100 Subject: Restrict universe context when declaring constants in obligations. --- test-suite/success/polymorphism.v | 7 +++++++ vernac/obligations.ml | 12 ++++++++---- 2 files changed, 15 insertions(+), 4 deletions(-) (limited to 'vernac') diff --git a/test-suite/success/polymorphism.v b/test-suite/success/polymorphism.v index b3f9a4994..d76b30791 100644 --- a/test-suite/success/polymorphism.v +++ b/test-suite/success/polymorphism.v @@ -461,3 +461,10 @@ Section test_letin_subtyping. Qed. End test_letin_subtyping. + +Module ObligationRegression. + (** Test for a regression encountered when fixing obligations for + stronger restriction of universe context. *) + Require Import CMorphisms. + Check trans_co_eq_inv_arrow_morphism@{_ _ _ _ _ _ _ _}. +End ObligationRegression. diff --git a/vernac/obligations.ml b/vernac/obligations.ml index 17a596c42..570f2e5b3 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -475,13 +475,17 @@ let declare_definition prg = (Evd.evar_universe_context_subst prg.prg_ctx) in let opaque = prg.prg_opaque in let fix_exn = Hook.get get_fix_exn () in - let univs = UState.check_univ_decl ~poly:(pi2 prg.prg_kind) prg.prg_ctx prg.prg_univdecl in - let ce = definition_entry ~fix_exn ~opaque ~types:(nf typ) ~univs (nf body) in + let typ = nf typ in + let body = nf body in + let uvars = Univ.LSet.union (Univops.universes_of_constr typ) (Univops.universes_of_constr body) in + let uctx = UState.restrict prg.prg_ctx uvars in + let univs = UState.check_univ_decl ~poly:(pi2 prg.prg_kind) uctx prg.prg_univdecl in + let ce = definition_entry ~fix_exn ~opaque ~types:typ ~univs body in let () = progmap_remove prg in - let ubinders = UState.universe_binders prg.prg_ctx in + let ubinders = UState.universe_binders uctx in DeclareDef.declare_definition prg.prg_name prg.prg_kind ce ubinders prg.prg_implicits - (Lemmas.mk_hook (fun l r -> Lemmas.call_hook fix_exn prg.prg_hook l r prg.prg_ctx; r)) + (Lemmas.mk_hook (fun l r -> Lemmas.call_hook fix_exn prg.prg_hook l r uctx; r)) let rec lam_index n t acc = match Constr.kind t with -- cgit v1.2.3