diff options
| author | Emilio Jesus Gallego Arias | 2019-10-24 03:43:04 +0200 |
|---|---|---|
| committer | Emilio Jesus Gallego Arias | 2019-10-24 21:33:58 +0200 |
| commit | 43f037b5f3af7ab642bed4c6767bf7845156f92f (patch) | |
| tree | 70614c100f24cdc76caf34541f784bc8994e37db | |
| parent | 4c779c4fee1134c5d632885de60db73d56021df4 (diff) | |
[declare] Split universe declaration code to vernac/
The code is self-contained and only used by commands; this also
highlights the several `Libobject.obj` registered for each
declaration.
| -rw-r--r-- | tactics/declare.ml | 103 | ||||
| -rw-r--r-- | tactics/declare.mli | 10 | ||||
| -rw-r--r-- | vernac/classes.ml | 4 | ||||
| -rw-r--r-- | vernac/comAssumption.ml | 2 | ||||
| -rw-r--r-- | vernac/comInductive.ml | 2 | ||||
| -rw-r--r-- | vernac/declareDef.ml | 2 | ||||
| -rw-r--r-- | vernac/declareUniv.ml | 110 | ||||
| -rw-r--r-- | vernac/declareUniv.mli | 17 | ||||
| -rw-r--r-- | vernac/lemmas.ml | 4 | ||||
| -rw-r--r-- | vernac/vernac.mllib | 1 | ||||
| -rw-r--r-- | vernac/vernacentries.ml | 4 |
11 files changed, 139 insertions, 120 deletions
diff --git a/tactics/declare.ml b/tactics/declare.ml index 7d32f1a7e8..412f735983 100644 --- a/tactics/declare.ml +++ b/tactics/declare.ml @@ -538,106 +538,3 @@ let assumption_message id = the type of the object than to the name of the object (see discussion on coqdev: "Chapter 4 of the Reference Manual", 8/10/2015) *) Flags.if_verbose Feedback.msg_info (Id.print id ++ str " is declared") - -(** Global universes are not substitutive objects but global objects - bound at the *library* or *module* level. The polymorphic flag is - used to distinguish universes declared in polymorphic sections, which - are discharged and do not remain in scope. *) - -type universe_source = - | BoundUniv (* polymorphic universe, bound in a function (this will go away someday) *) - | QualifiedUniv of Id.t (* global universe introduced by some global value *) - | UnqualifiedUniv (* other global universe *) - -type universe_name_decl = universe_source * (Id.t * Univ.Level.UGlobal.t) list - -let check_exists_universe sp = - if Nametab.exists_universe sp then - raise (AlreadyDeclared (Some "Universe", Libnames.basename sp)) - else () - -let qualify_univ i dp src id = - match src with - | BoundUniv | UnqualifiedUniv -> - i, Libnames.make_path dp id - | QualifiedUniv l -> - let dp = DirPath.repr dp in - Nametab.map_visibility succ i, Libnames.make_path (DirPath.make (l::dp)) id - -let do_univ_name ~check i dp src (id,univ) = - let i, sp = qualify_univ i dp src id in - if check then check_exists_universe sp; - Nametab.push_universe i sp univ - -let cache_univ_names ((sp, _), (src, univs)) = - let depth = sections_depth () in - let dp = Libnames.pop_dirpath_n depth (Libnames.dirpath sp) in - List.iter (do_univ_name ~check:true (Nametab.Until 1) dp src) univs - -let load_univ_names i ((sp, _), (src, univs)) = - List.iter (do_univ_name ~check:false (Nametab.Until i) (Libnames.dirpath sp) src) univs - -let open_univ_names i ((sp, _), (src, univs)) = - List.iter (do_univ_name ~check:false (Nametab.Exactly i) (Libnames.dirpath sp) src) univs - -let discharge_univ_names = function - | _, (BoundUniv, _) -> None - | _, ((QualifiedUniv _ | UnqualifiedUniv), _ as x) -> Some x - -let input_univ_names : universe_name_decl -> Libobject.obj = - declare_object - { (default_object "Global universe name state") with - cache_function = cache_univ_names; - load_function = load_univ_names; - open_function = open_univ_names; - discharge_function = discharge_univ_names; - subst_function = (fun (subst, a) -> (* Actually the name is generated once and for all. *) a); - classify_function = (fun a -> Substitute a) } - -let declare_univ_binders gr pl = - if Global.is_polymorphic gr then - () - else - let l = let open GlobRef in match gr with - | ConstRef c -> Label.to_id @@ Constant.label c - | IndRef (c, _) -> Label.to_id @@ MutInd.label c - | VarRef id -> - CErrors.anomaly ~label:"declare_univ_binders" Pp.(str "declare_univ_binders on variable " ++ Id.print id ++ str".") - | ConstructRef _ -> - CErrors.anomaly ~label:"declare_univ_binders" - Pp.(str "declare_univ_binders on an constructor reference") - in - let univs = Id.Map.fold (fun id univ univs -> - match Univ.Level.name univ with - | None -> assert false (* having Prop/Set/Var as binders is nonsense *) - | Some univ -> (id,univ)::univs) pl [] - in - Lib.add_anonymous_leaf (input_univ_names (QualifiedUniv l, univs)) - -let do_universe ~poly l = - let in_section = Global.sections_are_opened () in - let () = - if poly && not in_section then - CErrors.user_err ~hdr:"Constraint" - (str"Cannot declare polymorphic universes outside sections") - in - let l = List.map (fun {CAst.v=id} -> (id, UnivGen.new_univ_global ())) l in - let ctx = List.fold_left (fun ctx (_,qid) -> Univ.LSet.add (Univ.Level.make qid) ctx) - Univ.LSet.empty l, Univ.Constraint.empty - in - let src = if poly then BoundUniv else UnqualifiedUniv in - let () = Lib.add_anonymous_leaf (input_univ_names (src, l)) in - declare_universe_context ~poly ctx - -let do_constraint ~poly l = - let open Univ in - let u_of_id x = - Pretyping.interp_known_glob_level (Evd.from_env (Global.env ())) x - in - let constraints = List.fold_left (fun acc (l, d, r) -> - let lu = u_of_id l and ru = u_of_id r in - Constraint.add (lu, d, ru) acc) - Constraint.empty l - in - let uctx = ContextSet.add_constraints constraints ContextSet.empty in - declare_universe_context ~poly uctx diff --git a/tactics/declare.mli b/tactics/declare.mli index a6c1374a77..bf5877d4c4 100644 --- a/tactics/declare.mli +++ b/tactics/declare.mli @@ -43,6 +43,8 @@ type 'a constant_entry = | ParameterEntry of parameter_entry | PrimitiveEntry of primitive_entry +val declare_universe_context : poly:bool -> Univ.ContextSet.t -> unit + val declare_variable : name:variable -> kind:Decls.logical_kind @@ -102,13 +104,5 @@ val recursive_message : bool (** true = fixpoint *) -> val exists_name : Id.t -> bool -(** Global universe contexts, names and constraints *) -val declare_univ_binders : GlobRef.t -> UnivNames.universe_binders -> unit - -val declare_universe_context : poly:bool -> Univ.ContextSet.t -> unit - -val do_universe : poly:bool -> lident list -> unit -val do_constraint : poly:bool -> Glob_term.glob_constraint list -> unit - (* Used outside this module only in indschemes *) exception AlreadyDeclared of (string option * Id.t) diff --git a/vernac/classes.ml b/vernac/classes.ml index 0a8c4e6b0f..4715b5b101 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -325,7 +325,7 @@ let declare_instance_constant info global imps ?hook name decl poly sigma term t let entry = Declare.definition_entry ~types:termtype ~univs:uctx term in let kn = Declare.declare_constant ~name ~kind (Declare.DefinitionEntry entry) in Declare.definition_message name; - Declare.declare_univ_binders (GlobRef.ConstRef kn) (Evd.universe_binders sigma); + DeclareUniv.declare_univ_binders (GlobRef.ConstRef kn) (Evd.universe_binders sigma); instance_hook info global imps ?hook (GlobRef.ConstRef kn) let do_declare_instance sigma ~global ~poly k u ctx ctx' pri decl imps subst name = @@ -338,7 +338,7 @@ let do_declare_instance sigma ~global ~poly k u ctx ctx' pri decl imps subst nam let sigma, entry = DeclareDef.prepare_parameter ~allow_evars:false ~poly sigma decl termtype in let cst = Declare.declare_constant ~name ~kind:Decls.(IsAssumption Logical) (Declare.ParameterEntry entry) in - Declare.declare_univ_binders (GlobRef.ConstRef cst) (Evd.universe_binders sigma); + DeclareUniv.declare_univ_binders (GlobRef.ConstRef cst) (Evd.universe_binders sigma); instance_hook pri global imps (GlobRef.ConstRef cst) let declare_instance_program env sigma ~global ~poly id pri imps decl term termtype = diff --git a/vernac/comAssumption.ml b/vernac/comAssumption.ml index f9b73a59eb..8c3c206c71 100644 --- a/vernac/comAssumption.ml +++ b/vernac/comAssumption.ml @@ -69,7 +69,7 @@ let declare_axiom is_coe ~poly ~local ~kind typ (univs, pl) imps nl {CAst.v=name let kn = Declare.declare_constant ~name ~local ~kind decl in let gr = GlobRef.ConstRef kn in let () = maybe_declare_manual_implicits false gr imps in - let () = Declare.declare_univ_binders gr pl in + let () = DeclareUniv.declare_univ_binders gr pl in let () = Declare.assumption_message name in let env = Global.env () in let sigma = Evd.from_env env in diff --git a/vernac/comInductive.ml b/vernac/comInductive.ml index 98b869d72e..9270165379 100644 --- a/vernac/comInductive.ml +++ b/vernac/comInductive.ml @@ -577,7 +577,7 @@ let declare_mutual_inductive_with_eliminations ?(primitive_expected=false) mie p let (_, kn), prim = declare_mind mie in let mind = Global.mind_of_delta_kn kn in if primitive_expected && not prim then warn_non_primitive_record (mind,0); - Declare.declare_univ_binders (GlobRef.IndRef (mind,0)) pl; + DeclareUniv.declare_univ_binders (GlobRef.IndRef (mind,0)) pl; List.iteri (fun i (indimpls, constrimpls) -> let ind = (mind,i) in let gr = GlobRef.IndRef ind in diff --git a/vernac/declareDef.ml b/vernac/declareDef.ml index 1926faaf0e..300bb7452e 100644 --- a/vernac/declareDef.ml +++ b/vernac/declareDef.ml @@ -54,7 +54,7 @@ let declare_definition ~name ~scope ~kind ?hook_data udecl ce imps = | Global local -> let kn = declare_constant ~name ~local ~kind:Decls.(IsDefinition kind) (DefinitionEntry ce) in let gr = Names.GlobRef.ConstRef kn in - let () = Declare.declare_univ_binders gr udecl in + let () = DeclareUniv.declare_univ_binders gr udecl in gr in let () = maybe_declare_manual_implicits false gr imps in diff --git a/vernac/declareUniv.ml b/vernac/declareUniv.ml new file mode 100644 index 0000000000..69ba9d76ec --- /dev/null +++ b/vernac/declareUniv.ml @@ -0,0 +1,110 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(* * (see LICENSE file for the text of the license) *) +(************************************************************************) + +open Names + +type universe_source = + | BoundUniv (* polymorphic universe, bound in a function (this will go away someday) *) + | QualifiedUniv of Id.t (* global universe introduced by some global value *) + | UnqualifiedUniv (* other global universe *) + +type universe_name_decl = universe_source * (Id.t * Univ.Level.UGlobal.t) list + +let check_exists_universe sp = + if Nametab.exists_universe sp then + raise (Declare.AlreadyDeclared (Some "Universe", Libnames.basename sp)) + else () + +let qualify_univ i dp src id = + match src with + | BoundUniv | UnqualifiedUniv -> + i, Libnames.make_path dp id + | QualifiedUniv l -> + let dp = DirPath.repr dp in + Nametab.map_visibility succ i, Libnames.make_path (DirPath.make (l::dp)) id + +let do_univ_name ~check i dp src (id,univ) = + let i, sp = qualify_univ i dp src id in + if check then check_exists_universe sp; + Nametab.push_universe i sp univ + +let cache_univ_names ((sp, _), (src, univs)) = + let depth = Lib.sections_depth () in + let dp = Libnames.pop_dirpath_n depth (Libnames.dirpath sp) in + List.iter (do_univ_name ~check:true (Nametab.Until 1) dp src) univs + +let load_univ_names i ((sp, _), (src, univs)) = + List.iter (do_univ_name ~check:false (Nametab.Until i) (Libnames.dirpath sp) src) univs + +let open_univ_names i ((sp, _), (src, univs)) = + List.iter (do_univ_name ~check:false (Nametab.Exactly i) (Libnames.dirpath sp) src) univs + +let discharge_univ_names = function + | _, (BoundUniv, _) -> None + | _, ((QualifiedUniv _ | UnqualifiedUniv), _ as x) -> Some x + +let input_univ_names : universe_name_decl -> Libobject.obj = + let open Libobject in + declare_object + { (default_object "Global universe name state") with + cache_function = cache_univ_names; + load_function = load_univ_names; + open_function = open_univ_names; + discharge_function = discharge_univ_names; + subst_function = (fun (subst, a) -> (* Actually the name is generated once and for all. *) a); + classify_function = (fun a -> Substitute a) } + +let declare_univ_binders gr pl = + if Global.is_polymorphic gr then + () + else + let l = let open GlobRef in match gr with + | ConstRef c -> Label.to_id @@ Constant.label c + | IndRef (c, _) -> Label.to_id @@ MutInd.label c + | VarRef id -> + CErrors.anomaly ~label:"declare_univ_binders" Pp.(str "declare_univ_binders on variable " ++ Id.print id ++ str".") + | ConstructRef _ -> + CErrors.anomaly ~label:"declare_univ_binders" + Pp.(str "declare_univ_binders on an constructor reference") + in + let univs = Id.Map.fold (fun id univ univs -> + match Univ.Level.name univ with + | None -> assert false (* having Prop/Set/Var as binders is nonsense *) + | Some univ -> (id,univ)::univs) pl [] + in + Lib.add_anonymous_leaf (input_univ_names (QualifiedUniv l, univs)) + +let do_universe ~poly l = + let in_section = Global.sections_are_opened () in + let () = + if poly && not in_section then + CErrors.user_err ~hdr:"Constraint" + (Pp.str"Cannot declare polymorphic universes outside sections") + in + let l = List.map (fun {CAst.v=id} -> (id, UnivGen.new_univ_global ())) l in + let ctx = List.fold_left (fun ctx (_,qid) -> Univ.LSet.add (Univ.Level.make qid) ctx) + Univ.LSet.empty l, Univ.Constraint.empty + in + let src = if poly then BoundUniv else UnqualifiedUniv in + let () = Lib.add_anonymous_leaf (input_univ_names (src, l)) in + Declare.declare_universe_context ~poly ctx + +let do_constraint ~poly l = + let open Univ in + let u_of_id x = + Pretyping.interp_known_glob_level (Evd.from_env (Global.env ())) x + in + let constraints = List.fold_left (fun acc (l, d, r) -> + let lu = u_of_id l and ru = u_of_id r in + Constraint.add (lu, d, ru) acc) + Constraint.empty l + in + let uctx = ContextSet.add_constraints constraints ContextSet.empty in + Declare.declare_universe_context ~poly uctx diff --git a/vernac/declareUniv.mli b/vernac/declareUniv.mli new file mode 100644 index 0000000000..ce2a6e225c --- /dev/null +++ b/vernac/declareUniv.mli @@ -0,0 +1,17 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2019 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \VV/ **************************************************************) +(* // * This file is distributed under the terms of the *) +(* * GNU Lesser General Public License Version 2.1 *) +(* * (see LICENSE file for the text of the license) *) +(************************************************************************) + +open Names + +(** Global universe contexts, names and constraints *) +val declare_univ_binders : GlobRef.t -> UnivNames.universe_binders -> unit + +val do_universe : poly:bool -> lident list -> unit +val do_constraint : poly:bool -> Glob_term.glob_constraint list -> unit diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index e49277c51b..5ace8b917c 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -335,7 +335,7 @@ let finish_admitted env sigma ~name ~poly ~scope pe ctx hook ~udecl impargs othe in let kn = Declare.declare_constant ~name ~local ~kind:Decls.(IsAssumption Conjectural) (Declare.ParameterEntry pe) in let () = Declare.assumption_message name in - Declare.declare_univ_binders (GlobRef.ConstRef kn) (UState.universe_binders ctx); + DeclareUniv.declare_univ_binders (GlobRef.ConstRef kn) (UState.universe_binders ctx); (* This takes care of the implicits and hook for the current constant*) process_recthms ?fix_exn:None ?hook env sigma ctx ~udecl ~poly ~scope:(Global local) (GlobRef.ConstRef kn) impargs other_thms @@ -425,7 +425,7 @@ let finish_proved env sigma idopt po info = then Proof_using.suggest_constant (Global.env ()) kn in let gr = GlobRef.ConstRef kn in - Declare.declare_univ_binders gr (UState.universe_binders universes); + DeclareUniv.declare_univ_binders gr (UState.universe_binders universes); gr in Declare.definition_message name; diff --git a/vernac/vernac.mllib b/vernac/vernac.mllib index afc701edbc..e9bcef762e 100644 --- a/vernac/vernac.mllib +++ b/vernac/vernac.mllib @@ -13,6 +13,7 @@ Ppvernac Proof_using Egramcoq Metasyntax +DeclareUniv DeclareDef DeclareObl Canonical diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 430cee62c2..4ecd815dd2 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -814,14 +814,14 @@ let vernac_universe ~poly l = user_err ~hdr:"vernac_universe" (str"Polymorphic universes can only be declared inside sections, " ++ str "use Monomorphic Universe instead"); - Declare.do_universe ~poly l + DeclareUniv.do_universe ~poly l let vernac_constraint ~poly l = if poly && not (Global.sections_are_opened ()) then user_err ~hdr:"vernac_constraint" (str"Polymorphic universe constraints can only be declared" ++ str " inside sections, use Monomorphic Constraint instead"); - Declare.do_constraint ~poly l + DeclareUniv.do_constraint ~poly l (**********************) (* Modules *) |
