From 47ce63d23b8efe35babe0f4429c550400afd6b4f Mon Sep 17 00:00:00 2001 From: Amin Timany Date: Thu, 4 May 2017 17:53:12 +0200 Subject: Move univops from kernel to library --- kernel/kernel.mllib | 3 +-- kernel/univops.ml | 70 --------------------------------------------------- kernel/univops.mli | 17 ------------- library/library.mllib | 1 + library/univops.ml | 70 +++++++++++++++++++++++++++++++++++++++++++++++++++ library/univops.mli | 17 +++++++++++++ 6 files changed, 89 insertions(+), 89 deletions(-) delete mode 100644 kernel/univops.ml delete mode 100644 kernel/univops.mli create mode 100644 library/univops.ml create mode 100644 library/univops.mli diff --git a/kernel/kernel.mllib b/kernel/kernel.mllib index 8132d6685..0813315b5 100644 --- a/kernel/kernel.mllib +++ b/kernel/kernel.mllib @@ -41,6 +41,5 @@ Nativelibrary Safe_typing Vm Csymtable -Vconv Declarations -Univops +Vconv diff --git a/kernel/univops.ml b/kernel/univops.ml deleted file mode 100644 index e9383c6d9..000000000 --- a/kernel/univops.ml +++ /dev/null @@ -1,70 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* - LSet.fold LSet.add (Instance.levels u) s - | Sort u when not (Sorts.is_small u) -> - let u = univ_of_sort u in - LSet.fold LSet.add (Universe.levels u) s - | _ -> fold_constr aux s c - in aux LSet.empty c - -let universes_of_inductive mind = - if mind.mind_polymorphic then - begin - let u = Univ.UContext.instance (Univ.UInfoInd.univ_context mind.mind_universes) in - let univ_of_one_ind oind = - let arity_univs = - Context.Rel.fold_outside - (fun decl unvs -> - Univ.LSet.union - (Context.Rel.Declaration.fold_constr - (fun cnstr unvs -> - let cnstr = Vars.subst_instance_constr u cnstr in - Univ.LSet.union - (universes_of_constr cnstr) unvs) - decl Univ.LSet.empty) unvs) - oind.mind_arity_ctxt ~init:Univ.LSet.empty - in - Array.fold_left (fun unvs cns -> - let cns = Vars.subst_instance_constr u cns in - Univ.LSet.union (universes_of_constr cns) unvs) arity_univs - oind.mind_nf_lc - in - let univs = Array.fold_left (fun unvs pk -> Univ.LSet.union (univ_of_one_ind pk) unvs) Univ.LSet.empty mind.mind_packets in - let mindcnt = Univ.UContext.constraints (Univ.instantiate_univ_context (Univ.UInfoInd.univ_context mind.mind_universes)) in - let univs = Univ.LSet.union univs (Univ.universes_of_constraints mindcnt) in - univs - end - else LSet.empty - -let restrict_universe_context (univs,csts) s = - (* Universes that are not necessary to typecheck the term. - E.g. univs introduced by tactics and not used in the proof term. *) - let diff = LSet.diff univs s in - let rec aux diff candid univs ness = - let (diff', candid', univs', ness') = - Constraint.fold - (fun (l, d, r as c) (diff, candid, univs, csts) -> - if not (LSet.mem l diff) then - (LSet.remove r diff, candid, univs, Constraint.add c csts) - else if not (LSet.mem r diff) then - (LSet.remove l diff, candid, univs, Constraint.add c csts) - else (diff, Constraint.add c candid, univs, csts)) - candid (diff, Constraint.empty, univs, ness) - in - if ness' == ness then (LSet.diff univs diff', ness) - else aux diff' candid' univs' ness' - in aux diff csts univs Constraint.empty diff --git a/kernel/univops.mli b/kernel/univops.mli deleted file mode 100644 index 5b499c75b..000000000 --- a/kernel/univops.mli +++ /dev/null @@ -1,17 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* universe_set -val universes_of_inductive : mutual_inductive_body -> universe_set -val restrict_universe_context : universe_context_set -> universe_set -> universe_context_set diff --git a/library/library.mllib b/library/library.mllib index 6f433b77d..d94fc2291 100644 --- a/library/library.mllib +++ b/library/library.mllib @@ -1,3 +1,4 @@ +Univops Nameops Libnames Globnames diff --git a/library/univops.ml b/library/univops.ml new file mode 100644 index 000000000..e9383c6d9 --- /dev/null +++ b/library/univops.ml @@ -0,0 +1,70 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* + LSet.fold LSet.add (Instance.levels u) s + | Sort u when not (Sorts.is_small u) -> + let u = univ_of_sort u in + LSet.fold LSet.add (Universe.levels u) s + | _ -> fold_constr aux s c + in aux LSet.empty c + +let universes_of_inductive mind = + if mind.mind_polymorphic then + begin + let u = Univ.UContext.instance (Univ.UInfoInd.univ_context mind.mind_universes) in + let univ_of_one_ind oind = + let arity_univs = + Context.Rel.fold_outside + (fun decl unvs -> + Univ.LSet.union + (Context.Rel.Declaration.fold_constr + (fun cnstr unvs -> + let cnstr = Vars.subst_instance_constr u cnstr in + Univ.LSet.union + (universes_of_constr cnstr) unvs) + decl Univ.LSet.empty) unvs) + oind.mind_arity_ctxt ~init:Univ.LSet.empty + in + Array.fold_left (fun unvs cns -> + let cns = Vars.subst_instance_constr u cns in + Univ.LSet.union (universes_of_constr cns) unvs) arity_univs + oind.mind_nf_lc + in + let univs = Array.fold_left (fun unvs pk -> Univ.LSet.union (univ_of_one_ind pk) unvs) Univ.LSet.empty mind.mind_packets in + let mindcnt = Univ.UContext.constraints (Univ.instantiate_univ_context (Univ.UInfoInd.univ_context mind.mind_universes)) in + let univs = Univ.LSet.union univs (Univ.universes_of_constraints mindcnt) in + univs + end + else LSet.empty + +let restrict_universe_context (univs,csts) s = + (* Universes that are not necessary to typecheck the term. + E.g. univs introduced by tactics and not used in the proof term. *) + let diff = LSet.diff univs s in + let rec aux diff candid univs ness = + let (diff', candid', univs', ness') = + Constraint.fold + (fun (l, d, r as c) (diff, candid, univs, csts) -> + if not (LSet.mem l diff) then + (LSet.remove r diff, candid, univs, Constraint.add c csts) + else if not (LSet.mem r diff) then + (LSet.remove l diff, candid, univs, Constraint.add c csts) + else (diff, Constraint.add c candid, univs, csts)) + candid (diff, Constraint.empty, univs, ness) + in + if ness' == ness then (LSet.diff univs diff', ness) + else aux diff' candid' univs' ness' + in aux diff csts univs Constraint.empty diff --git a/library/univops.mli b/library/univops.mli new file mode 100644 index 000000000..5b499c75b --- /dev/null +++ b/library/univops.mli @@ -0,0 +1,17 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* universe_set +val universes_of_inductive : mutual_inductive_body -> universe_set +val restrict_universe_context : universe_context_set -> universe_set -> universe_context_set -- cgit v1.2.3