From f72a67569ec8cb9160d161699302b67919da5686 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Thu, 27 Jul 2017 14:54:41 +0200 Subject: Allow declaring universe constraints at definition level. Introduce a "+" modifier for universe and constraint declarations to indicate that these can be extended in the final definition/proof. By default [Definition f] is equivalent to [Definition f@{+|+}], i.e universes can be introduced and constraints as well. For [f@{}] or [f@{i j}], the constraints can be extended, no universe introduced, to maintain compatibility with existing developments. Use [f@{i j | }] to indicate that no constraint (nor universe) can be introduced. These kind of definitions could benefit from asynchronous processing. Declarations of universe binders and constraints also works for monomorphic definitions. --- pretyping/pretyping.mllib | 1 + pretyping/univdecls.ml | 64 +++++++++++++++++++++++++++++++++++++++++++++++ pretyping/univdecls.mli | 19 ++++++++++++++ 3 files changed, 84 insertions(+) create mode 100644 pretyping/univdecls.ml create mode 100644 pretyping/univdecls.mli (limited to 'pretyping') diff --git a/pretyping/pretyping.mllib b/pretyping/pretyping.mllib index c8b3307d7..d04dcb8e3 100644 --- a/pretyping/pretyping.mllib +++ b/pretyping/pretyping.mllib @@ -29,3 +29,4 @@ Indrec Cases Pretyping Unification +Univdecls diff --git a/pretyping/univdecls.ml b/pretyping/univdecls.ml new file mode 100644 index 000000000..d7c42d03a --- /dev/null +++ b/pretyping/univdecls.ml @@ -0,0 +1,64 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Loc.tag Univ.Level.prop + | GSet -> Loc.tag Univ.Level.set + | GType None | GType (Some (_, Anonymous)) -> + user_err ~hdr:"interp_constraint" + (str "Cannot declare constraints on anonymous universes") + | GType (Some (loc, Name id)) -> + try loc, Evd.universe_of_name evd (Id.to_string id) + with Not_found -> + user_err ?loc ~hdr:"interp_constraint" (str "Undeclared universe " ++ pr_id id) + in + let interp (evd,cstrs) (u, d, u') = + let lloc, ul = u_of_id u and rloc, u'l = u_of_id u' in + let cstr = (ul,d,u'l) in + let cstrs' = Univ.Constraint.add cstr cstrs in + try let evd = Evd.add_constraints evd (Univ.Constraint.singleton cstr) in + evd, cstrs' + with Univ.UniverseInconsistency e -> + user_err ~hdr:"interp_constraint" (str "Universe inconsistency" (* TODO *)) + in + List.fold_left interp (evd,Univ.Constraint.empty) cstrs + +let interp_univ_decl env decl = + let open Misctypes in + let pl = decl.univdecl_instance in + let evd = Evd.from_ctx (Evd.make_evar_universe_context env (Some pl)) in + let evd, cstrs = interp_univ_constraints env evd decl.univdecl_constraints in + let decl = { univdecl_instance = pl; + univdecl_extensible_instance = decl.univdecl_extensible_instance; + univdecl_constraints = cstrs; + univdecl_extensible_constraints = decl.univdecl_extensible_constraints } + in evd, decl + +let interp_univ_decl_opt env l = + match l with + | None -> Evd.from_env env, default_univ_decl + | Some decl -> interp_univ_decl env decl diff --git a/pretyping/univdecls.mli b/pretyping/univdecls.mli new file mode 100644 index 000000000..0c3b749cb --- /dev/null +++ b/pretyping/univdecls.mli @@ -0,0 +1,19 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Vernacexpr.universe_decl_expr -> + Evd.evar_map * universe_decl + +val interp_univ_decl_opt : Environ.env -> Vernacexpr.universe_decl_expr option -> + Evd.evar_map * universe_decl -- cgit v1.2.3